| author | bulwahn | 
| Fri, 11 Nov 2011 08:32:48 +0100 | |
| changeset 45451 | 74515e8e6046 | 
| parent 44789 | 5a062c23c7db | 
| child 47572 | 1e18bbfb40cb | 
| permissions | -rw-r--r-- | 
| 24590 | 1 | (* Title: Tools/nbe.ML | 
| 24155 | 2 | Authors: Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen | 
| 3 | ||
| 24839 | 4 | Normalization by evaluation, based on generic code generator. | 
| 24155 | 5 | *) | 
| 6 | ||
| 7 | signature NBE = | |
| 8 | sig | |
| 41247 | 9 | val dynamic_conv: theory -> conv | 
| 41184 | 10 | val dynamic_value: theory -> term -> term | 
| 11 | val static_conv: theory -> string list -> conv | |
| 12 | val static_value: theory -> string list -> term -> term | |
| 25101 | 13 | |
| 25204 | 14 | datatype Univ = | 
| 26064 | 15 | Const of int * Univ list (*named (uninterpreted) constants*) | 
| 25924 | 16 | | DFree of string * int (*free (uninterpreted) dictionary parameters*) | 
| 24155 | 17 | | BVar of int * Univ list | 
| 28350 | 18 | | Abs of (int * (Univ list -> Univ)) * Univ list | 
| 25924 | 19 | val apps: Univ -> Univ list -> Univ (*explicit applications*) | 
| 25944 | 20 | val abss: int -> (Univ list -> Univ) -> Univ | 
| 28337 | 21 | (*abstractions as closures*) | 
| 41037 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 22 | val same: Univ * Univ -> bool | 
| 24155 | 23 | |
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39290diff
changeset | 24 | val put_result: (unit -> Univ list -> Univ list) -> Proof.context -> Proof.context | 
| 32740 | 25 | val trace: bool Unsynchronized.ref | 
| 25924 | 26 | |
| 26747 
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
 haftmann parents: 
26739diff
changeset | 27 | val setup: theory -> theory | 
| 32544 | 28 | val add_const_alias: thm -> theory -> theory | 
| 24155 | 29 | end; | 
| 30 | ||
| 31 | structure Nbe: NBE = | |
| 32 | struct | |
| 33 | ||
| 34 | (* generic non-sense *) | |
| 35 | ||
| 32740 | 36 | val trace = Unsynchronized.ref false; | 
| 32544 | 37 | fun traced f x = if !trace then (tracing (f x); x) else x; | 
| 24155 | 38 | |
| 39 | ||
| 32544 | 40 | (** certificates and oracle for "trivial type classes" **) | 
| 41 | ||
| 33522 | 42 | structure Triv_Class_Data = Theory_Data | 
| 32544 | 43 | ( | 
| 44 | type T = (class * thm) list; | |
| 45 | val empty = []; | |
| 46 | val extend = I; | |
| 33522 | 47 | fun merge data : T = AList.merge (op =) (K true) data; | 
| 32544 | 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);
 | |
| 40362 
82a066bff182
Code.check_const etc.: reject too specific types
 haftmann parents: 
39911diff
changeset | 68 | val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs | 
| 32544 | 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 | |
| 43619 
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
 wenzelm parents: 
43329diff
changeset | 81 |   (Thm.add_oracle (@{binding triv_of_class}, fn (thy, T, class) =>
 | 
| 32544 | 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; | |
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 89 | val certT = Thm.ctyp_of thy; | 
| 32544 | 90 | val triv_classes = get_triv_classes thy; | 
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 91 | fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 92 | fun mk_entry (v, sort) = | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 93 | let | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 94 | val T = TFree (v, sort); | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 95 | val cT = certT T; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 96 | val triv_sort = additional_classes sort; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 97 | in | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 98 | (v, (Sorts.inter_sort algebra (sort, triv_sort), | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 99 | (cT, AList.make (fn class => Thm.of_class (cT, class)) sort | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 100 | @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort))) | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 101 | end; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 102 | val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []); | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 103 | fun instantiate thm = | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 104 | let | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 105 | val cert_tvars = map (certT o TVar) (Term.add_tvars | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 106 | ((fst o Logic.dest_equals o Logic.strip_imp_concl o Thm.prop_of) thm) []); | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 107 | val instantiation = | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 108 | map2 (fn cert_tvar => fn (_, (_, (cT, _))) => (cert_tvar, cT)) cert_tvars vs_tab; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 109 | in Thm.instantiate (instantiation, []) thm end; | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 110 | fun of_class (TFree (v, _), class) = | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 111 | the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class) | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 112 |       | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ_global thy T);
 | 
| 32544 | 113 | fun strip_of_class thm = | 
| 114 | let | |
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 115 | val prems_of_class = Thm.prop_of thm | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 116 | |> Logic.strip_imp_prems | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 117 | |> map (Logic.dest_of_class #> of_class); | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 118 | in fold Thm.elim_implies prems_of_class thm end; | 
| 36770 
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
 wenzelm parents: 
36768diff
changeset | 119 | in | 
| 
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
 wenzelm parents: 
36768diff
changeset | 120 | ct | 
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 121 | |> (Drule.cterm_fun o map_types o map_type_tfree) | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 122 | (fn (v, sort) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v)) | 
| 32544 | 123 | |> conv | 
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 124 | |> Thm.strip_shyps | 
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35500diff
changeset | 125 | |> Thm.varifyT_global | 
| 36982 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 126 | |> Thm.unconstrainT | 
| 
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
 haftmann parents: 
36960diff
changeset | 127 | |> instantiate | 
| 32544 | 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 **) | |
| 24155 | 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 | ||
| 25204 | 167 | datatype Univ = | 
| 26064 | 168 | Const of int * Univ list (*named (uninterpreted) constants*) | 
| 25924 | 169 | | DFree of string * int (*free (uninterpreted) dictionary parameters*) | 
| 27499 | 170 | | BVar of int * Univ list (*bound variables, named*) | 
| 24155 | 171 | | Abs of (int * (Univ list -> Univ)) * Univ list | 
| 27499 | 172 | (*abstractions as closures*); | 
| 24155 | 173 | |
| 35371 | 174 | |
| 24155 | 175 | (* constructor functions *) | 
| 176 | ||
| 25944 | 177 | fun abss n f = Abs ((n, f), []); | 
| 25924 | 178 | fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in | 
| 27499 | 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 | |
| 25924 | 184 | | apps (Const (name, xs)) ys = Const (name, ys @ xs) | 
| 27499 | 185 | | apps (BVar (n, xs)) ys = BVar (n, ys @ xs); | 
| 24155 | 186 | |
| 41037 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 187 | fun same (Const (k, xs), Const (l, ys)) = k = l andalso eq_list same (xs, ys) | 
| 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 188 | | same (DFree (s, k), DFree (t, l)) = s = t andalso k = l | 
| 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 189 | | same (BVar (k, xs), BVar (l, ys)) = k = l andalso eq_list same (xs, ys) | 
| 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 190 | | same _ = false; | 
| 40730 
2aa0390a2da7
nbe decides equality of abstractions by extensionality
 haftmann parents: 
40564diff
changeset | 191 | |
| 24155 | 192 | |
| 193 | (** assembling and compiling ML code from terms **) | |
| 194 | ||
| 195 | (* abstract ML syntax *) | |
| 196 | ||
| 197 | infix 9 `$` `$$`; | |
| 198 | fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")";
 | |
| 25101 | 199 | fun e `$$` [] = e | 
| 200 |   | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")";
 | |
| 24590 | 201 | fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")"; | 
| 24155 | 202 | |
| 203 | fun ml_cases t cs = | |
| 204 | "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")"; | |
| 25944 | 205 | fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end"; | 
| 28337 | 206 | fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")";
 | 
| 24155 | 207 | |
| 28350 | 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 | ||
| 24155 | 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 | |
| 26931 | 225 | |> cat_lines | 
| 24155 | 226 | |> suffix "\n" | 
| 227 | end; | |
| 228 | ||
| 35371 | 229 | |
| 25944 | 230 | (* nbe specific syntax and sandbox communication *) | 
| 231 | ||
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41346diff
changeset | 232 | structure Univs = Proof_Data | 
| 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41346diff
changeset | 233 | ( | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39290diff
changeset | 234 | type T = unit -> Univ list -> Univ list | 
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41346diff
changeset | 235 | (* FIXME avoid user error with non-user text *) | 
| 39399 | 236 | fun init _ () = error "Univs" | 
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39290diff
changeset | 237 | ); | 
| 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39290diff
changeset | 238 | val put_result = Univs.put; | 
| 24155 | 239 | |
| 240 | local | |
| 41068 | 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"; | |
| 24155 | 248 | in | 
| 249 | ||
| 39388 
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
 haftmann parents: 
39290diff
changeset | 250 | val univs_cookie = (Univs.get, put_result, name_put); | 
| 25944 | 251 | |
| 28337 | 252 | fun nbe_fun 0 "" = "nbe_value" | 
| 253 | | nbe_fun i c = "c_" ^ translate_string (fn "." => "_" | c => c) c ^ "_" ^ string_of_int i; | |
| 25101 | 254 | fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; | 
| 24155 | 255 | fun nbe_bound v = "v_" ^ v; | 
| 31888 | 256 | fun nbe_bound_optional NONE = "_" | 
| 35500 | 257 | | nbe_bound_optional (SOME v) = nbe_bound v; | 
| 28337 | 258 | fun nbe_default v = "w_" ^ v; | 
| 24155 | 259 | |
| 25924 | 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)]; | |
| 28337 | 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;
 | |
| 25924 | 269 | |
| 24155 | 270 | fun nbe_abss 0 f = f `$` ml_list [] | 
| 25944 | 271 | | nbe_abss n f = name_abss `$$` [string_of_int n, f]; | 
| 24155 | 272 | |
| 41037 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 273 | fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))";
 | 
| 28350 | 274 | |
| 24155 | 275 | end; | 
| 276 | ||
| 28054 | 277 | open Basic_Code_Thingol; | 
| 24155 | 278 | |
| 35371 | 279 | |
| 25865 | 280 | (* code generation *) | 
| 24155 | 281 | |
| 26064 | 282 | fun assemble_eqnss idx_of deps eqnss = | 
| 25944 | 283 | let | 
| 284 | fun prep_eqns (c, (vs, eqns)) = | |
| 25924 | 285 | let | 
| 25944 | 286 | val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; | 
| 28337 | 287 | val num_args = length dicts + ((length o fst o hd) eqns); | 
| 25944 | 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 | |
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 293 | val ts' = (maps o map) assemble_dict dss @ ts; | 
| 25944 | 294 | in case AList.lookup (op =) eqnss' c | 
| 28337 | 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' | |
| 25944 | 299 | | NONE => if member (op =) deps c | 
| 28337 | 300 | then nbe_apps (nbe_fun 0 c) ts' | 
| 301 | else nbe_apps_constr idx_of c ts' | |
| 25924 | 302 | end | 
| 41100 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
41068diff
changeset | 303 | and assemble_classrels classrels = | 
| 
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
 haftmann parents: 
41068diff
changeset | 304 | fold_rev (fn classrel => assemble_constapp classrel [] o single) classrels | 
| 41118 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 305 | and assemble_dict (Dict (classrels, x)) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 306 | assemble_classrels classrels (assemble_plain_dict x) | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 307 | and assemble_plain_dict (Dict_Const (inst, dss)) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 308 | assemble_constapp inst dss [] | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 309 | | assemble_plain_dict (Dict_Var (v, (n, _))) = | 
| 
b290841cd3b0
separated dictionary weakning into separate type
 haftmann parents: 
41100diff
changeset | 310 | nbe_dict v n | 
| 25924 | 311 | |
| 28350 | 312 | fun assemble_iterm constapp = | 
| 24155 | 313 | let | 
| 28350 | 314 | fun of_iterm match_cont t = | 
| 25944 | 315 | let | 
| 28054 | 316 | val (t', ts) = Code_Thingol.unfold_app t | 
| 28350 | 317 | in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end | 
| 44788 
8b935f1b3cf8
changing const type to pass along if typing annotations are necessary for disambigous terms
 bulwahn parents: 
44338diff
changeset | 318 | and of_iapp match_cont (IConst (c, (((_, dss), _), _))) ts = constapp c dss ts | 
| 31889 | 319 | | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts | 
| 31724 | 320 | | of_iapp match_cont ((v, _) `|=> t) ts = | 
| 31888 | 321 | nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts | 
| 28350 | 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
 | |
| 25944 | 326 | in of_iterm end; | 
| 24155 | 327 | |
| 28350 | 328 | fun subst_nonlin_vars args = | 
| 329 | let | |
| 330 | val vs = (fold o Code_Thingol.fold_varnames) | |
| 33002 | 331 | (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args []; | 
| 28350 | 332 | val names = Name.make_context (map fst vs); | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 333 | fun declare v k ctxt = | 
| 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 334 | let val vs = Name.invent ctxt v k | 
| 28350 | 335 | in (vs, fold Name.declare vs ctxt) end; | 
| 336 | val (vs_renames, _) = fold_map (fn (v, k) => if k > 1 | |
| 337 | then declare v (k - 1) #>> (fn vs => (v, vs)) | |
| 338 | else pair (v, [])) vs names; | |
| 339 | val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames; | |
| 340 | fun subst_vars (t as IConst _) samepairs = (t, samepairs) | |
| 31889 | 341 | | subst_vars (t as IVar NONE) samepairs = (t, samepairs) | 
| 342 | | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v | |
| 343 | of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs) | |
| 28350 | 344 | | NONE => (t, samepairs)) | 
| 345 | | subst_vars (t1 `$ t2) samepairs = samepairs | |
| 346 | |> subst_vars t1 | |
| 347 | ||>> subst_vars t2 | |
| 348 | |>> (op `$) | |
| 349 | | subst_vars (ICase (_, t)) samepairs = subst_vars t samepairs; | |
| 350 | val (args', _) = fold_map subst_vars args samepairs; | |
| 351 | in (samepairs, args') end; | |
| 352 | ||
| 28337 | 353 | fun assemble_eqn c dicts default_args (i, (args, rhs)) = | 
| 354 | let | |
| 355 | val is_eval = (c = ""); | |
| 356 | val default_rhs = nbe_apps_local (i+1) c (dicts @ default_args); | |
| 357 | val match_cont = if is_eval then NONE else SOME default_rhs; | |
| 28350 | 358 | val assemble_arg = assemble_iterm | 
| 359 | (fn c => fn _ => fn ts => nbe_apps_constr idx_of c ts) NONE; | |
| 35371 | 360 | val assemble_rhs = assemble_iterm assemble_constapp match_cont; | 
| 28350 | 361 | val (samepairs, args') = subst_nonlin_vars args; | 
| 362 | val s_args = map assemble_arg args'; | |
| 363 | val s_rhs = if null samepairs then assemble_rhs rhs | |
| 41037 
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
 haftmann parents: 
40730diff
changeset | 364 | else ml_if (ml_and (map nbe_same samepairs)) | 
| 28350 | 365 | (assemble_rhs rhs) default_rhs; | 
| 28337 | 366 | val eqns = if is_eval then | 
| 28350 | 367 | [([ml_list (rev (dicts @ s_args))], s_rhs)] | 
| 28337 | 368 | else | 
| 28350 | 369 | [([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs), | 
| 28337 | 370 | ([ml_list (rev (dicts @ default_args))], default_rhs)] | 
| 371 | in (nbe_fun i c, eqns) end; | |
| 372 | ||
| 25944 | 373 | fun assemble_eqns (c, (num_args, (dicts, eqns))) = | 
| 374 | let | |
| 28337 | 375 | val default_args = map nbe_default | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 376 | (Name.invent Name.context "a" (num_args - length dicts)); | 
| 28337 | 377 | val eqns' = map_index (assemble_eqn c dicts default_args) eqns | 
| 378 | @ (if c = "" then [] else [(nbe_fun (length eqns) c, | |
| 379 | [([ml_list (rev (dicts @ default_args))], | |
| 380 | nbe_apps_constr idx_of c (dicts @ default_args))])]); | |
| 381 | in (eqns', nbe_abss num_args (nbe_fun 0 c)) end; | |
| 24155 | 382 | |
| 25944 | 383 | val (fun_vars, fun_vals) = map_split assemble_eqns eqnss'; | 
| 28337 | 384 | val deps_vars = ml_list (map (nbe_fun 0) deps); | 
| 385 | in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end; | |
| 25944 | 386 | |
| 35371 | 387 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 388 | (* compile equations *) | 
| 25944 | 389 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 390 | fun compile_eqnss thy nbe_program raw_deps [] = [] | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 391 | | compile_eqnss thy nbe_program raw_deps eqnss = | 
| 24155 | 392 | let | 
| 42361 | 393 | val ctxt = Proof_Context.init_global thy; | 
| 26064 | 394 | val (deps, deps_vals) = split_list (map_filter | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 395 | (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Graph.get_node nbe_program dep)))) raw_deps); | 
| 26064 | 396 | val idx_of = raw_deps | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 397 | |> map (fn dep => (dep, snd (Graph.get_node nbe_program dep))) | 
| 26064 | 398 | |> AList.lookup (op =) | 
| 399 | |> (fn f => the o f); | |
| 400 | val s = assemble_eqnss idx_of deps eqnss; | |
| 24155 | 401 | val cs = map fst eqnss; | 
| 25944 | 402 | in | 
| 403 | s | |
| 32544 | 404 | |> traced (fn s => "\n--- code to be evaluated:\n" ^ s) | 
| 38931 
5e84c11c4b8a
evaluate takes ml context and ml expression parameter
 haftmann parents: 
38676diff
changeset | 405 | |> pair "" | 
| 39911 | 406 | |> Code_Runtime.value ctxt univs_cookie | 
| 25944 | 407 | |> (fn f => f deps_vals) | 
| 408 | |> (fn univs => cs ~~ univs) | |
| 409 | end; | |
| 25190 | 410 | |
| 30672 
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
 wenzelm parents: 
30288diff
changeset | 411 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 412 | (* extract equations from statements *) | 
| 24155 | 413 | |
| 37437 | 414 | fun eqns_of_stmt (_, Code_Thingol.Fun (_, ((_, []), _))) = | 
| 25101 | 415 | [] | 
| 37437 | 416 | | eqns_of_stmt (const, Code_Thingol.Fun (_, (((vs, _), eqns), _))) = | 
| 25101 | 417 | [(const, (vs, map fst eqns))] | 
| 28054 | 418 | | eqns_of_stmt (_, Code_Thingol.Datatypecons _) = | 
| 25101 | 419 | [] | 
| 28054 | 420 | | eqns_of_stmt (_, Code_Thingol.Datatype _) = | 
| 25101 | 421 | [] | 
| 37447 
ad3e04f289b6
transitive superclasses were also only a misunderstanding
 haftmann parents: 
37446diff
changeset | 422 | | eqns_of_stmt (class, Code_Thingol.Class (_, (v, (super_classes, classparams)))) = | 
| 25101 | 423 | let | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37146diff
changeset | 424 | val names = map snd super_classes @ map fst classparams; | 
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 425 | val params = Name.invent Name.context "d" (length names); | 
| 25101 | 426 | fun mk (k, name) = | 
| 427 | (name, ([(v, [])], | |
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 428 | [([IConst (class, ((([], []), ([], ITyVar "")), false)) `$$ map (IVar o SOME) params], | 
| 31889 | 429 | IVar (SOME (nth params k)))])); | 
| 25101 | 430 | in map_index mk names end | 
| 28054 | 431 | | eqns_of_stmt (_, Code_Thingol.Classrel _) = | 
| 25101 | 432 | [] | 
| 28054 | 433 | | eqns_of_stmt (_, Code_Thingol.Classparam _) = | 
| 25101 | 434 | [] | 
| 37449 | 435 | | eqns_of_stmt (inst, Code_Thingol.Classinst ((class, (_, arity_args)), (super_instances, (classparam_instances, _)))) = | 
| 44789 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 436 | [(inst, (arity_args, [([], IConst (class, ((([], []), ([], ITyVar "")), false)) `$$ | 
| 
5a062c23c7db
adding the body type as well to the code generation for constants as it is required for type annotations of constants
 bulwahn parents: 
44788diff
changeset | 437 | map (fn (_, (_, (inst, dss))) => IConst (inst, ((([], dss), ([], ITyVar "")), false))) super_instances | 
| 37384 
5aba26803073
more consistent naming aroud type classes and instances
 haftmann parents: 
37146diff
changeset | 438 | @ map (IConst o snd o fst) classparam_instances)]))]; | 
| 24155 | 439 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 440 | |
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 441 | (* compile whole programs *) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 442 | |
| 39399 | 443 | fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) = | 
| 444 | if can (Graph.get_node nbe_program) name | |
| 445 | then (nbe_program, (maxidx, idx_tab)) | |
| 446 | else (Graph.new_node (name, (NONE, maxidx)) nbe_program, | |
| 447 | (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab)); | |
| 448 | ||
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 449 | fun compile_stmts thy stmts_deps = | 
| 25101 | 450 | let | 
| 451 | val names = map (fst o fst) stmts_deps; | |
| 452 | val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; | |
| 453 | val eqnss = maps (eqns_of_stmt o fst) stmts_deps; | |
| 26064 | 454 | val refl_deps = names_deps | 
| 25190 | 455 | |> maps snd | 
| 456 | |> distinct (op =) | |
| 26064 | 457 | |> fold (insert (op =)) names; | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 458 | fun compile nbe_program = eqnss | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 459 | |> compile_eqnss thy nbe_program refl_deps | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 460 | |> rpair nbe_program; | 
| 25101 | 461 | in | 
| 39399 | 462 | fold ensure_const_idx refl_deps | 
| 26064 | 463 | #> apfst (fold (fn (name, deps) => fold (curry Graph.add_edge name) deps) names_deps | 
| 464 | #> compile | |
| 465 | #-> fold (fn (name, univ) => (Graph.map_node name o apfst) (K (SOME univ)))) | |
| 25101 | 466 | end; | 
| 24155 | 467 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 468 | fun compile_program thy program = | 
| 25101 | 469 | let | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 470 | fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Graph.get_node) nbe_program) names | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 471 | then (nbe_program, (maxidx, idx_tab)) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 472 | else (nbe_program, (maxidx, idx_tab)) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 473 | |> compile_stmts thy (map (fn name => ((name, Graph.get_node program name), | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
43619diff
changeset | 474 | Graph.immediate_succs program name)) names); | 
| 28706 | 475 | in | 
| 476 | fold_rev add_stmts (Graph.strong_conn program) | |
| 477 | end; | |
| 24155 | 478 | |
| 25944 | 479 | |
| 480 | (** evaluation **) | |
| 481 | ||
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 482 | (* term evaluation by compilation *) | 
| 25944 | 483 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 484 | fun compile_term thy nbe_program deps (vs : (string * sort) list, t) = | 
| 25924 | 485 | let | 
| 486 | val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; | |
| 487 | in | |
| 31064 | 488 |     ("", (vs, [([], t)]))
 | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 489 | |> singleton (compile_eqnss thy nbe_program deps) | 
| 25924 | 490 | |> snd | 
| 31064 | 491 | |> (fn t => apps t (rev dict_frees)) | 
| 25924 | 492 | end; | 
| 24155 | 493 | |
| 35371 | 494 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 495 | (* reconstruction *) | 
| 24155 | 496 | |
| 30947 | 497 | fun typ_of_itype program vs (ityco `%% itys) = | 
| 498 | let | |
| 499 | val Code_Thingol.Datatype (tyco, _) = Graph.get_node program ityco; | |
| 500 | in Type (tyco, map (typ_of_itype program vs) itys) end | |
| 501 | | typ_of_itype program vs (ITyVar v) = | |
| 502 | let | |
| 503 | val sort = (the o AList.lookup (op =) vs) v; | |
| 504 |       in TFree ("'" ^ v, sort) end;
 | |
| 505 | ||
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 506 | fun term_of_univ thy program idx_tab t = | 
| 24155 | 507 | let | 
| 25101 | 508 | fun take_until f [] = [] | 
| 509 | | take_until f (x::xs) = if f x then [] else x :: take_until f xs; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 510 | fun is_dict (Const (idx, _)) = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 511 | of Code_Thingol.Class _ => true | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 512 | | Code_Thingol.Classrel _ => true | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 513 | | Code_Thingol.Classinst _ => true | 
| 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 514 | | _ => false) | 
| 25101 | 515 | | is_dict (DFree _) = true | 
| 516 | | is_dict _ = false; | |
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 517 | fun const_of_idx idx = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx | 
| 35371 | 518 | of Code_Thingol.Fun (c, _) => c | 
| 519 | | Code_Thingol.Datatypecons (c, _) => c | |
| 520 | | Code_Thingol.Classparam (c, _) => c); | |
| 24155 | 521 | fun of_apps bounds (t, ts) = | 
| 522 | fold_map (of_univ bounds) ts | |
| 523 | #>> (fn ts' => list_comb (t, rev ts')) | |
| 26064 | 524 | and of_univ bounds (Const (idx, ts)) typidx = | 
| 24155 | 525 | let | 
| 25101 | 526 | val ts' = take_until is_dict ts; | 
| 28663 
bd8438543bf2
code identifier namings are no longer imperative
 haftmann parents: 
28423diff
changeset | 527 | val c = const_of_idx idx; | 
| 31957 | 528 | val T = map_type_tvar (fn ((v, i), _) => | 
| 37145 
01aa36932739
renamed structure TypeInfer to Type_Infer, keeping the old name as legacy alias for some time;
 wenzelm parents: 
36982diff
changeset | 529 | Type_Infer.param typidx (v ^ string_of_int i, [])) | 
| 31957 | 530 | (Sign.the_const_type thy c); | 
| 30022 
1d8b8fa19074
maintain order of constructors in datatypes; clarified conventions for type schemes
 haftmann parents: 
29272diff
changeset | 531 | val typidx' = typidx + 1; | 
| 31957 | 532 | in of_apps bounds (Term.Const (c, T), ts') typidx' end | 
| 27499 | 533 | | of_univ bounds (BVar (n, ts)) typidx = | 
| 534 | of_apps bounds (Bound (bounds - n - 1), ts) typidx | |
| 24155 | 535 | | of_univ bounds (t as Abs _) typidx = | 
| 536 | typidx | |
| 25924 | 537 | |> of_univ (bounds + 1) (apps t [BVar (bounds, [])]) | 
| 24155 | 538 |           |-> (fn t' => pair (Term.Abs ("u", dummyT, t')))
 | 
| 539 | in of_univ 0 t 0 |> fst end; | |
| 540 | ||
| 35371 | 541 | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 542 | (* evaluation with type reconstruction *) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 543 | |
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 544 | fun eval_term thy program (nbe_program, idx_tab) ((vs0, (vs, ty)), t) deps = | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 545 | let | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 546 | val ctxt = Syntax.init_pretty_global thy; | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 547 | val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 548 | val ty' = typ_of_itype program vs0 ty; | 
| 42405 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 wenzelm parents: 
42361diff
changeset | 549 | fun type_infer t = | 
| 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 wenzelm parents: 
42361diff
changeset | 550 | Syntax.check_term (Config.put Type_Infer_Context.const_sorts false ctxt) | 
| 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 wenzelm parents: 
42361diff
changeset | 551 | (Type.constraint ty' t); | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 552 | fun check_tvars t = | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 553 | if null (Term.add_tvars t []) then t | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 554 |       else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t);
 | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 555 | in | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 556 | compile_term thy nbe_program deps (vs, t) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 557 | |> term_of_univ thy program idx_tab | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 558 | |> traced (fn t => "Normalized:\n" ^ string_of_term t) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 559 | |> type_infer | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 560 | |> traced (fn t => "Types inferred:\n" ^ string_of_term t) | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 561 | |> check_tvars | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 562 | |> traced (fn _ => "---\n") | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 563 | end; | 
| 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 564 | |
| 39436 | 565 | |
| 25101 | 566 | (* function store *) | 
| 567 | ||
| 34173 
458ced35abb8
reduced code generator cache to the baremost minimum
 haftmann parents: 
33522diff
changeset | 568 | structure Nbe_Functions = Code_Data | 
| 25101 | 569 | ( | 
| 35500 | 570 | type T = (Univ option * int) Graph.T * (int * string Inttab.table); | 
| 571 | val empty = (Graph.empty, (0, Inttab.empty)); | |
| 25101 | 572 | ); | 
| 573 | ||
| 39399 | 574 | fun compile ignore_cache thy program = | 
| 24155 | 575 | let | 
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 576 | val (nbe_program, (_, idx_tab)) = | 
| 39399 | 577 | Nbe_Functions.change (if ignore_cache then NONE else SOME thy) | 
| 578 | (compile_program thy program); | |
| 39392 
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
 haftmann parents: 
39388diff
changeset | 579 | in (nbe_program, idx_tab) end; | 
| 24155 | 580 | |
| 32544 | 581 | |
| 39396 | 582 | (* dynamic evaluation oracle *) | 
| 24155 | 583 | |
| 30947 | 584 | fun mk_equals thy lhs raw_rhs = | 
| 26747 
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
 haftmann parents: 
26739diff
changeset | 585 | let | 
| 30947 | 586 | val ty = Thm.typ_of (Thm.ctyp_of_term lhs); | 
| 587 |     val eq = Thm.cterm_of thy (Term.Const ("==", ty --> ty --> propT));
 | |
| 588 | val rhs = Thm.cterm_of thy raw_rhs; | |
| 589 | in Thm.mk_binop eq lhs rhs end; | |
| 26747 
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
 haftmann parents: 
26739diff
changeset | 590 | |
| 38669 
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
 haftmann parents: 
37449diff
changeset | 591 | val (_, raw_oracle) = Context.>>> (Context.map_theory_result | 
| 43619 
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
 wenzelm parents: 
43329diff
changeset | 592 |   (Thm.add_oracle (@{binding normalization_by_evaluation},
 | 
| 39399 | 593 | fn (thy, program, nbe_program_idx_tab, vsp_ty_t, deps, ct) => | 
| 594 | mk_equals thy ct (eval_term thy program nbe_program_idx_tab vsp_ty_t deps)))); | |
| 31064 | 595 | |
| 39399 | 596 | fun oracle thy program nbe_program_idx_tab vsp_ty_t deps ct = | 
| 597 | raw_oracle (thy, program, nbe_program_idx_tab, vsp_ty_t, deps, ct); | |
| 30942 
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
 haftmann parents: 
30672diff
changeset | 598 | |
| 41247 | 599 | fun dynamic_conv thy = lift_triv_classes_conv thy (Code_Thingol.dynamic_conv thy | 
| 600 | (K (fn program => oracle thy program (compile false thy program)))); | |
| 24155 | 601 | |
| 41184 | 602 | fun dynamic_value thy = lift_triv_classes_rew thy | 
| 603 | (Code_Thingol.dynamic_value thy I | |
| 39606 
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
 haftmann parents: 
39475diff
changeset | 604 | (K (fn program => eval_term thy program (compile false thy program)))); | 
| 39399 | 605 | |
| 41184 | 606 | fun static_conv thy consts = | 
| 607 | lift_triv_classes_conv thy (Code_Thingol.static_conv thy consts | |
| 41346 | 608 | (K (fn program => fn _ => let val nbe_program = compile true thy program | 
| 609 | in oracle thy program nbe_program end))); | |
| 24839 | 610 | |
| 41184 | 611 | fun static_value thy consts = lift_triv_classes_rew thy | 
| 612 | (Code_Thingol.static_value thy I consts | |
| 41346 | 613 | (K (fn program => fn _ => let val nbe_program = compile true thy program | 
| 614 | in eval_term thy program (compile false thy program) end))); | |
| 41184 | 615 | |
| 35371 | 616 | |
| 39396 | 617 | (** setup **) | 
| 24155 | 618 | |
| 42361 | 619 | val setup = Value.add_evaluator ("nbe", dynamic_value o Proof_Context.theory_of);
 | 
| 24155 | 620 | |
| 621 | end; | |
| 28337 | 622 |