1(*  Title:      Tools/nbe.ML
2    Authors:    Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen
3
4Normalization by evaluation, based on generic code generator.
5*)
6
7signature NBE =
8sig
9  val dynamic_conv: Proof.context -> conv
10  val dynamic_value: Proof.context -> term -> term
11  val static_conv: { ctxt: Proof.context, consts: string list }
12    -> Proof.context -> conv
13  val static_value: { ctxt: Proof.context, consts: string list }
14    -> Proof.context -> term -> term
15
16  datatype Univ =
17      Const of int * Univ list               (*named (uninterpreted) constants*)
18    | DFree of string * int                  (*free (uninterpreted) dictionary parameters*)
19    | BVar of int * Univ list
20    | Abs of (int * (Univ list -> Univ)) * Univ list
21  val apps: Univ -> Univ list -> Univ        (*explicit applications*)
22  val abss: int -> (Univ list -> Univ) -> Univ
23                                             (*abstractions as closures*)
24  val same: Univ * Univ -> bool
25
26  val put_result: (unit -> Univ list -> Univ list)
27    -> Proof.context -> Proof.context
28  val trace: bool Config.T
29
30  val add_const_alias: thm -> theory -> theory
31end;
32
33structure Nbe: NBE =
34struct
35
36(* generic non-sense *)
37
38val trace = Attrib.setup_config_bool \<^binding>\<open>nbe_trace\<close> (K false);
39fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x;
40
41
42(** certificates and oracle for "trivial type classes" **)
43
44structure Triv_Class_Data = Theory_Data
45(
46  type T = (class * thm) list;
47  val empty = [];
48  val extend = I;
49  fun merge data : T = AList.merge (op =) (K true) data;
50);
51
52fun add_const_alias thm thy =
53  let
54    val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
55     of SOME ofclass_eq => ofclass_eq
56      | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm);
57    val (T, class) = case try Logic.dest_of_class ofclass
58     of SOME T_class => T_class
59      | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm);
60    val tvar = case try Term.dest_TVar T
61     of SOME (tvar as (_, sort)) => if null (filter (can (Axclass.get_info thy)) sort)
62          then tvar
63          else error ("Bad sort: " ^ Thm.string_of_thm_global thy thm)
64      | _ => error ("Bad type: " ^ Thm.string_of_thm_global thy thm);
65    val _ = if Term.add_tvars eqn [] = [tvar] then ()
66      else error ("Inconsistent type: " ^ Thm.string_of_thm_global thy thm);
67    val lhs_rhs = case try Logic.dest_equals eqn
68     of SOME lhs_rhs => lhs_rhs
69      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
70    val c_c' = case try (apply2 (Axclass.unoverload_const thy o dest_Const)) lhs_rhs
71     of SOME c_c' => c_c'
72      | _ => error ("Not an equation with two constants: "
73          ^ Syntax.string_of_term_global thy eqn);
74    val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then ()
75      else error ("Inconsistent class: " ^ Thm.string_of_thm_global thy thm);
76  in Triv_Class_Data.map (AList.update (op =) (class, Thm.trim_context thm)) thy end;
77
78local
79
80val get_triv_classes = map fst o Triv_Class_Data.get;
81
82val (_, triv_of_class) = Context.>>> (Context.map_theory_result
83  (Thm.add_oracle (\<^binding>\<open>triv_of_class\<close>, fn (thy, T, class) =>
84    Thm.global_cterm_of thy (Logic.mk_of_class (T, class)))));
85
86in
87
88fun lift_triv_classes_conv orig_ctxt conv ct =
89  let
90    val thy = Proof_Context.theory_of orig_ctxt;
91    val ctxt = Proof_Context.init_global thy;
92      (*FIXME quasi-global context*)
93    val algebra = Sign.classes_of thy;
94    val triv_classes = get_triv_classes thy;
95    fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes;
96    fun mk_entry (v, sort) =
97      let
98        val T = TFree (v, sort);
99        val cT = Thm.ctyp_of ctxt T;
100        val triv_sort = additional_classes sort;
101      in
102        (v, (Sorts.inter_sort algebra (sort, triv_sort),
103          (cT, AList.make (fn class => Thm.of_class (cT, class)) sort
104            @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort)))
105      end;
106    val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []);
107    fun instantiate thm =
108      let
109        val tvars =
110          Term.add_tvars (#1 (Logic.dest_equals (Logic.strip_imp_concl (Thm.prop_of thm)))) [];
111        val instT = map2 (fn v => fn (_, (_, (cT, _))) => (v, cT)) tvars vs_tab;
112      in Thm.instantiate (instT, []) thm end;
113    fun of_class (TFree (v, _), class) =
114          the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class)
115      | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ ctxt T);
116    fun strip_of_class thm =
117      let
118        val prems_of_class = Thm.prop_of thm
119          |> Logic.strip_imp_prems
120          |> map (Logic.dest_of_class #> of_class);
121      in fold Thm.elim_implies prems_of_class thm end;
122  in
123    ct
124    |> Thm.term_of
125    |> (map_types o map_type_tfree)
126        (fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v))
127    |> Thm.cterm_of ctxt
128    |> conv ctxt
129    |> Thm.strip_shyps
130    |> Thm.varifyT_global
131    |> Thm.unconstrainT
132    |> instantiate
133    |> strip_of_class
134  end;
135
136fun lift_triv_classes_rew ctxt rew t =
137  let
138    val thy = Proof_Context.theory_of ctxt;
139    val algebra = Sign.classes_of thy;
140    val triv_classes = get_triv_classes thy;
141    val vs = Term.add_tfrees t [];
142  in
143    t
144    |> (map_types o map_type_tfree)
145        (fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes)))
146    |> rew
147    |> (map_types o map_type_tfree)
148        (fn (v, sort) => TFree (v, the_default sort (AList.lookup (op =) vs v)))
149  end;
150
151end;
152
153
154(** the semantic universe **)
155
156(*
157   Functions are given by their semantical function value. To avoid
158   trouble with the ML-type system, these functions have the most
159   generic type, that is "Univ list -> Univ". The calling convention is
160   that the arguments come as a list, the last argument first. In
161   other words, a function call that usually would look like
162
163   f x_1 x_2 ... x_n   or   f(x_1,x_2, ..., x_n)
164
165   would be in our convention called as
166
167              f [x_n,..,x_2,x_1]
168
169   Moreover, to handle functions that are still waiting for some
170   arguments we have additionally a list of arguments collected to far
171   and the number of arguments we're still waiting for.
172*)
173
174datatype Univ =
175    Const of int * Univ list           (*named (uninterpreted) constants*)
176  | DFree of string * int              (*free (uninterpreted) dictionary parameters*)
177  | BVar of int * Univ list            (*bound variables, named*)
178  | Abs of (int * (Univ list -> Univ)) * Univ list
179                                       (*abstractions as closures*);
180
181
182(* constructor functions *)
183
184fun abss n f = Abs ((n, f), []);
185fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in
186      case int_ord (k, 0)
187       of EQUAL => f (ys @ xs)
188        | LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end
189        | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*)
190      end
191  | apps (Const (name, xs)) ys = Const (name, ys @ xs)
192  | apps (BVar (n, xs)) ys = BVar (n, ys @ xs);
193
194fun same (Const (k, xs), Const (l, ys)) = k = l andalso eq_list same (xs, ys)
195  | same (DFree (s, k), DFree (t, l)) = s = t andalso k = l
196  | same (BVar (k, xs), BVar (l, ys)) = k = l andalso eq_list same (xs, ys)
197  | same _ = false;
198
199
200(** assembling and compiling ML code from terms **)
201
202(* abstract ML syntax *)
203
204infix 9 `$` `$$`;
205fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")";
206fun e `$$` [] = e
207  | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")";
208fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")";
209
210fun ml_cases t cs =
211  "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")";
212fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end";
213fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")";
214
215fun ml_and [] = "true"
216  | ml_and [x] = x
217  | ml_and xs = "(" ^ space_implode " andalso " xs ^ ")";
218fun ml_if b x y = "(if " ^ b ^ " then " ^ x ^ " else " ^ y ^ ")";
219
220fun ml_list es = "[" ^ commas es ^ "]";
221
222fun ml_fundefs ([(name, [([], e)])]) =
223      "val " ^ name ^ " = " ^ e ^ "\n"
224  | ml_fundefs (eqs :: eqss) =
225      let
226        fun fundef (name, eqs) =
227          let
228            fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e
229          in space_implode "\n  | " (map eqn eqs) end;
230      in
231        (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss
232        |> cat_lines
233        |> suffix "\n"
234      end;
235
236
237(* nbe specific syntax and sandbox communication *)
238
239structure Univs = Proof_Data
240(
241  type T = unit -> Univ list -> Univ list;
242  val empty: T = fn () => raise Fail "Univs";
243  fun init _ = empty;
244);
245val get_result = Univs.get;
246val put_result = Univs.put;
247
248local
249  val prefix = "Nbe.";
250  val name_put = prefix ^ "put_result";
251  val name_const = prefix ^ "Const";
252  val name_abss = prefix ^ "abss";
253  val name_apps = prefix ^ "apps";
254  val name_same = prefix ^ "same";
255in
256
257val univs_cookie = (get_result, put_result, name_put);
258
259fun nbe_fun idx_of 0 (Code_Symbol.Constant "") = "nbe_value"
260  | nbe_fun idx_of i sym = "c_" ^ string_of_int (idx_of sym)
261      ^ "_" ^ Code_Symbol.default_base sym ^ "_" ^ string_of_int i;
262fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n;
263fun nbe_bound v = "v_" ^ v;
264fun nbe_bound_optional NONE = "_"
265  | nbe_bound_optional (SOME v) = nbe_bound v;
266fun nbe_default v = "w_" ^ v;
267
268(*note: these three are the "turning spots" where proper argument order is established!*)
269fun nbe_apps t [] = t
270  | nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)];
271fun nbe_apps_local idx_of i c ts = nbe_fun idx_of i c `$` ml_list (rev ts);
272fun nbe_apps_constr ctxt idx_of c ts =
273  let
274    val c' = if Config.get ctxt trace
275      then string_of_int (idx_of c) ^ " (*" ^ Code_Symbol.default_base c ^ "*)"
276      else string_of_int (idx_of c);
277  in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end;
278
279fun nbe_abss 0 f = f `$` ml_list []
280  | nbe_abss n f = name_abss `$$` [string_of_int n, f];
281
282fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))";
283
284end;
285
286open Basic_Code_Symbol;
287open Basic_Code_Thingol;
288
289
290(* code generation *)
291
292fun assemble_eqnss ctxt idx_of deps eqnss =
293  let
294    fun prep_eqns (c, (vs, eqns)) =
295      let
296        val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs;
297        val num_args = length dicts + ((length o fst o hd) eqns);
298      in (c, (num_args, (dicts, eqns))) end;
299    val eqnss' = map prep_eqns eqnss;
300
301    fun assemble_constapp sym dss ts = 
302      let
303        val ts' = (maps o map) assemble_dict dss @ ts;
304      in case AList.lookup (op =) eqnss' sym
305       of SOME (num_args, _) => if num_args <= length ts'
306            then let val (ts1, ts2) = chop num_args ts'
307            in nbe_apps (nbe_apps_local idx_of 0 sym ts1) ts2
308            end else nbe_apps (nbe_abss num_args (nbe_fun idx_of 0 sym)) ts'
309        | NONE => if member (op =) deps sym
310            then nbe_apps (nbe_fun idx_of 0 sym) ts'
311            else nbe_apps_constr ctxt idx_of sym ts'
312      end
313    and assemble_classrels classrels =
314      fold_rev (fn classrel => assemble_constapp (Class_Relation classrel) [] o single) classrels
315    and assemble_dict (Dict (classrels, x)) =
316          assemble_classrels classrels (assemble_plain_dict x)
317    and assemble_plain_dict (Dict_Const (inst, dss)) =
318          assemble_constapp (Class_Instance inst) dss []
319      | assemble_plain_dict (Dict_Var { var, index, ... }) =
320          nbe_dict var index
321
322    fun assemble_iterm constapp =
323      let
324        fun of_iterm match_cont t =
325          let
326            val (t', ts) = Code_Thingol.unfold_app t
327          in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end
328        and of_iapp match_cont (IConst { sym, dicts = dss, ... }) ts = constapp sym dss ts
329          | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts
330          | of_iapp match_cont ((v, _) `|=> t) ts =
331              nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts
332          | of_iapp match_cont (ICase { term = t, clauses = clauses, primitive = t0, ... }) ts =
333              nbe_apps (ml_cases (of_iterm NONE t)
334                (map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) clauses
335                  @ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts
336      in of_iterm end;
337
338    fun subst_nonlin_vars args =
339      let
340        val vs = (fold o Code_Thingol.fold_varnames)
341          (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args [];
342        val names = Name.make_context (map fst vs);
343        fun declare v k ctxt =
344          let val vs = Name.invent ctxt v k
345          in (vs, fold Name.declare vs ctxt) end;
346        val (vs_renames, _) = fold_map (fn (v, k) => if k > 1
347          then declare v (k - 1) #>> (fn vs => (v, vs))
348          else pair (v, [])) vs names;
349        val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames;
350        fun subst_vars (t as IConst _) samepairs = (t, samepairs)
351          | subst_vars (t as IVar NONE) samepairs = (t, samepairs)
352          | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v
353             of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs)
354              | NONE => (t, samepairs))
355          | subst_vars (t1 `$ t2) samepairs = samepairs
356              |> subst_vars t1
357              ||>> subst_vars t2
358              |>> (op `$)
359          | subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs;
360        val (args', _) = fold_map subst_vars args samepairs;
361      in (samepairs, args') end;
362
363    fun assemble_eqn sym dicts default_args (i, (args, rhs)) =
364      let
365        val match_cont = if Code_Symbol.is_value sym then NONE
366          else SOME (nbe_apps_local idx_of (i + 1) sym (dicts @ default_args));
367        val assemble_arg = assemble_iterm
368          (fn sym' => fn dss => fn ts => nbe_apps_constr ctxt idx_of sym' ((maps o map) (K "_")
369            dss @ ts)) NONE;
370        val assemble_rhs = assemble_iterm assemble_constapp match_cont;
371        val (samepairs, args') = subst_nonlin_vars args;
372        val s_args = map assemble_arg args';
373        val s_rhs = if null samepairs then assemble_rhs rhs
374          else ml_if (ml_and (map nbe_same samepairs))
375            (assemble_rhs rhs) (the match_cont);
376        val eqns = case match_cont
377         of NONE => [([ml_list (rev (dicts @ s_args))], s_rhs)]
378          | SOME default_rhs =>
379              [([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs),
380                ([ml_list (rev (dicts @ default_args))], default_rhs)]
381      in (nbe_fun idx_of i sym, eqns) end;
382
383    fun assemble_eqns (sym, (num_args, (dicts, eqns))) =
384      let
385        val default_args = map nbe_default
386          (Name.invent Name.context "a" (num_args - length dicts));
387        val eqns' = map_index (assemble_eqn sym dicts default_args) eqns
388          @ (if Code_Symbol.is_value sym then [] else [(nbe_fun idx_of (length eqns) sym,
389            [([ml_list (rev (dicts @ default_args))],
390              nbe_apps_constr ctxt idx_of sym (dicts @ default_args))])]);
391      in (eqns', nbe_abss num_args (nbe_fun idx_of 0 sym)) end;
392
393    val (fun_vars, fun_vals) = map_split assemble_eqns eqnss';
394    val deps_vars = ml_list (map (nbe_fun idx_of 0) deps);
395  in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end;
396
397
398(* compilation of equations *)
399
400fun compile_eqnss ctxt nbe_program raw_deps [] = []
401  | compile_eqnss ctxt nbe_program raw_deps eqnss =
402      let
403        val (deps, deps_vals) = split_list (map_filter
404          (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps);
405        val idx_of = raw_deps
406          |> map (fn dep => (dep, snd (Code_Symbol.Graph.get_node nbe_program dep)))
407          |> AList.lookup (op =)
408          |> (fn f => the o f);
409        val s = assemble_eqnss ctxt idx_of deps eqnss;
410        val cs = map fst eqnss;
411      in
412        s
413        |> traced ctxt (fn s => "\n--- code to be evaluated:\n" ^ s)
414        |> pair ""
415        |> Code_Runtime.value ctxt univs_cookie
416        |> (fn f => f deps_vals)
417        |> (fn univs => cs ~~ univs)
418      end;
419
420
421(* extraction of equations from statements *)
422
423fun dummy_const sym dss =
424  IConst { sym = sym, typargs = [], dicts = dss,
425    dom = [], annotation = NONE };
426
427fun eqns_of_stmt (_, Code_Thingol.NoStmt) =
428      []
429  | eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) =
430      []
431  | eqns_of_stmt (sym_const, Code_Thingol.Fun (((vs, _), eqns), _)) =
432      [(sym_const, (vs, map fst eqns))]
433  | eqns_of_stmt (_, Code_Thingol.Datatypecons _) =
434      []
435  | eqns_of_stmt (_, Code_Thingol.Datatype _) =
436      []
437  | eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) =
438      let
439        val syms = map Class_Relation classrels @ map (Constant o fst) classparams;
440        val params = Name.invent Name.context "d" (length syms);
441        fun mk (k, sym) =
442          (sym, ([(v, [])],
443            [([dummy_const sym_class [] `$$ map (IVar o SOME) params],
444              IVar (SOME (nth params k)))]));
445      in map_index mk syms end
446  | eqns_of_stmt (_, Code_Thingol.Classrel _) =
447      []
448  | eqns_of_stmt (_, Code_Thingol.Classparam _) =
449      []
450  | eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) =
451      [(sym_inst, (vs, [([], dummy_const (Type_Class class) [] `$$
452        map (fn (class, dss) => dummy_const (Class_Instance (tyco, class)) dss) superinsts
453        @ map (IConst o fst o snd o fst) inst_params)]))];
454
455
456(* compilation of whole programs *)
457
458fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) =
459  if can (Code_Symbol.Graph.get_node nbe_program) name
460  then (nbe_program, (maxidx, idx_tab))
461  else (Code_Symbol.Graph.new_node (name, (NONE, maxidx)) nbe_program,
462    (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab));
463
464fun compile_stmts ctxt stmts_deps =
465  let
466    val names = map (fst o fst) stmts_deps;
467    val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps;
468    val eqnss = maps (eqns_of_stmt o fst) stmts_deps;
469    val refl_deps = names_deps
470      |> maps snd
471      |> distinct (op =)
472      |> fold (insert (op =)) names;
473    fun compile nbe_program = eqnss
474      |> compile_eqnss ctxt nbe_program refl_deps
475      |> rpair nbe_program;
476  in
477    fold ensure_const_idx refl_deps
478    #> apfst (fold (fn (name, deps) => fold (curry Code_Symbol.Graph.add_edge name) deps) names_deps
479      #> compile
480      #-> fold (fn (name, univ) => (Code_Symbol.Graph.map_node name o apfst) (K (SOME univ))))
481  end;
482
483fun compile_program { ctxt, program } =
484  let
485    fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names
486      then (nbe_program, (maxidx, idx_tab))
487      else (nbe_program, (maxidx, idx_tab))
488        |> compile_stmts ctxt (map (fn name => ((name, Code_Symbol.Graph.get_node program name),
489          Code_Symbol.Graph.immediate_succs program name)) names);
490  in
491    fold_rev add_stmts (Code_Symbol.Graph.strong_conn program)
492  end;
493
494
495(** normalization **)
496
497(* compilation and reconstruction of terms *)
498
499fun compile_term { ctxt, nbe_program, deps, term = (vs, t) } =
500  let 
501    val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs;
502  in
503    (Code_Symbol.value, (vs, [([], t)]))
504    |> singleton (compile_eqnss ctxt nbe_program deps)
505    |> snd
506    |> (fn t => apps t (rev dict_frees))
507  end;
508
509fun reconstruct_term ctxt (idx_tab : Code_Symbol.T Inttab.table) t =
510  let
511    fun take_until f [] = []
512      | take_until f (x :: xs) = if f x then [] else x :: take_until f xs;
513    fun is_dict (Const (idx, _)) =
514          (case Inttab.lookup idx_tab idx of
515            SOME (Constant _) => false
516          | _ => true)
517      | is_dict (DFree _) = true
518      | is_dict _ = false;
519    fun const_of_idx idx =
520      case Inttab.lookup idx_tab idx of SOME (Constant const) => const;
521    fun of_apps bounds (t, ts) =
522      fold_map (of_univ bounds) ts
523      #>> (fn ts' => list_comb (t, rev ts'))
524    and of_univ bounds (Const (idx, ts)) typidx =
525          let
526            val ts' = take_until is_dict ts;
527            val const = const_of_idx idx;
528            val T = map_type_tvar (fn ((v, i), _) =>
529              Type_Infer.param typidx (v ^ string_of_int i, []))
530                (Sign.the_const_type (Proof_Context.theory_of ctxt) const);
531            val typidx' = typidx + 1;
532          in of_apps bounds (Term.Const (const, T), ts') typidx' end
533      | of_univ bounds (BVar (n, ts)) typidx =
534          of_apps bounds (Bound (bounds - n - 1), ts) typidx
535      | of_univ bounds (t as Abs _) typidx =
536          typidx
537          |> of_univ (bounds + 1) (apps t [BVar (bounds, [])])
538          |-> (fn t' => pair (Term.Abs ("u", dummyT, t')))
539  in of_univ 0 t 0 |> fst end;
540
541fun compile_and_reconstruct_term { ctxt, nbe_program, idx_tab, deps, term } =
542  compile_term
543    { ctxt = ctxt, nbe_program = nbe_program, deps = deps, term = term }
544  |> reconstruct_term ctxt idx_tab;
545
546fun normalize_term (nbe_program, idx_tab) raw_ctxt t_original ((vs, ty) : typscheme, t) deps =
547  let
548    val ctxt = Syntax.init_pretty_global (Proof_Context.theory_of raw_ctxt);
549    val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt);
550    fun type_infer t' =
551      Syntax.check_term
552        (ctxt
553          |> Config.put Type_Infer.object_logic false
554          |> Config.put Type_Infer_Context.const_sorts false)
555        (Type.constraint (fastype_of t_original) t');
556    fun check_tvars t' =
557      if null (Term.add_tvars t' []) then t'
558      else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t');
559  in
560    Code_Preproc.timed "computing NBE expression" #ctxt compile_and_reconstruct_term
561      { ctxt = ctxt, nbe_program = nbe_program, idx_tab = idx_tab, deps = deps, term = (vs, t) }
562    |> traced ctxt (fn t => "Normalized:\n" ^ string_of_term t)
563    |> type_infer
564    |> traced ctxt (fn t => "Types inferred:\n" ^ string_of_term t)
565    |> check_tvars
566    |> traced ctxt (fn _ => "---\n")
567  end;
568
569
570(* function store *)
571
572structure Nbe_Functions = Code_Data
573(
574  type T = (Univ option * int) Code_Symbol.Graph.T * (int * Code_Symbol.T Inttab.table);
575  val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty));
576);
577
578fun compile ignore_cache ctxt program =
579  let
580    val (nbe_program, (_, idx_tab)) =
581      Nbe_Functions.change (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt))
582        (Code_Preproc.timed "compiling NBE program" #ctxt
583          compile_program { ctxt = ctxt, program = program });
584  in (nbe_program, idx_tab) end;
585
586
587(* evaluation oracle *)
588
589fun mk_equals ctxt lhs raw_rhs =
590  let
591    val ty = Thm.typ_of_cterm lhs;
592    val eq = Thm.cterm_of ctxt (Term.Const (\<^const_name>\<open>Pure.eq\<close>, ty --> ty --> propT));
593    val rhs = Thm.cterm_of ctxt raw_rhs;
594  in Thm.mk_binop eq lhs rhs end;
595
596val (_, raw_oracle) = Context.>>> (Context.map_theory_result
597  (Thm.add_oracle (\<^binding>\<open>normalization_by_evaluation\<close>,
598    fn (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct) =>
599      mk_equals ctxt ct (normalize_term nbe_program_idx_tab ctxt (Thm.term_of ct) vs_ty_t deps))));
600
601fun oracle nbe_program_idx_tab ctxt vs_ty_t deps ct =
602  raw_oracle (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct);
603
604fun dynamic_conv ctxt = lift_triv_classes_conv ctxt
605  (fn ctxt' => Code_Thingol.dynamic_conv ctxt' (fn program =>
606    oracle (compile false ctxt program) ctxt'));
607
608fun dynamic_value ctxt = lift_triv_classes_rew ctxt
609  (Code_Thingol.dynamic_value ctxt I (fn program =>
610    normalize_term (compile false ctxt program) ctxt));
611
612fun static_conv (ctxt_consts as { ctxt, ... }) =
613  let
614    val conv = Code_Thingol.static_conv_thingol ctxt_consts
615      (fn { program, deps = _ } => oracle (compile true ctxt program));
616  in fn ctxt' => lift_triv_classes_conv ctxt' conv end;
617
618fun static_value { ctxt, consts } =
619  let
620    val comp = Code_Thingol.static_value { ctxt = ctxt, lift_postproc = I, consts = consts }
621      (fn { program, deps = _ } => normalize_term (compile false ctxt program));
622  in fn ctxt' => lift_triv_classes_rew ctxt' (comp ctxt') end;
623
624end;
625