author | huffman |
Wed, 03 Nov 2010 17:22:25 -0700 | |
changeset 40435 | a26503ac7c87 |
parent 39911 | 2b4430847310 |
child 40362 | 82a066bff182 |
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 |
|
38673 | 9 |
val dynamic_eval_conv: conv |
38669
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents:
37449
diff
changeset
|
10 |
val dynamic_eval_value: theory -> term -> term |
39399 | 11 |
val static_eval_conv: theory -> string list -> conv |
25101 | 12 |
|
25204 | 13 |
datatype Univ = |
26064 | 14 |
Const of int * Univ list (*named (uninterpreted) constants*) |
25924 | 15 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
24155 | 16 |
| BVar of int * Univ list |
28350 | 17 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
25924 | 18 |
val apps: Univ -> Univ list -> Univ (*explicit applications*) |
25944 | 19 |
val abss: int -> (Univ list -> Univ) -> Univ |
28337 | 20 |
(*abstractions as closures*) |
28350 | 21 |
val same: Univ -> Univ -> bool |
24155 | 22 |
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
23 |
val put_result: (unit -> Univ list -> Univ list) -> Proof.context -> Proof.context |
32740 | 24 |
val trace: bool Unsynchronized.ref |
25924 | 25 |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
26 |
val setup: theory -> theory |
32544 | 27 |
val add_const_alias: thm -> theory -> theory |
24155 | 28 |
end; |
29 |
||
30 |
structure Nbe: NBE = |
|
31 |
struct |
|
32 |
||
33 |
(* generic non-sense *) |
|
34 |
||
32740 | 35 |
val trace = Unsynchronized.ref false; |
32544 | 36 |
fun traced f x = if !trace then (tracing (f x); x) else x; |
24155 | 37 |
|
38 |
||
32544 | 39 |
(** certificates and oracle for "trivial type classes" **) |
40 |
||
33522 | 41 |
structure Triv_Class_Data = Theory_Data |
32544 | 42 |
( |
43 |
type T = (class * thm) list; |
|
44 |
val empty = []; |
|
45 |
val extend = I; |
|
33522 | 46 |
fun merge data : T = AList.merge (op =) (K true) data; |
32544 | 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 (Code.check_const thy)) 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 |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
80 |
(Thm.add_oracle (Binding.name "triv_of_class", fn (thy, T, class) => |
32544 | 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; |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
88 |
val certT = Thm.ctyp_of thy; |
32544 | 89 |
val triv_classes = get_triv_classes thy; |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
90 |
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:
36960
diff
changeset
|
91 |
fun mk_entry (v, sort) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
92 |
let |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
93 |
val T = TFree (v, sort); |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
94 |
val cT = certT T; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
95 |
val triv_sort = additional_classes sort; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
96 |
in |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
97 |
(v, (Sorts.inter_sort algebra (sort, triv_sort), |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
98 |
(cT, AList.make (fn class => Thm.of_class (cT, class)) sort |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
99 |
@ 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:
36960
diff
changeset
|
100 |
end; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
101 |
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:
36960
diff
changeset
|
102 |
fun instantiate thm = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
103 |
let |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
104 |
val cert_tvars = map (certT o TVar) (Term.add_tvars |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
105 |
((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:
36960
diff
changeset
|
106 |
val instantiation = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
107 |
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:
36960
diff
changeset
|
108 |
in Thm.instantiate (instantiation, []) thm end; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
109 |
fun of_class (TFree (v, _), class) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
110 |
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:
36960
diff
changeset
|
111 |
| of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ_global thy T); |
32544 | 112 |
fun strip_of_class thm = |
113 |
let |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
114 |
val prems_of_class = Thm.prop_of thm |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
115 |
|> Logic.strip_imp_prems |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
116 |
|> map (Logic.dest_of_class #> of_class); |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
117 |
in fold Thm.elim_implies prems_of_class thm end; |
36770
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents:
36768
diff
changeset
|
118 |
in |
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents:
36768
diff
changeset
|
119 |
ct |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
120 |
|> (Drule.cterm_fun o map_types o map_type_tfree) |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
121 |
(fn (v, sort) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v)) |
32544 | 122 |
|> conv |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
123 |
|> 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:
35500
diff
changeset
|
124 |
|> Thm.varifyT_global |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
125 |
|> Thm.unconstrainT |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
126 |
|> instantiate |
32544 | 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 **) |
|
24155 | 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 |
||
25204 | 166 |
datatype Univ = |
26064 | 167 |
Const of int * Univ list (*named (uninterpreted) constants*) |
25924 | 168 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
27499 | 169 |
| BVar of int * Univ list (*bound variables, named*) |
24155 | 170 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
27499 | 171 |
(*abstractions as closures*); |
24155 | 172 |
|
28350 | 173 |
fun same (Const (k, xs)) (Const (l, ys)) = k = l andalso sames xs ys |
174 |
| same (DFree (s, k)) (DFree (t, l)) = s = t andalso k = l |
|
175 |
| same (BVar (k, xs)) (BVar (l, ys)) = k = l andalso sames xs ys |
|
176 |
| same _ _ = false |
|
177 |
and sames xs ys = length xs = length ys andalso forall (uncurry same) (xs ~~ ys); |
|
178 |
||
35371 | 179 |
|
24155 | 180 |
(* constructor functions *) |
181 |
||
25944 | 182 |
fun abss n f = Abs ((n, f), []); |
25924 | 183 |
fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in |
27499 | 184 |
case int_ord (k, 0) |
185 |
of EQUAL => f (ys @ xs) |
|
186 |
| LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end |
|
187 |
| GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*) |
|
188 |
end |
|
25924 | 189 |
| apps (Const (name, xs)) ys = Const (name, ys @ xs) |
27499 | 190 |
| apps (BVar (n, xs)) ys = BVar (n, ys @ xs); |
24155 | 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 ^ ")"; |
|
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 |
||
39399 | 232 |
structure Univs = Proof_Data ( |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
233 |
type T = unit -> Univ list -> Univ list |
39399 | 234 |
fun init _ () = error "Univs" |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
235 |
); |
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
236 |
val put_result = Univs.put; |
24155 | 237 |
|
238 |
local |
|
28350 | 239 |
val prefix = "Nbe."; |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
240 |
val name_put = prefix ^ "put_result"; |
28350 | 241 |
val name_ref = prefix ^ "univs_ref"; |
242 |
val name_const = prefix ^ "Const"; |
|
243 |
val name_abss = prefix ^ "abss"; |
|
244 |
val name_apps = prefix ^ "apps"; |
|
245 |
val name_same = prefix ^ "same"; |
|
24155 | 246 |
in |
247 |
||
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
248 |
val univs_cookie = (Univs.get, put_result, name_put); |
25944 | 249 |
|
28337 | 250 |
fun nbe_fun 0 "" = "nbe_value" |
251 |
| nbe_fun i c = "c_" ^ translate_string (fn "." => "_" | c => c) c ^ "_" ^ string_of_int i; |
|
25101 | 252 |
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; |
24155 | 253 |
fun nbe_bound v = "v_" ^ v; |
31888 | 254 |
fun nbe_bound_optional NONE = "_" |
35500 | 255 |
| nbe_bound_optional (SOME v) = nbe_bound v; |
28337 | 256 |
fun nbe_default v = "w_" ^ v; |
24155 | 257 |
|
25924 | 258 |
(*note: these three are the "turning spots" where proper argument order is established!*) |
259 |
fun nbe_apps t [] = t |
|
260 |
| nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)]; |
|
28337 | 261 |
fun nbe_apps_local i c ts = nbe_fun i c `$` ml_list (rev ts); |
262 |
fun nbe_apps_constr idx_of c ts = |
|
263 |
let |
|
264 |
val c' = if !trace then string_of_int (idx_of c) ^ " (*" ^ c ^ "*)" |
|
265 |
else string_of_int (idx_of c); |
|
266 |
in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end; |
|
25924 | 267 |
|
24155 | 268 |
fun nbe_abss 0 f = f `$` ml_list [] |
25944 | 269 |
| nbe_abss n f = name_abss `$$` [string_of_int n, f]; |
24155 | 270 |
|
28350 | 271 |
fun nbe_same v1 v2 = "(" ^ name_same ^ " " ^ nbe_bound v1 ^ " " ^ nbe_bound v2 ^ ")"; |
272 |
||
24155 | 273 |
end; |
274 |
||
28054 | 275 |
open Basic_Code_Thingol; |
24155 | 276 |
|
35371 | 277 |
|
25865 | 278 |
(* code generation *) |
24155 | 279 |
|
26064 | 280 |
fun assemble_eqnss idx_of deps eqnss = |
25944 | 281 |
let |
282 |
fun prep_eqns (c, (vs, eqns)) = |
|
25924 | 283 |
let |
25944 | 284 |
val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; |
28337 | 285 |
val num_args = length dicts + ((length o fst o hd) eqns); |
25944 | 286 |
in (c, (num_args, (dicts, eqns))) end; |
287 |
val eqnss' = map prep_eqns eqnss; |
|
288 |
||
289 |
fun assemble_constapp c dss ts = |
|
290 |
let |
|
291 |
val ts' = (maps o map) assemble_idict dss @ ts; |
|
292 |
in case AList.lookup (op =) eqnss' c |
|
28337 | 293 |
of SOME (num_args, _) => if num_args <= length ts' |
294 |
then let val (ts1, ts2) = chop num_args ts' |
|
295 |
in nbe_apps (nbe_apps_local 0 c ts1) ts2 |
|
296 |
end else nbe_apps (nbe_abss num_args (nbe_fun 0 c)) ts' |
|
25944 | 297 |
| NONE => if member (op =) deps c |
28337 | 298 |
then nbe_apps (nbe_fun 0 c) ts' |
299 |
else nbe_apps_constr idx_of c ts' |
|
25924 | 300 |
end |
25944 | 301 |
and assemble_idict (DictConst (inst, dss)) = |
302 |
assemble_constapp inst dss [] |
|
303 |
| assemble_idict (DictVar (supers, (v, (n, _)))) = |
|
304 |
fold_rev (fn super => assemble_constapp super [] o single) supers (nbe_dict v n); |
|
25924 | 305 |
|
28350 | 306 |
fun assemble_iterm constapp = |
24155 | 307 |
let |
28350 | 308 |
fun of_iterm match_cont t = |
25944 | 309 |
let |
28054 | 310 |
val (t', ts) = Code_Thingol.unfold_app t |
28350 | 311 |
in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end |
31049 | 312 |
and of_iapp match_cont (IConst (c, ((_, dss), _))) ts = constapp c dss ts |
31889 | 313 |
| of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts |
31724 | 314 |
| of_iapp match_cont ((v, _) `|=> t) ts = |
31888 | 315 |
nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts |
28350 | 316 |
| of_iapp match_cont (ICase (((t, _), cs), t0)) ts = |
317 |
nbe_apps (ml_cases (of_iterm NONE t) |
|
318 |
(map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) cs |
|
319 |
@ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts |
|
25944 | 320 |
in of_iterm end; |
24155 | 321 |
|
28350 | 322 |
fun subst_nonlin_vars args = |
323 |
let |
|
324 |
val vs = (fold o Code_Thingol.fold_varnames) |
|
33002 | 325 |
(fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args []; |
28350 | 326 |
val names = Name.make_context (map fst vs); |
327 |
fun declare v k ctxt = let val vs = Name.invents ctxt v k |
|
328 |
in (vs, fold Name.declare vs ctxt) end; |
|
329 |
val (vs_renames, _) = fold_map (fn (v, k) => if k > 1 |
|
330 |
then declare v (k - 1) #>> (fn vs => (v, vs)) |
|
331 |
else pair (v, [])) vs names; |
|
332 |
val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames; |
|
333 |
fun subst_vars (t as IConst _) samepairs = (t, samepairs) |
|
31889 | 334 |
| subst_vars (t as IVar NONE) samepairs = (t, samepairs) |
335 |
| subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v |
|
336 |
of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs) |
|
28350 | 337 |
| NONE => (t, samepairs)) |
338 |
| subst_vars (t1 `$ t2) samepairs = samepairs |
|
339 |
|> subst_vars t1 |
|
340 |
||>> subst_vars t2 |
|
341 |
|>> (op `$) |
|
342 |
| subst_vars (ICase (_, t)) samepairs = subst_vars t samepairs; |
|
343 |
val (args', _) = fold_map subst_vars args samepairs; |
|
344 |
in (samepairs, args') end; |
|
345 |
||
28337 | 346 |
fun assemble_eqn c dicts default_args (i, (args, rhs)) = |
347 |
let |
|
348 |
val is_eval = (c = ""); |
|
349 |
val default_rhs = nbe_apps_local (i+1) c (dicts @ default_args); |
|
350 |
val match_cont = if is_eval then NONE else SOME default_rhs; |
|
28350 | 351 |
val assemble_arg = assemble_iterm |
352 |
(fn c => fn _ => fn ts => nbe_apps_constr idx_of c ts) NONE; |
|
35371 | 353 |
val assemble_rhs = assemble_iterm assemble_constapp match_cont; |
28350 | 354 |
val (samepairs, args') = subst_nonlin_vars args; |
355 |
val s_args = map assemble_arg args'; |
|
356 |
val s_rhs = if null samepairs then assemble_rhs rhs |
|
357 |
else ml_if (ml_and (map (uncurry nbe_same) samepairs)) |
|
358 |
(assemble_rhs rhs) default_rhs; |
|
28337 | 359 |
val eqns = if is_eval then |
28350 | 360 |
[([ml_list (rev (dicts @ s_args))], s_rhs)] |
28337 | 361 |
else |
28350 | 362 |
[([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs), |
28337 | 363 |
([ml_list (rev (dicts @ default_args))], default_rhs)] |
364 |
in (nbe_fun i c, eqns) end; |
|
365 |
||
25944 | 366 |
fun assemble_eqns (c, (num_args, (dicts, eqns))) = |
367 |
let |
|
28337 | 368 |
val default_args = map nbe_default |
369 |
(Name.invent_list [] "a" (num_args - length dicts)); |
|
370 |
val eqns' = map_index (assemble_eqn c dicts default_args) eqns |
|
371 |
@ (if c = "" then [] else [(nbe_fun (length eqns) c, |
|
372 |
[([ml_list (rev (dicts @ default_args))], |
|
373 |
nbe_apps_constr idx_of c (dicts @ default_args))])]); |
|
374 |
in (eqns', nbe_abss num_args (nbe_fun 0 c)) end; |
|
24155 | 375 |
|
25944 | 376 |
val (fun_vars, fun_vals) = map_split assemble_eqns eqnss'; |
28337 | 377 |
val deps_vars = ml_list (map (nbe_fun 0) deps); |
378 |
in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end; |
|
25944 | 379 |
|
35371 | 380 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
381 |
(* compile equations *) |
25944 | 382 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
383 |
fun compile_eqnss thy nbe_program raw_deps [] = [] |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
384 |
| compile_eqnss thy nbe_program raw_deps eqnss = |
24155 | 385 |
let |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
386 |
val ctxt = ProofContext.init_global thy; |
26064 | 387 |
val (deps, deps_vals) = split_list (map_filter |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
388 |
(fn dep => Option.map (fn univ => (dep, univ)) (fst ((Graph.get_node nbe_program dep)))) raw_deps); |
26064 | 389 |
val idx_of = raw_deps |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
390 |
|> map (fn dep => (dep, snd (Graph.get_node nbe_program dep))) |
26064 | 391 |
|> AList.lookup (op =) |
392 |
|> (fn f => the o f); |
|
393 |
val s = assemble_eqnss idx_of deps eqnss; |
|
24155 | 394 |
val cs = map fst eqnss; |
25944 | 395 |
in |
396 |
s |
|
32544 | 397 |
|> traced (fn s => "\n--- code to be evaluated:\n" ^ s) |
38931
5e84c11c4b8a
evaluate takes ml context and ml expression parameter
haftmann
parents:
38676
diff
changeset
|
398 |
|> pair "" |
39911 | 399 |
|> Code_Runtime.value ctxt univs_cookie |
25944 | 400 |
|> (fn f => f deps_vals) |
401 |
|> (fn univs => cs ~~ univs) |
|
402 |
end; |
|
25190 | 403 |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30288
diff
changeset
|
404 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
405 |
(* extract equations from statements *) |
24155 | 406 |
|
37437 | 407 |
fun eqns_of_stmt (_, Code_Thingol.Fun (_, ((_, []), _))) = |
25101 | 408 |
[] |
37437 | 409 |
| eqns_of_stmt (const, Code_Thingol.Fun (_, (((vs, _), eqns), _))) = |
25101 | 410 |
[(const, (vs, map fst eqns))] |
28054 | 411 |
| eqns_of_stmt (_, Code_Thingol.Datatypecons _) = |
25101 | 412 |
[] |
28054 | 413 |
| eqns_of_stmt (_, Code_Thingol.Datatype _) = |
25101 | 414 |
[] |
37447
ad3e04f289b6
transitive superclasses were also only a misunderstanding
haftmann
parents:
37446
diff
changeset
|
415 |
| eqns_of_stmt (class, Code_Thingol.Class (_, (v, (super_classes, classparams)))) = |
25101 | 416 |
let |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37146
diff
changeset
|
417 |
val names = map snd super_classes @ map fst classparams; |
25101 | 418 |
val params = Name.invent_list [] "d" (length names); |
419 |
fun mk (k, name) = |
|
420 |
(name, ([(v, [])], |
|
31889 | 421 |
[([IConst (class, (([], []), [])) `$$ map (IVar o SOME) params], |
422 |
IVar (SOME (nth params k)))])); |
|
25101 | 423 |
in map_index mk names end |
28054 | 424 |
| eqns_of_stmt (_, Code_Thingol.Classrel _) = |
25101 | 425 |
[] |
28054 | 426 |
| eqns_of_stmt (_, Code_Thingol.Classparam _) = |
25101 | 427 |
[] |
37449 | 428 |
| eqns_of_stmt (inst, Code_Thingol.Classinst ((class, (_, arity_args)), (super_instances, (classparam_instances, _)))) = |
37384
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37146
diff
changeset
|
429 |
[(inst, (arity_args, [([], IConst (class, (([], []), [])) `$$ |
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37146
diff
changeset
|
430 |
map (fn (_, (_, (inst, dss))) => IConst (inst, (([], dss), []))) super_instances |
5aba26803073
more consistent naming aroud type classes and instances
haftmann
parents:
37146
diff
changeset
|
431 |
@ map (IConst o snd o fst) classparam_instances)]))]; |
24155 | 432 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
433 |
|
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
434 |
(* compile whole programs *) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
435 |
|
39399 | 436 |
fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) = |
437 |
if can (Graph.get_node nbe_program) name |
|
438 |
then (nbe_program, (maxidx, idx_tab)) |
|
439 |
else (Graph.new_node (name, (NONE, maxidx)) nbe_program, |
|
440 |
(maxidx + 1, Inttab.update_new (maxidx, name) idx_tab)); |
|
441 |
||
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
442 |
fun compile_stmts thy stmts_deps = |
25101 | 443 |
let |
444 |
val names = map (fst o fst) stmts_deps; |
|
445 |
val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; |
|
446 |
val eqnss = maps (eqns_of_stmt o fst) stmts_deps; |
|
26064 | 447 |
val refl_deps = names_deps |
25190 | 448 |
|> maps snd |
449 |
|> distinct (op =) |
|
26064 | 450 |
|> fold (insert (op =)) names; |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
451 |
fun compile nbe_program = eqnss |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
452 |
|> compile_eqnss thy nbe_program refl_deps |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
453 |
|> rpair nbe_program; |
25101 | 454 |
in |
39399 | 455 |
fold ensure_const_idx refl_deps |
26064 | 456 |
#> apfst (fold (fn (name, deps) => fold (curry Graph.add_edge name) deps) names_deps |
457 |
#> compile |
|
458 |
#-> fold (fn (name, univ) => (Graph.map_node name o apfst) (K (SOME univ)))) |
|
25101 | 459 |
end; |
24155 | 460 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
461 |
fun compile_program thy program = |
25101 | 462 |
let |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
463 |
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:
39388
diff
changeset
|
464 |
then (nbe_program, (maxidx, idx_tab)) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
465 |
else (nbe_program, (maxidx, idx_tab)) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
466 |
|> compile_stmts thy (map (fn name => ((name, Graph.get_node program name), |
27103 | 467 |
Graph.imm_succs program name)) names); |
28706 | 468 |
in |
469 |
fold_rev add_stmts (Graph.strong_conn program) |
|
470 |
end; |
|
24155 | 471 |
|
25944 | 472 |
|
473 |
(** evaluation **) |
|
474 |
||
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
475 |
(* term evaluation by compilation *) |
25944 | 476 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
477 |
fun compile_term thy nbe_program deps (vs : (string * sort) list, t) = |
25924 | 478 |
let |
479 |
val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; |
|
480 |
in |
|
31064 | 481 |
("", (vs, [([], t)])) |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
482 |
|> singleton (compile_eqnss thy nbe_program deps) |
25924 | 483 |
|> snd |
31064 | 484 |
|> (fn t => apps t (rev dict_frees)) |
25924 | 485 |
end; |
24155 | 486 |
|
35371 | 487 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
488 |
(* reconstruction *) |
24155 | 489 |
|
30947 | 490 |
fun typ_of_itype program vs (ityco `%% itys) = |
491 |
let |
|
492 |
val Code_Thingol.Datatype (tyco, _) = Graph.get_node program ityco; |
|
493 |
in Type (tyco, map (typ_of_itype program vs) itys) end |
|
494 |
| typ_of_itype program vs (ITyVar v) = |
|
495 |
let |
|
496 |
val sort = (the o AList.lookup (op =) vs) v; |
|
497 |
in TFree ("'" ^ v, sort) end; |
|
498 |
||
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
499 |
fun term_of_univ thy program idx_tab t = |
24155 | 500 |
let |
25101 | 501 |
fun take_until f [] = [] |
502 |
| 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:
28423
diff
changeset
|
503 |
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:
28423
diff
changeset
|
504 |
of Code_Thingol.Class _ => true |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
505 |
| Code_Thingol.Classrel _ => true |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
506 |
| Code_Thingol.Classinst _ => true |
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
507 |
| _ => false) |
25101 | 508 |
| is_dict (DFree _) = true |
509 |
| is_dict _ = false; |
|
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
510 |
fun const_of_idx idx = (case (Graph.get_node program o the o Inttab.lookup idx_tab) idx |
35371 | 511 |
of Code_Thingol.Fun (c, _) => c |
512 |
| Code_Thingol.Datatypecons (c, _) => c |
|
513 |
| Code_Thingol.Classparam (c, _) => c); |
|
24155 | 514 |
fun of_apps bounds (t, ts) = |
515 |
fold_map (of_univ bounds) ts |
|
516 |
#>> (fn ts' => list_comb (t, rev ts')) |
|
26064 | 517 |
and of_univ bounds (Const (idx, ts)) typidx = |
24155 | 518 |
let |
25101 | 519 |
val ts' = take_until is_dict ts; |
28663
bd8438543bf2
code identifier namings are no longer imperative
haftmann
parents:
28423
diff
changeset
|
520 |
val c = const_of_idx idx; |
31957 | 521 |
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:
36982
diff
changeset
|
522 |
Type_Infer.param typidx (v ^ string_of_int i, [])) |
31957 | 523 |
(Sign.the_const_type thy c); |
30022
1d8b8fa19074
maintain order of constructors in datatypes; clarified conventions for type schemes
haftmann
parents:
29272
diff
changeset
|
524 |
val typidx' = typidx + 1; |
31957 | 525 |
in of_apps bounds (Term.Const (c, T), ts') typidx' end |
27499 | 526 |
| of_univ bounds (BVar (n, ts)) typidx = |
527 |
of_apps bounds (Bound (bounds - n - 1), ts) typidx |
|
24155 | 528 |
| of_univ bounds (t as Abs _) typidx = |
529 |
typidx |
|
25924 | 530 |
|> of_univ (bounds + 1) (apps t [BVar (bounds, [])]) |
24155 | 531 |
|-> (fn t' => pair (Term.Abs ("u", dummyT, t'))) |
532 |
in of_univ 0 t 0 |> fst end; |
|
533 |
||
35371 | 534 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
535 |
(* evaluation with type reconstruction *) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
536 |
|
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
537 |
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:
39388
diff
changeset
|
538 |
let |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
539 |
val ctxt = Syntax.init_pretty_global thy; |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
540 |
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:
39388
diff
changeset
|
541 |
val ty' = typ_of_itype program vs0 ty; |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
542 |
fun type_infer t = singleton |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
543 |
(Type_Infer.infer_types ctxt (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE)) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
544 |
(Type.constraint ty' t); |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
545 |
fun check_tvars t = |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
546 |
if null (Term.add_tvars t []) then t |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
547 |
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:
39388
diff
changeset
|
548 |
in |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
549 |
compile_term thy nbe_program deps (vs, t) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
550 |
|> term_of_univ thy program idx_tab |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
551 |
|> traced (fn t => "Normalized:\n" ^ string_of_term t) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
552 |
|> type_infer |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
553 |
|> traced (fn t => "Types inferred:\n" ^ string_of_term t) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
554 |
|> check_tvars |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
555 |
|> traced (fn _ => "---\n") |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
556 |
end; |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
557 |
|
39436 | 558 |
|
25101 | 559 |
(* function store *) |
560 |
||
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33522
diff
changeset
|
561 |
structure Nbe_Functions = Code_Data |
25101 | 562 |
( |
35500 | 563 |
type T = (Univ option * int) Graph.T * (int * string Inttab.table); |
564 |
val empty = (Graph.empty, (0, Inttab.empty)); |
|
25101 | 565 |
); |
566 |
||
39399 | 567 |
fun compile ignore_cache thy program = |
24155 | 568 |
let |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
569 |
val (nbe_program, (_, idx_tab)) = |
39399 | 570 |
Nbe_Functions.change (if ignore_cache then NONE else SOME thy) |
571 |
(compile_program thy program); |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
572 |
in (nbe_program, idx_tab) end; |
24155 | 573 |
|
32544 | 574 |
|
39396 | 575 |
(* dynamic evaluation oracle *) |
24155 | 576 |
|
30947 | 577 |
fun mk_equals thy lhs raw_rhs = |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
578 |
let |
30947 | 579 |
val ty = Thm.typ_of (Thm.ctyp_of_term lhs); |
580 |
val eq = Thm.cterm_of thy (Term.Const ("==", ty --> ty --> propT)); |
|
581 |
val rhs = Thm.cterm_of thy raw_rhs; |
|
582 |
in Thm.mk_binop eq lhs rhs end; |
|
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
583 |
|
38669
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents:
37449
diff
changeset
|
584 |
val (_, raw_oracle) = Context.>>> (Context.map_theory_result |
39399 | 585 |
(Thm.add_oracle (Binding.name "normalization_by_evaluation", |
586 |
fn (thy, program, nbe_program_idx_tab, vsp_ty_t, deps, ct) => |
|
587 |
mk_equals thy ct (eval_term thy program nbe_program_idx_tab vsp_ty_t deps)))); |
|
31064 | 588 |
|
39399 | 589 |
fun oracle thy program nbe_program_idx_tab vsp_ty_t deps ct = |
590 |
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:
30672
diff
changeset
|
591 |
|
39606
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
592 |
val dynamic_eval_conv = Conv.tap_thy (fn thy => |
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
593 |
lift_triv_classes_conv thy (Code_Thingol.dynamic_eval_conv thy |
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
594 |
(K (fn program => oracle thy program (compile false thy program))))); |
24155 | 595 |
|
38669
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents:
37449
diff
changeset
|
596 |
fun dynamic_eval_value thy = lift_triv_classes_rew thy |
39606
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
597 |
(Code_Thingol.dynamic_eval_value thy I |
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
598 |
(K (fn program => eval_term thy program (compile false thy program)))); |
39399 | 599 |
|
39606
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
600 |
fun static_eval_conv thy consts = |
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
601 |
lift_triv_classes_conv thy (Code_Thingol.static_eval_conv thy consts |
39475 | 602 |
(K (fn program => let val nbe_program = compile true thy program |
39606
7af0441a3a47
no_frees_* is subsumed by new framework mechanisms in Code_Preproc
haftmann
parents:
39475
diff
changeset
|
603 |
in fn thy => oracle thy program nbe_program end))); |
24839 | 604 |
|
35371 | 605 |
|
39396 | 606 |
(** setup **) |
24155 | 607 |
|
38669
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents:
37449
diff
changeset
|
608 |
val setup = Value.add_evaluator ("nbe", dynamic_eval_value o ProofContext.theory_of); |
24155 | 609 |
|
610 |
end; |
|
28337 | 611 |