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