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