author | paulson <lp15@cam.ac.uk> |
Tue, 13 Sep 2022 18:56:39 +0100 | |
changeset 76139 | 3190ee65139b |
parent 74561 | 8e6c973003c8 |
child 77233 | 6bdd125d932b |
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 |
|
55757 | 9 |
val dynamic_conv: Proof.context -> conv |
10 |
val dynamic_value: Proof.context -> term -> term |
|
63164 | 11 |
val static_conv: { ctxt: Proof.context, consts: string list } |
12 |
-> Proof.context -> conv |
|
13 |
val static_value: { ctxt: Proof.context, consts: string list } |
|
14 |
-> Proof.context -> term -> term |
|
25101 | 15 |
|
25204 | 16 |
datatype Univ = |
26064 | 17 |
Const of int * Univ list (*named (uninterpreted) constants*) |
25924 | 18 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
24155 | 19 |
| BVar of int * Univ list |
28350 | 20 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
25924 | 21 |
val apps: Univ -> Univ list -> Univ (*explicit applications*) |
25944 | 22 |
val abss: int -> (Univ list -> Univ) -> Univ |
28337 | 23 |
(*abstractions as closures*) |
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
24 |
val same: Univ * Univ -> bool |
24155 | 25 |
|
63164 | 26 |
val put_result: (unit -> Univ list -> Univ list) |
27 |
-> Proof.context -> Proof.context |
|
57435 | 28 |
val trace: bool Config.T |
25924 | 29 |
|
32544 | 30 |
val add_const_alias: thm -> theory -> theory |
24155 | 31 |
end; |
32 |
||
33 |
structure Nbe: NBE = |
|
34 |
struct |
|
35 |
||
36 |
(* generic non-sense *) |
|
37 |
||
67149 | 38 |
val trace = Attrib.setup_config_bool \<^binding>\<open>nbe_trace\<close> (K false); |
57435 | 39 |
fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x; |
24155 | 40 |
|
41 |
||
32544 | 42 |
(** certificates and oracle for "trivial type classes" **) |
43 |
||
33522 | 44 |
structure Triv_Class_Data = Theory_Data |
32544 | 45 |
( |
46 |
type T = (class * thm) list; |
|
47 |
val empty = []; |
|
33522 | 48 |
fun merge data : T = AList.merge (op =) (K true) data; |
32544 | 49 |
); |
50 |
||
51 |
fun add_const_alias thm thy = |
|
52 |
let |
|
53 |
val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm) |
|
54 |
of SOME ofclass_eq => ofclass_eq |
|
61268 | 55 |
| _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); |
32544 | 56 |
val (T, class) = case try Logic.dest_of_class ofclass |
57 |
of SOME T_class => T_class |
|
61268 | 58 |
| _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); |
32544 | 59 |
val tvar = case try Term.dest_TVar T |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
48072
diff
changeset
|
60 |
of SOME (tvar as (_, sort)) => if null (filter (can (Axclass.get_info thy)) sort) |
32544 | 61 |
then tvar |
61268 | 62 |
else error ("Bad sort: " ^ Thm.string_of_thm_global thy thm) |
63 |
| _ => error ("Bad type: " ^ Thm.string_of_thm_global thy thm); |
|
32544 | 64 |
val _ = if Term.add_tvars eqn [] = [tvar] then () |
61268 | 65 |
else error ("Inconsistent type: " ^ Thm.string_of_thm_global thy thm); |
32544 | 66 |
val lhs_rhs = case try Logic.dest_equals eqn |
67 |
of SOME lhs_rhs => lhs_rhs |
|
68 |
| _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn); |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58397
diff
changeset
|
69 |
val c_c' = case try (apply2 (Axclass.unoverload_const thy o dest_Const)) lhs_rhs |
32544 | 70 |
of SOME c_c' => c_c' |
71 |
| _ => error ("Not an equation with two constants: " |
|
72 |
^ Syntax.string_of_term_global thy eqn); |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
48072
diff
changeset
|
73 |
val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then () |
61268 | 74 |
else error ("Inconsistent class: " ^ Thm.string_of_thm_global thy thm); |
61096 | 75 |
in Triv_Class_Data.map (AList.update (op =) (class, Thm.trim_context thm)) thy end; |
32544 | 76 |
|
77 |
local |
|
78 |
||
79 |
val get_triv_classes = map fst o Triv_Class_Data.get; |
|
80 |
||
81 |
val (_, triv_of_class) = Context.>>> (Context.map_theory_result |
|
67149 | 82 |
(Thm.add_oracle (\<^binding>\<open>triv_of_class\<close>, fn (thy, T, class) => |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59586
diff
changeset
|
83 |
Thm.global_cterm_of thy (Logic.mk_of_class (T, class))))); |
32544 | 84 |
|
85 |
in |
|
86 |
||
63158
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
87 |
fun lift_triv_classes_conv orig_ctxt conv ct = |
32544 | 88 |
let |
63158
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
89 |
val thy = Proof_Context.theory_of orig_ctxt; |
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
90 |
val ctxt = Proof_Context.init_global thy; |
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
91 |
(*FIXME quasi-global context*) |
32544 | 92 |
val algebra = Sign.classes_of thy; |
93 |
val triv_classes = get_triv_classes thy; |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
94 |
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
|
95 |
fun mk_entry (v, sort) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
96 |
let |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
97 |
val T = TFree (v, sort); |
60322 | 98 |
val cT = Thm.ctyp_of ctxt T; |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
99 |
val triv_sort = additional_classes sort; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
100 |
in |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
101 |
(v, (Sorts.inter_sort algebra (sort, triv_sort), |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
102 |
(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
|
103 |
@ 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
|
104 |
end; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
105 |
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
|
106 |
fun instantiate thm = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
107 |
let |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60322
diff
changeset
|
108 |
val tvars = |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60322
diff
changeset
|
109 |
Term.add_tvars (#1 (Logic.dest_equals (Logic.strip_imp_concl (Thm.prop_of thm)))) []; |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
60322
diff
changeset
|
110 |
val instT = map2 (fn v => fn (_, (_, (cT, _))) => (v, cT)) tvars vs_tab; |
74282 | 111 |
in Thm.instantiate (TVars.make instT, Vars.empty) thm end; |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
112 |
fun of_class (TFree (v, _), class) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
113 |
the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class) |
60322 | 114 |
| of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ ctxt T); |
32544 | 115 |
fun strip_of_class thm = |
116 |
let |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
117 |
val prems_of_class = Thm.prop_of thm |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
118 |
|> Logic.strip_imp_prems |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
119 |
|> map (Logic.dest_of_class #> of_class); |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
120 |
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
|
121 |
in |
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents:
36768
diff
changeset
|
122 |
ct |
60322 | 123 |
|> Thm.term_of |
124 |
|> (map_types o map_type_tfree) |
|
47609 | 125 |
(fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v)) |
60322 | 126 |
|> Thm.cterm_of ctxt |
63158
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
127 |
|> conv ctxt |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
128 |
|> 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
|
129 |
|> Thm.varifyT_global |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
130 |
|> Thm.unconstrainT |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
131 |
|> instantiate |
32544 | 132 |
|> strip_of_class |
133 |
end; |
|
134 |
||
55757 | 135 |
fun lift_triv_classes_rew ctxt rew t = |
32544 | 136 |
let |
55757 | 137 |
val thy = Proof_Context.theory_of ctxt; |
32544 | 138 |
val algebra = Sign.classes_of thy; |
139 |
val triv_classes = get_triv_classes thy; |
|
140 |
val vs = Term.add_tfrees t []; |
|
52473
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
141 |
in |
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
142 |
t |
32544 | 143 |
|> (map_types o map_type_tfree) |
144 |
(fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes))) |
|
145 |
|> rew |
|
146 |
|> (map_types o map_type_tfree) |
|
52473
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
147 |
(fn (v, sort) => TFree (v, the_default sort (AList.lookup (op =) vs v))) |
32544 | 148 |
end; |
149 |
||
150 |
end; |
|
151 |
||
152 |
||
153 |
(** the semantic universe **) |
|
24155 | 154 |
|
155 |
(* |
|
156 |
Functions are given by their semantical function value. To avoid |
|
157 |
trouble with the ML-type system, these functions have the most |
|
158 |
generic type, that is "Univ list -> Univ". The calling convention is |
|
159 |
that the arguments come as a list, the last argument first. In |
|
160 |
other words, a function call that usually would look like |
|
161 |
||
162 |
f x_1 x_2 ... x_n or f(x_1,x_2, ..., x_n) |
|
163 |
||
164 |
would be in our convention called as |
|
165 |
||
166 |
f [x_n,..,x_2,x_1] |
|
167 |
||
168 |
Moreover, to handle functions that are still waiting for some |
|
169 |
arguments we have additionally a list of arguments collected to far |
|
170 |
and the number of arguments we're still waiting for. |
|
171 |
*) |
|
172 |
||
25204 | 173 |
datatype Univ = |
26064 | 174 |
Const of int * Univ list (*named (uninterpreted) constants*) |
25924 | 175 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
27499 | 176 |
| BVar of int * Univ list (*bound variables, named*) |
24155 | 177 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
27499 | 178 |
(*abstractions as closures*); |
24155 | 179 |
|
35371 | 180 |
|
24155 | 181 |
(* constructor functions *) |
182 |
||
25944 | 183 |
fun abss n f = Abs ((n, f), []); |
25924 | 184 |
fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in |
27499 | 185 |
case int_ord (k, 0) |
186 |
of EQUAL => f (ys @ xs) |
|
187 |
| LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end |
|
188 |
| GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*) |
|
189 |
end |
|
25924 | 190 |
| apps (Const (name, xs)) ys = Const (name, ys @ xs) |
27499 | 191 |
| apps (BVar (n, xs)) ys = BVar (n, ys @ xs); |
24155 | 192 |
|
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
193 |
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:
40730
diff
changeset
|
194 |
| 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:
40730
diff
changeset
|
195 |
| 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:
40730
diff
changeset
|
196 |
| same _ = false; |
40730
2aa0390a2da7
nbe decides equality of abstractions by extensionality
haftmann
parents:
40564
diff
changeset
|
197 |
|
24155 | 198 |
|
199 |
(** assembling and compiling ML code from terms **) |
|
200 |
||
201 |
(* abstract ML syntax *) |
|
202 |
||
203 |
infix 9 `$` `$$`; |
|
204 |
fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")"; |
|
25101 | 205 |
fun e `$$` [] = e |
206 |
| e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")"; |
|
24590 | 207 |
fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")"; |
24155 | 208 |
|
209 |
fun ml_cases t cs = |
|
210 |
"(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")"; |
|
25944 | 211 |
fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end"; |
28337 | 212 |
fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")"; |
24155 | 213 |
|
28350 | 214 |
fun ml_and [] = "true" |
215 |
| ml_and [x] = x |
|
216 |
| ml_and xs = "(" ^ space_implode " andalso " xs ^ ")"; |
|
217 |
fun ml_if b x y = "(if " ^ b ^ " then " ^ x ^ " else " ^ y ^ ")"; |
|
218 |
||
24155 | 219 |
fun ml_list es = "[" ^ commas es ^ "]"; |
220 |
||
221 |
fun ml_fundefs ([(name, [([], e)])]) = |
|
222 |
"val " ^ name ^ " = " ^ e ^ "\n" |
|
223 |
| ml_fundefs (eqs :: eqss) = |
|
224 |
let |
|
225 |
fun fundef (name, eqs) = |
|
226 |
let |
|
227 |
fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e |
|
228 |
in space_implode "\n | " (map eqn eqs) end; |
|
229 |
in |
|
230 |
(prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss |
|
26931 | 231 |
|> cat_lines |
24155 | 232 |
|> suffix "\n" |
233 |
end; |
|
234 |
||
35371 | 235 |
|
25944 | 236 |
(* nbe specific syntax and sandbox communication *) |
237 |
||
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41346
diff
changeset
|
238 |
structure Univs = Proof_Data |
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41346
diff
changeset
|
239 |
( |
59153 | 240 |
type T = unit -> Univ list -> Univ list; |
241 |
val empty: T = fn () => raise Fail "Univs"; |
|
242 |
fun init _ = empty; |
|
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
243 |
); |
59153 | 244 |
val get_result = Univs.get; |
39388
fdbb2c55ffc2
replaced ML_Context.evaluate by ML_Context.value -- using context data instead of bare metal references
haftmann
parents:
39290
diff
changeset
|
245 |
val put_result = Univs.put; |
24155 | 246 |
|
247 |
local |
|
59151
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
248 |
val prefix = "Nbe."; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
249 |
val name_put = prefix ^ "put_result"; |
41068 | 250 |
val name_const = prefix ^ "Const"; |
59151
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
251 |
val name_abss = prefix ^ "abss"; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
252 |
val name_apps = prefix ^ "apps"; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
253 |
val name_same = prefix ^ "same"; |
24155 | 254 |
in |
255 |
||
59153 | 256 |
val univs_cookie = (get_result, put_result, name_put); |
25944 | 257 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
258 |
fun nbe_fun idx_of 0 (Code_Symbol.Constant "") = "nbe_value" |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
259 |
| nbe_fun idx_of i sym = "c_" ^ string_of_int (idx_of sym) |
55150 | 260 |
^ "_" ^ Code_Symbol.default_base sym ^ "_" ^ string_of_int i; |
25101 | 261 |
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; |
24155 | 262 |
fun nbe_bound v = "v_" ^ v; |
31888 | 263 |
fun nbe_bound_optional NONE = "_" |
35500 | 264 |
| nbe_bound_optional (SOME v) = nbe_bound v; |
28337 | 265 |
fun nbe_default v = "w_" ^ v; |
24155 | 266 |
|
25924 | 267 |
(*note: these three are the "turning spots" where proper argument order is established!*) |
268 |
fun nbe_apps t [] = t |
|
269 |
| nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)]; |
|
55043
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
270 |
fun nbe_apps_local idx_of i c ts = nbe_fun idx_of i c `$` ml_list (rev ts); |
57435 | 271 |
fun nbe_apps_constr ctxt idx_of c ts = |
28337 | 272 |
let |
57435 | 273 |
val c' = if Config.get ctxt trace |
274 |
then string_of_int (idx_of c) ^ " (*" ^ Code_Symbol.default_base c ^ "*)" |
|
28337 | 275 |
else string_of_int (idx_of c); |
276 |
in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end; |
|
25924 | 277 |
|
24155 | 278 |
fun nbe_abss 0 f = f `$` ml_list [] |
25944 | 279 |
| nbe_abss n f = name_abss `$$` [string_of_int n, f]; |
24155 | 280 |
|
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
281 |
fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))"; |
28350 | 282 |
|
24155 | 283 |
end; |
284 |
||
55150 | 285 |
open Basic_Code_Symbol; |
28054 | 286 |
open Basic_Code_Thingol; |
24155 | 287 |
|
35371 | 288 |
|
25865 | 289 |
(* code generation *) |
24155 | 290 |
|
57435 | 291 |
fun assemble_eqnss ctxt idx_of deps eqnss = |
25944 | 292 |
let |
293 |
fun prep_eqns (c, (vs, eqns)) = |
|
25924 | 294 |
let |
25944 | 295 |
val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; |
28337 | 296 |
val num_args = length dicts + ((length o fst o hd) eqns); |
25944 | 297 |
in (c, (num_args, (dicts, eqns))) end; |
298 |
val eqnss' = map prep_eqns eqnss; |
|
299 |
||
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
300 |
fun assemble_constapp sym dss ts = |
25944 | 301 |
let |
41118
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
302 |
val ts' = (maps o map) assemble_dict dss @ ts; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
303 |
in case AList.lookup (op =) eqnss' sym |
28337 | 304 |
of SOME (num_args, _) => if num_args <= length ts' |
305 |
then let val (ts1, ts2) = chop num_args ts' |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
306 |
in nbe_apps (nbe_apps_local idx_of 0 sym ts1) ts2 |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
307 |
end else nbe_apps (nbe_abss num_args (nbe_fun idx_of 0 sym)) ts' |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
308 |
| NONE => if member (op =) deps sym |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
309 |
then nbe_apps (nbe_fun idx_of 0 sym) ts' |
57435 | 310 |
else nbe_apps_constr ctxt idx_of sym ts' |
25924 | 311 |
end |
41100
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
haftmann
parents:
41068
diff
changeset
|
312 |
and assemble_classrels classrels = |
55150 | 313 |
fold_rev (fn classrel => assemble_constapp (Class_Relation classrel) [] o single) classrels |
41118
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
314 |
and assemble_dict (Dict (classrels, x)) = |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
315 |
assemble_classrels classrels (assemble_plain_dict x) |
b290841cd3b0
separated dictionary weakning into separate type
haftmann
parents:
41100
diff
changeset
|
316 |
and assemble_plain_dict (Dict_Const (inst, dss)) = |
55150 | 317 |
assemble_constapp (Class_Instance inst) dss [] |
62539 | 318 |
| assemble_plain_dict (Dict_Var { var, index, ... }) = |
319 |
nbe_dict var index |
|
25924 | 320 |
|
28350 | 321 |
fun assemble_iterm constapp = |
24155 | 322 |
let |
28350 | 323 |
fun of_iterm match_cont t = |
25944 | 324 |
let |
28054 | 325 |
val (t', ts) = Code_Thingol.unfold_app t |
28350 | 326 |
in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
327 |
and of_iapp match_cont (IConst { sym, dicts = dss, ... }) ts = constapp sym dss ts |
31889 | 328 |
| of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts |
31724 | 329 |
| of_iapp match_cont ((v, _) `|=> t) ts = |
31888 | 330 |
nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47609
diff
changeset
|
331 |
| of_iapp match_cont (ICase { term = t, clauses = clauses, primitive = t0, ... }) ts = |
28350 | 332 |
nbe_apps (ml_cases (of_iterm NONE t) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47609
diff
changeset
|
333 |
(map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) clauses |
28350 | 334 |
@ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts |
25944 | 335 |
in of_iterm end; |
24155 | 336 |
|
28350 | 337 |
fun subst_nonlin_vars args = |
338 |
let |
|
339 |
val vs = (fold o Code_Thingol.fold_varnames) |
|
33002 | 340 |
(fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args []; |
28350 | 341 |
val names = Name.make_context (map fst vs); |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43323
diff
changeset
|
342 |
fun declare v k ctxt = |
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43323
diff
changeset
|
343 |
let val vs = Name.invent ctxt v k |
28350 | 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) |
|
31889 | 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) |
|
28350 | 353 |
| NONE => (t, samepairs)) |
354 |
| subst_vars (t1 `$ t2) samepairs = samepairs |
|
355 |
|> subst_vars t1 |
|
356 |
||>> subst_vars t2 |
|
357 |
|>> (op `$) |
|
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47609
diff
changeset
|
358 |
| subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs; |
28350 | 359 |
val (args', _) = fold_map subst_vars args samepairs; |
360 |
in (samepairs, args') end; |
|
361 |
||
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
362 |
fun assemble_eqn sym dicts default_args (i, (args, rhs)) = |
28337 | 363 |
let |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
364 |
val match_cont = if Code_Symbol.is_value sym then NONE |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
365 |
else SOME (nbe_apps_local idx_of (i + 1) sym (dicts @ default_args)); |
28350 | 366 |
val assemble_arg = assemble_iterm |
57435 | 367 |
(fn sym' => fn dss => fn ts => nbe_apps_constr ctxt idx_of sym' ((maps o map) (K "_") |
368 |
dss @ ts)) NONE; |
|
35371 | 369 |
val assemble_rhs = assemble_iterm assemble_constapp match_cont; |
28350 | 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 |
|
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
373 |
else ml_if (ml_and (map nbe_same samepairs)) |
55043
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
374 |
(assemble_rhs rhs) (the match_cont); |
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
375 |
val eqns = case match_cont |
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
376 |
of NONE => [([ml_list (rev (dicts @ s_args))], s_rhs)] |
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
377 |
| SOME default_rhs => |
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
378 |
[([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs), |
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
379 |
([ml_list (rev (dicts @ default_args))], default_rhs)] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
380 |
in (nbe_fun idx_of i sym, eqns) end; |
28337 | 381 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
382 |
fun assemble_eqns (sym, (num_args, (dicts, eqns))) = |
25944 | 383 |
let |
28337 | 384 |
val default_args = map nbe_default |
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43323
diff
changeset
|
385 |
(Name.invent Name.context "a" (num_args - length dicts)); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
386 |
val eqns' = map_index (assemble_eqn sym dicts default_args) eqns |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
387 |
@ (if Code_Symbol.is_value sym then [] else [(nbe_fun idx_of (length eqns) sym, |
28337 | 388 |
[([ml_list (rev (dicts @ default_args))], |
57435 | 389 |
nbe_apps_constr ctxt idx_of sym (dicts @ default_args))])]); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
390 |
in (eqns', nbe_abss num_args (nbe_fun idx_of 0 sym)) end; |
24155 | 391 |
|
25944 | 392 |
val (fun_vars, fun_vals) = map_split assemble_eqns eqnss'; |
55043
acefda71629b
prefer indexes user for pattern matching to print concrete names for symbols, do not rely on printable unique identifiers
haftmann
parents:
54889
diff
changeset
|
393 |
val deps_vars = ml_list (map (nbe_fun idx_of 0) deps); |
28337 | 394 |
in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end; |
25944 | 395 |
|
35371 | 396 |
|
63162 | 397 |
(* compilation of equations *) |
25944 | 398 |
|
55757 | 399 |
fun compile_eqnss ctxt nbe_program raw_deps [] = [] |
400 |
| compile_eqnss ctxt nbe_program raw_deps eqnss = |
|
24155 | 401 |
let |
26064 | 402 |
val (deps, deps_vals) = split_list (map_filter |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
403 |
(fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps); |
26064 | 404 |
val idx_of = raw_deps |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
405 |
|> map (fn dep => (dep, snd (Code_Symbol.Graph.get_node nbe_program dep))) |
26064 | 406 |
|> AList.lookup (op =) |
407 |
|> (fn f => the o f); |
|
57435 | 408 |
val s = assemble_eqnss ctxt idx_of deps eqnss; |
24155 | 409 |
val cs = map fst eqnss; |
25944 | 410 |
in |
411 |
s |
|
57435 | 412 |
|> traced ctxt (fn s => "\n--- code to be evaluated:\n" ^ s) |
38931
5e84c11c4b8a
evaluate takes ml context and ml expression parameter
haftmann
parents:
38676
diff
changeset
|
413 |
|> pair "" |
39911 | 414 |
|> Code_Runtime.value ctxt univs_cookie |
25944 | 415 |
|> (fn f => f deps_vals) |
416 |
|> (fn univs => cs ~~ univs) |
|
417 |
end; |
|
25190 | 418 |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30288
diff
changeset
|
419 |
|
63162 | 420 |
(* extraction of equations from statements *) |
24155 | 421 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
422 |
fun dummy_const sym dss = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
423 |
IConst { sym = sym, typargs = [], dicts = dss, |
58397 | 424 |
dom = [], annotation = NONE }; |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47609
diff
changeset
|
425 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
426 |
fun eqns_of_stmt (_, Code_Thingol.NoStmt) = |
54889
4121d64fde90
explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents:
52519
diff
changeset
|
427 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
428 |
| eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) = |
25101 | 429 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
430 |
| eqns_of_stmt (sym_const, Code_Thingol.Fun (((vs, _), eqns), _)) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
431 |
[(sym_const, (vs, map fst eqns))] |
28054 | 432 |
| eqns_of_stmt (_, Code_Thingol.Datatypecons _) = |
25101 | 433 |
[] |
28054 | 434 |
| eqns_of_stmt (_, Code_Thingol.Datatype _) = |
25101 | 435 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
436 |
| eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) = |
25101 | 437 |
let |
55150 | 438 |
val syms = map Class_Relation classrels @ map (Constant o fst) classparams; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
439 |
val params = Name.invent Name.context "d" (length syms); |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
440 |
fun mk (k, sym) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
441 |
(sym, ([(v, [])], |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
442 |
[([dummy_const sym_class [] `$$ map (IVar o SOME) params], |
31889 | 443 |
IVar (SOME (nth params k)))])); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
444 |
in map_index mk syms end |
28054 | 445 |
| eqns_of_stmt (_, Code_Thingol.Classrel _) = |
25101 | 446 |
[] |
28054 | 447 |
| eqns_of_stmt (_, Code_Thingol.Classparam _) = |
25101 | 448 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
449 |
| eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) = |
55150 | 450 |
[(sym_inst, (vs, [([], dummy_const (Type_Class class) [] `$$ |
451 |
map (fn (class, dss) => dummy_const (Class_Instance (tyco, class)) dss) superinsts |
|
52519
598addf65209
explicit hint for domain of class parameters in instance statements
haftmann
parents:
52474
diff
changeset
|
452 |
@ map (IConst o fst o snd o fst) inst_params)]))]; |
24155 | 453 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
454 |
|
63162 | 455 |
(* compilation of whole programs *) |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
456 |
|
39399 | 457 |
fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) = |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
458 |
if can (Code_Symbol.Graph.get_node nbe_program) name |
39399 | 459 |
then (nbe_program, (maxidx, idx_tab)) |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
460 |
else (Code_Symbol.Graph.new_node (name, (NONE, maxidx)) nbe_program, |
39399 | 461 |
(maxidx + 1, Inttab.update_new (maxidx, name) idx_tab)); |
462 |
||
55757 | 463 |
fun compile_stmts ctxt stmts_deps = |
25101 | 464 |
let |
465 |
val names = map (fst o fst) stmts_deps; |
|
466 |
val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; |
|
467 |
val eqnss = maps (eqns_of_stmt o fst) stmts_deps; |
|
26064 | 468 |
val refl_deps = names_deps |
25190 | 469 |
|> maps snd |
470 |
|> distinct (op =) |
|
26064 | 471 |
|> fold (insert (op =)) names; |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
472 |
fun compile nbe_program = eqnss |
55757 | 473 |
|> compile_eqnss ctxt nbe_program refl_deps |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
474 |
|> rpair nbe_program; |
25101 | 475 |
in |
39399 | 476 |
fold ensure_const_idx refl_deps |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
477 |
#> apfst (fold (fn (name, deps) => fold (curry Code_Symbol.Graph.add_edge name) deps) names_deps |
26064 | 478 |
#> compile |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
479 |
#-> fold (fn (name, univ) => (Code_Symbol.Graph.map_node name o apfst) (K (SOME univ)))) |
25101 | 480 |
end; |
24155 | 481 |
|
63164 | 482 |
fun compile_program { ctxt, program } = |
25101 | 483 |
let |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
484 |
fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
485 |
then (nbe_program, (maxidx, idx_tab)) |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
486 |
else (nbe_program, (maxidx, idx_tab)) |
55757 | 487 |
|> compile_stmts ctxt (map (fn name => ((name, Code_Symbol.Graph.get_node program name), |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
488 |
Code_Symbol.Graph.immediate_succs program name)) names); |
28706 | 489 |
in |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
490 |
fold_rev add_stmts (Code_Symbol.Graph.strong_conn program) |
28706 | 491 |
end; |
24155 | 492 |
|
25944 | 493 |
|
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
494 |
(** normalization **) |
25944 | 495 |
|
63162 | 496 |
(* compilation and reconstruction of terms *) |
25944 | 497 |
|
63162 | 498 |
fun compile_term { ctxt, nbe_program, deps, term = (vs, t) } = |
25924 | 499 |
let |
500 |
val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; |
|
501 |
in |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
502 |
(Code_Symbol.value, (vs, [([], t)])) |
55757 | 503 |
|> singleton (compile_eqnss ctxt nbe_program deps) |
25924 | 504 |
|> snd |
31064 | 505 |
|> (fn t => apps t (rev dict_frees)) |
25924 | 506 |
end; |
24155 | 507 |
|
63162 | 508 |
fun reconstruct_term ctxt (idx_tab : Code_Symbol.T Inttab.table) t = |
24155 | 509 |
let |
25101 | 510 |
fun take_until f [] = [] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
511 |
| take_until f (x :: xs) = if f x then [] else x :: take_until f xs; |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
512 |
fun is_dict (Const (idx, _)) = |
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
513 |
(case Inttab.lookup idx_tab idx of |
55150 | 514 |
SOME (Constant _) => false |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
515 |
| _ => true) |
25101 | 516 |
| is_dict (DFree _) = true |
517 |
| is_dict _ = false; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
518 |
fun const_of_idx idx = |
55150 | 519 |
case Inttab.lookup idx_tab idx of SOME (Constant const) => const; |
24155 | 520 |
fun of_apps bounds (t, ts) = |
521 |
fold_map (of_univ bounds) ts |
|
522 |
#>> (fn ts' => list_comb (t, rev ts')) |
|
26064 | 523 |
and of_univ bounds (Const (idx, ts)) typidx = |
24155 | 524 |
let |
25101 | 525 |
val ts' = take_until is_dict ts; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
526 |
val const = const_of_idx idx; |
31957 | 527 |
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
|
528 |
Type_Infer.param typidx (v ^ string_of_int i, [])) |
55757 | 529 |
(Sign.the_const_type (Proof_Context.theory_of ctxt) const); |
30022
1d8b8fa19074
maintain order of constructors in datatypes; clarified conventions for type schemes
haftmann
parents:
29272
diff
changeset
|
530 |
val typidx' = typidx + 1; |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
531 |
in of_apps bounds (Term.Const (const, T), ts') typidx' end |
27499 | 532 |
| of_univ bounds (BVar (n, ts)) typidx = |
533 |
of_apps bounds (Bound (bounds - n - 1), ts) typidx |
|
24155 | 534 |
| of_univ bounds (t as Abs _) typidx = |
535 |
typidx |
|
25924 | 536 |
|> of_univ (bounds + 1) (apps t [BVar (bounds, [])]) |
24155 | 537 |
|-> (fn t' => pair (Term.Abs ("u", dummyT, t'))) |
538 |
in of_univ 0 t 0 |> fst end; |
|
539 |
||
63162 | 540 |
fun compile_and_reconstruct_term { ctxt, nbe_program, idx_tab, deps, term } = |
541 |
compile_term |
|
542 |
{ ctxt = ctxt, nbe_program = nbe_program, deps = deps, term = term } |
|
543 |
|> reconstruct_term ctxt idx_tab; |
|
35371 | 544 |
|
63162 | 545 |
fun normalize_term (nbe_program, idx_tab) raw_ctxt t_original ((vs, ty) : typscheme, t) deps = |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
546 |
let |
55757 | 547 |
val ctxt = Syntax.init_pretty_global (Proof_Context.theory_of raw_ctxt); |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
548 |
val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); |
56969
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
549 |
fun type_infer t' = |
62958
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62539
diff
changeset
|
550 |
Syntax.check_term |
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62539
diff
changeset
|
551 |
(ctxt |
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62539
diff
changeset
|
552 |
|> Config.put Type_Infer.object_logic false |
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62539
diff
changeset
|
553 |
|> Config.put Type_Infer_Context.const_sorts false) |
56969
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
554 |
(Type.constraint (fastype_of t_original) t'); |
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
555 |
fun check_tvars t' = |
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
556 |
if null (Term.add_tvars t' []) then t' |
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
557 |
else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t'); |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
558 |
in |
63164 | 559 |
Code_Preproc.timed "computing NBE expression" #ctxt compile_and_reconstruct_term |
63162 | 560 |
{ ctxt = ctxt, nbe_program = nbe_program, idx_tab = idx_tab, deps = deps, term = (vs, t) } |
57435 | 561 |
|> traced ctxt (fn t => "Normalized:\n" ^ string_of_term t) |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
562 |
|> type_infer |
57435 | 563 |
|> traced ctxt (fn t => "Types inferred:\n" ^ string_of_term t) |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
564 |
|> check_tvars |
57435 | 565 |
|> traced ctxt (fn _ => "---\n") |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
566 |
end; |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
567 |
|
39436 | 568 |
|
25101 | 569 |
(* function store *) |
570 |
||
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33522
diff
changeset
|
571 |
structure Nbe_Functions = Code_Data |
25101 | 572 |
( |
55150 | 573 |
type T = (Univ option * int) Code_Symbol.Graph.T * (int * Code_Symbol.T Inttab.table); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
574 |
val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty)); |
25101 | 575 |
); |
576 |
||
55757 | 577 |
fun compile ignore_cache ctxt program = |
24155 | 578 |
let |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
579 |
val (nbe_program, (_, idx_tab)) = |
55757 | 580 |
Nbe_Functions.change (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt)) |
63164 | 581 |
(Code_Preproc.timed "compiling NBE program" #ctxt |
582 |
compile_program { ctxt = ctxt, program = program }); |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
583 |
in (nbe_program, idx_tab) end; |
24155 | 584 |
|
32544 | 585 |
|
47572 | 586 |
(* evaluation oracle *) |
24155 | 587 |
|
55757 | 588 |
fun mk_equals ctxt lhs raw_rhs = |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
589 |
let |
59586 | 590 |
val ty = Thm.typ_of_cterm lhs; |
74295 | 591 |
val eq = Thm.cterm_of ctxt \<^Const>\<open>Pure.eq ty\<close>; |
59642 | 592 |
val rhs = Thm.cterm_of ctxt raw_rhs; |
30947 | 593 |
in Thm.mk_binop eq lhs rhs end; |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
594 |
|
38669
9ff76d0f0610
refined and unified naming convention for dynamic code evaluation techniques
haftmann
parents:
37449
diff
changeset
|
595 |
val (_, raw_oracle) = Context.>>> (Context.map_theory_result |
67149 | 596 |
(Thm.add_oracle (\<^binding>\<open>normalization_by_evaluation\<close>, |
56974
4ab498f41eee
accurate separation of static and dynamic context
haftmann
parents:
56973
diff
changeset
|
597 |
fn (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct) => |
63162 | 598 |
mk_equals ctxt ct (normalize_term nbe_program_idx_tab ctxt (Thm.term_of ct) vs_ty_t deps)))); |
55757 | 599 |
|
56974
4ab498f41eee
accurate separation of static and dynamic context
haftmann
parents:
56973
diff
changeset
|
600 |
fun oracle nbe_program_idx_tab ctxt vs_ty_t deps ct = |
4ab498f41eee
accurate separation of static and dynamic context
haftmann
parents:
56973
diff
changeset
|
601 |
raw_oracle (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct); |
31064 | 602 |
|
55757 | 603 |
fun dynamic_conv ctxt = lift_triv_classes_conv ctxt |
63158
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
604 |
(fn ctxt' => Code_Thingol.dynamic_conv ctxt' (fn program => |
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
605 |
oracle (compile false ctxt program) ctxt')); |
30942
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents:
30672
diff
changeset
|
606 |
|
55757 | 607 |
fun dynamic_value ctxt = lift_triv_classes_rew ctxt |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
608 |
(Code_Thingol.dynamic_value ctxt I (fn program => |
63162 | 609 |
normalize_term (compile false ctxt program) ctxt)); |
24155 | 610 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56972
diff
changeset
|
611 |
fun static_conv (ctxt_consts as { ctxt, ... }) = |
55757 | 612 |
let |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
613 |
val conv = Code_Thingol.static_conv_thingol ctxt_consts |
63164 | 614 |
(fn { program, deps = _ } => oracle (compile true ctxt program)); |
63158
534f16b0ca39
explicit quasi-global context for nbe conversions -- works around quasi-global type variable handling in lift_triv_classes_conv
haftmann
parents:
63157
diff
changeset
|
615 |
in fn ctxt' => lift_triv_classes_conv ctxt' conv end; |
39399 | 616 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56972
diff
changeset
|
617 |
fun static_value { ctxt, consts } = |
55757 | 618 |
let |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
619 |
val comp = Code_Thingol.static_value { ctxt = ctxt, lift_postproc = I, consts = consts } |
63164 | 620 |
(fn { program, deps = _ } => normalize_term (compile false ctxt program)); |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
621 |
in fn ctxt' => lift_triv_classes_rew ctxt' (comp ctxt') end; |
41184 | 622 |
|
24155 | 623 |
end; |