author | nipkow |
Thu, 17 Jul 2025 21:06:22 +0100 | |
changeset 82885 | 5d2a599f88af |
parent 82510 | 5601f5cce4c6 |
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 |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
16 |
datatype Type = |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
17 |
Type of string * Type list |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
18 |
| TParam of string |
25204 | 19 |
datatype Univ = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
20 |
Const of (int * Type list) * Univ list (*named (uninterpreted) constants*) |
25924 | 21 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
22 |
| BVar of int * Univ list (*bound variables, named*) |
28350 | 23 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
25924 | 24 |
val apps: Univ -> Univ list -> Univ (*explicit applications*) |
25944 | 25 |
val abss: int -> (Univ list -> Univ) -> Univ |
28337 | 26 |
(*abstractions as closures*) |
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
27 |
val same: Univ * Univ -> bool |
24155 | 28 |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
29 |
val put_result: (unit -> (Type list -> Univ) list -> (Type list -> Univ) list) |
63164 | 30 |
-> Proof.context -> Proof.context |
57435 | 31 |
val trace: bool Config.T |
25924 | 32 |
|
32544 | 33 |
val add_const_alias: thm -> theory -> theory |
24155 | 34 |
end; |
35 |
||
36 |
structure Nbe: NBE = |
|
37 |
struct |
|
38 |
||
39 |
(* generic non-sense *) |
|
40 |
||
67149 | 41 |
val trace = Attrib.setup_config_bool \<^binding>\<open>nbe_trace\<close> (K false); |
57435 | 42 |
fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x; |
24155 | 43 |
|
44 |
||
32544 | 45 |
(** certificates and oracle for "trivial type classes" **) |
46 |
||
33522 | 47 |
structure Triv_Class_Data = Theory_Data |
32544 | 48 |
( |
49 |
type T = (class * thm) list; |
|
50 |
val empty = []; |
|
33522 | 51 |
fun merge data : T = AList.merge (op =) (K true) data; |
32544 | 52 |
); |
53 |
||
54 |
fun add_const_alias thm thy = |
|
55 |
let |
|
56 |
val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm) |
|
57 |
of SOME ofclass_eq => ofclass_eq |
|
61268 | 58 |
| _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); |
32544 | 59 |
val (T, class) = case try Logic.dest_of_class ofclass |
60 |
of SOME T_class => T_class |
|
61268 | 61 |
| _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); |
32544 | 62 |
val tvar = case try Term.dest_TVar T |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
48072
diff
changeset
|
63 |
of SOME (tvar as (_, sort)) => if null (filter (can (Axclass.get_info thy)) sort) |
32544 | 64 |
then tvar |
61268 | 65 |
else error ("Bad sort: " ^ Thm.string_of_thm_global thy thm) |
66 |
| _ => error ("Bad type: " ^ Thm.string_of_thm_global thy thm); |
|
32544 | 67 |
val _ = if Term.add_tvars eqn [] = [tvar] then () |
61268 | 68 |
else error ("Inconsistent type: " ^ Thm.string_of_thm_global thy thm); |
32544 | 69 |
val lhs_rhs = case try Logic.dest_equals eqn |
70 |
of SOME lhs_rhs => lhs_rhs |
|
71 |
| _ => 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
|
72 |
val c_c' = case try (apply2 (Axclass.unoverload_const thy o dest_Const)) lhs_rhs |
32544 | 73 |
of SOME c_c' => c_c' |
74 |
| _ => error ("Not an equation with two constants: " |
|
75 |
^ Syntax.string_of_term_global thy eqn); |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
48072
diff
changeset
|
76 |
val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then () |
61268 | 77 |
else error ("Inconsistent class: " ^ Thm.string_of_thm_global thy thm); |
61096 | 78 |
in Triv_Class_Data.map (AList.update (op =) (class, Thm.trim_context thm)) thy end; |
32544 | 79 |
|
80 |
local |
|
81 |
||
82 |
val get_triv_classes = map fst o Triv_Class_Data.get; |
|
83 |
||
78795
f7e972d567f3
clarified signature: more concise variations on implicit theory setup;
wenzelm
parents:
77707
diff
changeset
|
84 |
val (_, triv_of_class) = |
f7e972d567f3
clarified signature: more concise variations on implicit theory setup;
wenzelm
parents:
77707
diff
changeset
|
85 |
Theory.setup_result (Thm.add_oracle (\<^binding>\<open>triv_of_class\<close>, |
f7e972d567f3
clarified signature: more concise variations on implicit theory setup;
wenzelm
parents:
77707
diff
changeset
|
86 |
fn (thy, T, class) => Thm.global_cterm_of thy (Logic.mk_of_class (T, class)))); |
32544 | 87 |
|
88 |
in |
|
89 |
||
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
|
90 |
fun lift_triv_classes_conv orig_ctxt conv ct = |
32544 | 91 |
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
|
92 |
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
|
93 |
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
|
94 |
(*FIXME quasi-global context*) |
32544 | 95 |
val algebra = Sign.classes_of thy; |
96 |
val triv_classes = get_triv_classes thy; |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
97 |
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
|
98 |
fun mk_entry (v, sort) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
99 |
let |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
100 |
val T = TFree (v, sort); |
60322 | 101 |
val cT = Thm.ctyp_of ctxt T; |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
102 |
val triv_sort = additional_classes sort; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
103 |
in |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
104 |
(v, (Sorts.inter_sort algebra (sort, triv_sort), |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
105 |
(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
|
106 |
@ 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
|
107 |
end; |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
108 |
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
|
109 |
fun instantiate thm = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
110 |
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
|
111 |
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
|
112 |
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
|
113 |
val instT = map2 (fn v => fn (_, (_, (cT, _))) => (v, cT)) tvars vs_tab; |
74282 | 114 |
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
|
115 |
fun of_class (TFree (v, _), class) = |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
116 |
the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class) |
60322 | 117 |
| of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ ctxt T); |
32544 | 118 |
fun strip_of_class thm = |
119 |
let |
|
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
120 |
val prems_of_class = Thm.prop_of thm |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
121 |
|> Logic.strip_imp_prems |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
122 |
|> map (Logic.dest_of_class #> of_class); |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
123 |
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
|
124 |
in |
c3a04614c710
reactivated Thm.legacy_unconstrainT for Nbe.lift_triv_classes_conv;
wenzelm
parents:
36768
diff
changeset
|
125 |
ct |
60322 | 126 |
|> Thm.term_of |
127 |
|> (map_types o map_type_tfree) |
|
47609 | 128 |
(fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v)) |
60322 | 129 |
|> 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
|
130 |
|> conv ctxt |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
131 |
|> 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
|
132 |
|> Thm.varifyT_global |
36982
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
133 |
|> Thm.unconstrainT |
1d4478a797c2
new version of triv_of_class machinery without legacy_unconstrain
haftmann
parents:
36960
diff
changeset
|
134 |
|> instantiate |
32544 | 135 |
|> strip_of_class |
136 |
end; |
|
137 |
||
55757 | 138 |
fun lift_triv_classes_rew ctxt rew t = |
32544 | 139 |
let |
55757 | 140 |
val thy = Proof_Context.theory_of ctxt; |
32544 | 141 |
val algebra = Sign.classes_of thy; |
142 |
val triv_classes = get_triv_classes thy; |
|
143 |
val vs = Term.add_tfrees t []; |
|
52473
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
144 |
in |
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
145 |
t |
32544 | 146 |
|> (map_types o map_type_tfree) |
147 |
(fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes))) |
|
148 |
|> rew |
|
149 |
|> (map_types o map_type_tfree) |
|
52473
a2407f62a29f
do not choke on type variables emerging during rewriting
haftmann
parents:
51685
diff
changeset
|
150 |
(fn (v, sort) => TFree (v, the_default sort (AList.lookup (op =) vs v))) |
32544 | 151 |
end; |
152 |
||
153 |
end; |
|
154 |
||
155 |
||
156 |
(** the semantic universe **) |
|
24155 | 157 |
|
158 |
(* |
|
82373 | 159 |
Functions are given by their semantic function value. To avoid |
24155 | 160 |
trouble with the ML-type system, these functions have the most |
161 |
generic type, that is "Univ list -> Univ". The calling convention is |
|
162 |
that the arguments come as a list, the last argument first. In |
|
163 |
other words, a function call that usually would look like |
|
164 |
||
165 |
f x_1 x_2 ... x_n or f(x_1,x_2, ..., x_n) |
|
166 |
||
167 |
would be in our convention called as |
|
168 |
||
169 |
f [x_n,..,x_2,x_1] |
|
170 |
||
171 |
Moreover, to handle functions that are still waiting for some |
|
172 |
arguments we have additionally a list of arguments collected to far |
|
173 |
and the number of arguments we're still waiting for. |
|
174 |
*) |
|
175 |
||
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
176 |
datatype Type = |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
177 |
Type of string * Type list |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
178 |
| TParam of string |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
179 |
|
25204 | 180 |
datatype Univ = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
181 |
Const of (int * Type list) * Univ list (*named (uninterpreted) constants*) |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
182 |
| DFree of string * int (*free (uninterpreted) dictionary parameters*) |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
183 |
| BVar of int * Univ list (*bound variables, named*) |
24155 | 184 |
| Abs of (int * (Univ list -> Univ)) * Univ list |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
185 |
(*abstractions as closures*) |
24155 | 186 |
|
187 |
(* constructor functions *) |
|
188 |
||
25944 | 189 |
fun abss n f = Abs ((n, f), []); |
82454 | 190 |
fun apps (Abs ((n, f), us)) ws = let val k = n - length ws in |
27499 | 191 |
case int_ord (k, 0) |
82454 | 192 |
of EQUAL => f (ws @ us) |
193 |
| LESS => let val (ws1, ws2) = chop (~ k) ws in apps (f (ws2 @ us)) ws1 end |
|
194 |
| GREATER => Abs ((k, f), ws @ us) (*note: reverse convention also for apps!*) |
|
27499 | 195 |
end |
25924 | 196 |
| apps (Const (name, xs)) ys = Const (name, ys @ xs) |
27499 | 197 |
| apps (BVar (n, xs)) ys = BVar (n, ys @ xs); |
24155 | 198 |
|
82454 | 199 |
fun same_type (Type (tyco1, ty1), Type (tyco2, tys2)) = |
200 |
(tyco1 = tyco2) andalso eq_list same_type (ty1, tys2) |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
201 |
| same_type (TParam v1, TParam v2) = (v1 = v2) |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
202 |
| same_type _ = false; |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
203 |
|
82454 | 204 |
fun same (Const ((k1, typargs1), us1), Const ((k2, typargs2), us2)) = |
205 |
(k1 = k2) andalso eq_list same_type (typargs1, typargs2) andalso eq_list same (us1, us2) |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
206 |
| same (DFree (n1, i1), DFree (n2, i2)) = (n1 = n2) andalso (i1 = i2) |
82454 | 207 |
| same (BVar (i1, us1), BVar (i2, us2)) = (i1 = i2) andalso eq_list same (us1, us2) |
41037
6d6f23b3a879
removed experimental equality checking of closures; acknowledge underapproximation of equality in function name
haftmann
parents:
40730
diff
changeset
|
208 |
| same _ = false; |
40730
2aa0390a2da7
nbe decides equality of abstractions by extensionality
haftmann
parents:
40564
diff
changeset
|
209 |
|
24155 | 210 |
|
211 |
(** assembling and compiling ML code from terms **) |
|
212 |
||
213 |
(* abstract ML syntax *) |
|
214 |
||
215 |
infix 9 `$` `$$`; |
|
82453 | 216 |
infix 8 `*`; |
82510 | 217 |
|
82453 | 218 |
fun e1 `$` e2 = enclose "(" ")" (e1 ^ " " ^ e2); |
25101 | 219 |
fun e `$$` [] = e |
82453 | 220 |
| e `$$` es = enclose "(" ")" (e ^ " " ^ implode_space es); |
221 |
fun ml_abs v e = enclose "(" ")" ("fn " ^ v ^ " => " ^ e); |
|
24155 | 222 |
|
82453 | 223 |
fun ml_cases e cs = enclose "(" ")" |
224 |
("case " ^ e ^ " of " ^ space_implode " | " (map (fn (p, e) => p ^ " => " ^ e) cs)); |
|
82506 | 225 |
fun ml_let d e = "\n let\n" ^ prefix_lines " " d ^ "\n in " ^ e ^ " end"; |
24155 | 226 |
|
28350 | 227 |
fun ml_and [] = "true" |
82453 | 228 |
| ml_and [e] = e |
229 |
| ml_and es = enclose "(" ")" (space_implode " andalso " es); |
|
82506 | 230 |
fun ml_if b e1 e2 = enclose "(" ")" (implode_space ["if", b, "then", e1, "else", e2]); |
28350 | 231 |
|
82453 | 232 |
fun e1 `*` e2 = enclose "(" ")" (e1 ^ ", " ^ e2); |
233 |
fun ml_list es = enclose "[" "]" (commas es); |
|
24155 | 234 |
|
82506 | 235 |
fun ml_exc s = enclose "(" ")" ("raise Fail " ^ quote s); |
236 |
||
24155 | 237 |
fun ml_fundefs ([(name, [([], e)])]) = |
82506 | 238 |
"val " ^ name ^ " = " ^ e ^ "" |
24155 | 239 |
| ml_fundefs (eqs :: eqss) = |
240 |
let |
|
241 |
fun fundef (name, eqs) = |
|
242 |
let |
|
80910 | 243 |
fun eqn (es, e) = name ^ " " ^ implode_space es ^ " = " ^ e |
24155 | 244 |
in space_implode "\n | " (map eqn eqs) end; |
245 |
in |
|
246 |
(prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss |
|
26931 | 247 |
|> cat_lines |
24155 | 248 |
end; |
249 |
||
35371 | 250 |
|
25944 | 251 |
(* nbe specific syntax and sandbox communication *) |
252 |
||
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41346
diff
changeset
|
253 |
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
|
254 |
( |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
255 |
type T = unit -> (Type list -> Univ) list -> (Type list -> Univ) list; |
59153 | 256 |
val empty: T = fn () => raise Fail "Univs"; |
257 |
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
|
258 |
); |
59153 | 259 |
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
|
260 |
val put_result = Univs.put; |
24155 | 261 |
|
262 |
local |
|
59151
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
263 |
val prefix = "Nbe."; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
264 |
val name_put = prefix ^ "put_result"; |
41068 | 265 |
val name_const = prefix ^ "Const"; |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
266 |
val name_type = prefix ^ "Type"; |
59151
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
267 |
val name_abss = prefix ^ "abss"; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
268 |
val name_apps = prefix ^ "apps"; |
a012574b78e7
proper exception for internal program failure, not user-error;
wenzelm
parents:
59058
diff
changeset
|
269 |
val name_same = prefix ^ "same"; |
24155 | 270 |
in |
271 |
||
59153 | 272 |
val univs_cookie = (get_result, put_result, name_put); |
25944 | 273 |
|
82454 | 274 |
(* |
275 |
Convention: parameters representing ("assembled") string representations of logical entities |
|
276 |
are prefixed with an "a_" -- unless they are an unqualified name ready to become part of |
|
277 |
an ML identifier. |
|
278 |
*) |
|
279 |
||
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
280 |
fun nbe_tparam v = "t_" ^ v; |
25101 | 281 |
fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; |
82506 | 282 |
fun nbe_global_param v = "w_" ^ v; |
24155 | 283 |
fun nbe_bound v = "v_" ^ v; |
31888 | 284 |
fun nbe_bound_optional NONE = "_" |
35500 | 285 |
| nbe_bound_optional (SOME v) = nbe_bound v; |
82454 | 286 |
fun nbe_type a_sym a_tys = name_type `$` (quote a_sym `*` ml_list a_tys); |
287 |
fun nbe_fun a_sym a_typargs = a_sym `$` ml_list a_typargs; |
|
24155 | 288 |
|
82506 | 289 |
(*note: these are the "turning spots" where proper argument order is established!*) |
82454 | 290 |
fun nbe_apps a_u [] = a_u |
291 |
| nbe_apps a_u a_us = name_apps `$$` [a_u, ml_list (rev a_us)]; |
|
292 |
fun nbe_apps_fun a_sym a_typargs a_us = nbe_fun a_sym a_typargs `$` ml_list (rev a_us); |
|
82508 | 293 |
fun nbe_apps_fallback_fun a_sym a_us = a_sym `$$` a_us; |
82510 | 294 |
fun nbe_apps_const a_sym a_typargs a_us = name_const `$` ((a_sym `*` ml_list a_typargs) `*` ml_list (rev a_us)); |
295 |
fun nbe_apps_constpat a_sym a_us = name_const `$` ((a_sym `*` "_") `*` ml_list (rev a_us)); |
|
25924 | 296 |
|
24155 | 297 |
fun nbe_abss 0 f = f `$` ml_list [] |
25944 | 298 |
| nbe_abss n f = name_abss `$$` [string_of_int n, f]; |
24155 | 299 |
|
82453 | 300 |
fun nbe_same (v1, v2) = enclose "(" ")" (name_same `$` (nbe_bound v1 `*` nbe_bound v2)); |
28350 | 301 |
|
24155 | 302 |
end; |
303 |
||
55150 | 304 |
open Basic_Code_Symbol; |
28054 | 305 |
open Basic_Code_Thingol; |
24155 | 306 |
|
35371 | 307 |
|
25865 | 308 |
(* code generation *) |
24155 | 309 |
|
82375 | 310 |
fun subst_nonlin_vars args = |
25944 | 311 |
let |
82375 | 312 |
val vs = (fold o Code_Thingol.fold_varnames) |
313 |
(fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args []; |
|
314 |
val names = Name.make_context (map fst vs); |
|
315 |
val (vs_renames, _) = fold_map (fn (v, k) => if k > 1 |
|
316 |
then Name.invent' v (k - 1) #>> (fn vs => (v, vs)) |
|
317 |
else pair (v, [])) vs names; |
|
318 |
val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames; |
|
319 |
fun subst_vars (t as IConst _) samepairs = (t, samepairs) |
|
320 |
| subst_vars (t as IVar NONE) samepairs = (t, samepairs) |
|
321 |
| subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v |
|
322 |
of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs) |
|
323 |
| NONE => (t, samepairs)) |
|
324 |
| subst_vars (t1 `$ t2) samepairs = samepairs |
|
325 |
|> subst_vars t1 |
|
326 |
||>> subst_vars t2 |
|
327 |
|>> (op `$) |
|
328 |
| subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs; |
|
329 |
val (args', _) = fold_map subst_vars args samepairs; |
|
330 |
in (samepairs, args') end; |
|
25944 | 331 |
|
82375 | 332 |
fun preprocess_eqns (sym, (vs, eqns)) = |
333 |
let |
|
82454 | 334 |
val a_tparams = map (fn (v, _) => nbe_tparam v) vs; |
335 |
val a_dict_params = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; |
|
336 |
val num_args = length a_dict_params + ((length o fst o hd) eqns); |
|
82506 | 337 |
val a_global_params = map nbe_global_param (Name.invent_global "a" (num_args - length a_dict_params)); |
338 |
in (sym, (num_args, (a_tparams, a_dict_params, a_global_params, (map o apfst) subst_nonlin_vars eqns))) end; |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
339 |
|
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
340 |
fun assemble_type (tyco `%% tys) = nbe_type tyco (map assemble_type tys) |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
341 |
| assemble_type (ITyVar v) = nbe_tparam v |
82375 | 342 |
|
82454 | 343 |
fun assemble_preprocessed_eqnss ctxt idx_of_sym deps eqnss = |
82375 | 344 |
let |
82507 | 345 |
fun suffixed_fun_ident suffix sym = space_implode "_" |
82510 | 346 |
["c", if Code_Symbol.is_value sym then "0" else string_of_int (idx_of_sym sym), |
82507 | 347 |
Code_Symbol.default_base sym, suffix]; |
348 |
val fun_ident = suffixed_fun_ident "nbe"; |
|
82508 | 349 |
fun fallback_fun_ident i = suffixed_fun_ident (string_of_int i); |
82510 | 350 |
fun const_ident sym = |
82374 | 351 |
if Config.get ctxt trace |
82454 | 352 |
then string_of_int (idx_of_sym sym) ^ " (*" ^ Code_Symbol.default_base sym ^ "*)" |
353 |
else string_of_int (idx_of_sym sym); |
|
82374 | 354 |
|
82506 | 355 |
fun assemble_fun sym = nbe_fun (fun_ident sym); |
356 |
fun assemble_app_fun sym = nbe_apps_fun (fun_ident sym); |
|
82508 | 357 |
fun assemble_app_fallback_fun i sym = nbe_apps_fallback_fun (fallback_fun_ident i sym); |
82510 | 358 |
fun assemble_app_const sym = nbe_apps_const (const_ident sym); |
359 |
fun assemble_app_constpat sym = nbe_apps_constpat (const_ident sym); |
|
82374 | 360 |
|
82454 | 361 |
fun assemble_constapp sym typargs dictss a_ts = |
25944 | 362 |
let |
82454 | 363 |
val a_typargs = map (assemble_type) typargs; |
364 |
val a_ts' = (maps o map) assemble_dict (map2 (fn ty => map (fn dict => (ty, dict))) typargs dictss) @ a_ts; |
|
82375 | 365 |
in case AList.lookup (op =) eqnss sym |
82454 | 366 |
of SOME (num_args, _) => if num_args <= length a_ts' |
367 |
then let val (a_ts1, a_ts2) = chop num_args a_ts' |
|
82506 | 368 |
in nbe_apps (assemble_app_fun sym a_typargs a_ts1) a_ts2 |
369 |
end else nbe_apps (nbe_abss num_args (assemble_fun sym a_typargs)) a_ts' |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
370 |
| NONE => if member (op =) deps sym |
82506 | 371 |
then nbe_apps (assemble_fun sym a_typargs) a_ts' |
82510 | 372 |
else assemble_app_const sym a_typargs a_ts' |
25924 | 373 |
end |
41100
6c0940392fb4
dictionary constants must permit explicit weakening of classes;
haftmann
parents:
41068
diff
changeset
|
374 |
and assemble_classrels classrels = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
375 |
fold_rev (fn classrel => assemble_constapp (Class_Relation classrel) [] [] o single) classrels |
82454 | 376 |
and assemble_dict (ty, Dict (classrels, dict)) = |
377 |
assemble_classrels classrels (assemble_plain_dict ty dict) |
|
82447 | 378 |
and assemble_plain_dict (_ `%% tys) (Dict_Const (inst, dictss)) = |
379 |
assemble_constapp (Class_Instance inst) tys (map snd dictss) [] |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
380 |
| assemble_plain_dict _ (Dict_Var { var, index, ... }) = |
82509 | 381 |
nbe_dict var index; |
25924 | 382 |
|
82509 | 383 |
fun assemble_iterm is_pat a_match_fallback t = |
24155 | 384 |
let |
82509 | 385 |
fun assemble_app (IConst { sym, typargs, dictss, ... }) = |
82510 | 386 |
if is_pat then fn a_ts => assemble_app_constpat sym ((maps o map) (K "_") dictss @ a_ts) |
82509 | 387 |
else assemble_constapp sym typargs dictss |
388 |
| assemble_app (IVar v) = nbe_apps (nbe_bound_optional v) |
|
389 |
| assemble_app ((v, _) `|=> (t, _)) = |
|
390 |
nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (assemble_iterm is_pat NONE t))) |
|
391 |
| assemble_app (ICase { term = t, clauses = clauses, primitive = t0, ... }) = |
|
392 |
nbe_apps (ml_cases (assemble_iterm is_pat NONE t) |
|
393 |
(map (fn (p, t) => (assemble_iterm true NONE p, assemble_iterm is_pat a_match_fallback t)) clauses |
|
394 |
@ [("_", case a_match_fallback of SOME s => s | NONE => assemble_iterm is_pat NONE t0)])) |
|
395 |
val (t', ts) = Code_Thingol.unfold_app t; |
|
396 |
val a_ts = fold_rev (cons o assemble_iterm is_pat NONE) ts []; |
|
397 |
in assemble_app t' a_ts end; |
|
82375 | 398 |
|
82508 | 399 |
fun assemble_fallback_fundef sym a_global_params ((samepairs, args), rhs) a_fallback_rhs = |
28350 | 400 |
let |
82509 | 401 |
val a_rhs_core = assemble_iterm false (SOME a_fallback_rhs) rhs; |
82508 | 402 |
val a_rhs = if null samepairs then a_rhs_core |
403 |
else ml_if (ml_and (map nbe_same samepairs)) a_rhs_core a_fallback_rhs; |
|
404 |
val a_fallback_eqn = if forall Code_Thingol.is_IVar args then NONE |
|
405 |
else SOME (replicate (length a_global_params) "_", a_fallback_rhs); |
|
82509 | 406 |
in (map (assemble_iterm true NONE) args, a_rhs) :: the_list a_fallback_eqn end; |
82506 | 407 |
|
82508 | 408 |
fun assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params [(([], []), rhs)] = |
82509 | 409 |
assemble_iterm false NONE rhs |
82508 | 410 |
| assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params eqns = |
411 |
let |
|
412 |
val a_fallback_syms = map_range (fn i => fallback_fun_ident i sym) (length eqns); |
|
413 |
val a_fallback_rhss = |
|
414 |
map_range (fn i => assemble_app_fallback_fun (i + 1) sym a_global_params) (length eqns - 1) |
|
82510 | 415 |
@ [assemble_app_const sym a_tparams (a_dict_params @ a_global_params)]; |
82508 | 416 |
in |
417 |
ml_let (ml_fundefs (a_fallback_syms ~~ |
|
418 |
map2 (assemble_fallback_fundef sym a_global_params) eqns a_fallback_rhss)) |
|
419 |
(assemble_app_fallback_fun 0 sym a_global_params) |
|
420 |
end; |
|
24155 | 421 |
|
82508 | 422 |
fun assemble_fundef (sym, (num_args, (a_tparams, a_dict_params, a_global_params, eqns))) = |
82506 | 423 |
let |
424 |
val a_lhs = [ml_list a_tparams, ml_list (rev (a_dict_params @ a_global_params))]; |
|
82508 | 425 |
val a_rhs = assemble_fallback_fundefs sym a_tparams a_dict_params a_global_params eqns; |
82506 | 426 |
val a_univ = ml_abs (ml_list a_tparams) (nbe_abss num_args (assemble_fun sym a_tparams)); |
427 |
in |
|
428 |
((fun_ident sym, [(a_lhs, a_rhs), (a_lhs, ml_exc (fun_ident sym))]), a_univ) |
|
429 |
end; |
|
82375 | 430 |
|
82508 | 431 |
val (a_fun_defs, a_fun_vals) = map_split assemble_fundef eqnss; |
82506 | 432 |
val dep_params = ml_list (map fun_ident deps); |
433 |
in ml_abs dep_params (ml_let (ml_fundefs a_fun_defs) (ml_list a_fun_vals)) end; |
|
25944 | 434 |
|
82454 | 435 |
fun assemble_eqnss ctxt idx_of_sym deps eqnss = |
436 |
assemble_preprocessed_eqnss ctxt idx_of_sym deps (map preprocess_eqns eqnss); |
|
82375 | 437 |
|
35371 | 438 |
|
63162 | 439 |
(* compilation of equations *) |
25944 | 440 |
|
55757 | 441 |
fun compile_eqnss ctxt nbe_program raw_deps [] = [] |
442 |
| compile_eqnss ctxt nbe_program raw_deps eqnss = |
|
24155 | 443 |
let |
26064 | 444 |
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
|
445 |
(fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps); |
82454 | 446 |
val idx_of_sym = raw_deps |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
447 |
|> map (fn dep => (dep, snd (Code_Symbol.Graph.get_node nbe_program dep))) |
26064 | 448 |
|> AList.lookup (op =) |
449 |
|> (fn f => the o f); |
|
82454 | 450 |
val s = assemble_eqnss ctxt idx_of_sym deps eqnss; |
82374 | 451 |
val syms = map fst eqnss; |
25944 | 452 |
in |
453 |
s |
|
57435 | 454 |
|> 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
|
455 |
|> pair "" |
39911 | 456 |
|> Code_Runtime.value ctxt univs_cookie |
82454 | 457 |
|> (fn dependent_fs => dependent_fs deps_vals) |
458 |
|> (fn fs => syms ~~ fs) |
|
25944 | 459 |
end; |
25190 | 460 |
|
30672
beaadd5af500
more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents:
30288
diff
changeset
|
461 |
|
63162 | 462 |
(* extraction of equations from statements *) |
24155 | 463 |
|
82454 | 464 |
fun dummy_const sym typargs dictss = |
465 |
IConst { sym = sym, typargs = typargs, dictss = dictss, |
|
77233
6bdd125d932b
explicit range types in abstractions
stuebinm <stuebinm@disroot.org>
parents:
74561
diff
changeset
|
466 |
dom = [], annotation = NONE, range = ITyVar "" }; |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47609
diff
changeset
|
467 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
468 |
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
|
469 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
470 |
| eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) = |
25101 | 471 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
472 |
| 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
|
473 |
[(sym_const, (vs, map fst eqns))] |
28054 | 474 |
| eqns_of_stmt (_, Code_Thingol.Datatypecons _) = |
25101 | 475 |
[] |
28054 | 476 |
| eqns_of_stmt (_, Code_Thingol.Datatype _) = |
25101 | 477 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
478 |
| eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) = |
25101 | 479 |
let |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
480 |
val syms = map (rpair [] o Class_Relation) classrels @ map (rpair [(v, [])] o Constant o fst) classparams; |
81507 | 481 |
val params = Name.invent_global "d" (length syms); |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
482 |
fun mk (k, (sym, vs)) = |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
483 |
(sym, (vs, |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
484 |
[([dummy_const sym_class [] [] `$$ map (IVar o SOME) params], |
31889 | 485 |
IVar (SOME (nth params k)))])); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
486 |
in map_index mk syms end |
28054 | 487 |
| eqns_of_stmt (_, Code_Thingol.Classrel _) = |
25101 | 488 |
[] |
28054 | 489 |
| eqns_of_stmt (_, Code_Thingol.Classparam _) = |
25101 | 490 |
[] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
491 |
| eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
492 |
[(sym_inst, (vs, [([], dummy_const (Type_Class class) [] [] `$$ |
82447 | 493 |
map (fn (class, dictss) => |
494 |
dummy_const (Class_Instance (tyco, class)) (map (ITyVar o fst) vs) (map snd dictss)) superinsts |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
495 |
@ map (IConst o fst o snd o fst) inst_params)]))]; |
24155 | 496 |
|
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
497 |
|
63162 | 498 |
(* compilation of whole programs *) |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
499 |
|
82454 | 500 |
fun ensure_sym_idx sym (nbe_program, (max_idx, sym_tab)) = |
501 |
if can (Code_Symbol.Graph.get_node nbe_program) sym |
|
502 |
then (nbe_program, (max_idx, sym_tab)) |
|
503 |
else (Code_Symbol.Graph.new_node (sym, (NONE, max_idx)) nbe_program, |
|
504 |
(max_idx + 1, Inttab.update_new (max_idx, sym) sym_tab)); |
|
39399 | 505 |
|
55757 | 506 |
fun compile_stmts ctxt stmts_deps = |
25101 | 507 |
let |
508 |
val names = map (fst o fst) stmts_deps; |
|
509 |
val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; |
|
510 |
val eqnss = maps (eqns_of_stmt o fst) stmts_deps; |
|
26064 | 511 |
val refl_deps = names_deps |
25190 | 512 |
|> maps snd |
513 |
|> distinct (op =) |
|
26064 | 514 |
|> fold (insert (op =)) names; |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
515 |
fun compile nbe_program = eqnss |
55757 | 516 |
|> compile_eqnss ctxt nbe_program refl_deps |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
517 |
|> rpair nbe_program; |
25101 | 518 |
in |
82454 | 519 |
fold ensure_sym_idx refl_deps |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
520 |
#> apfst (fold (fn (name, deps) => fold (curry Code_Symbol.Graph.add_edge name) deps) names_deps |
26064 | 521 |
#> compile |
82374 | 522 |
#-> fold (fn (sym, univ) => (Code_Symbol.Graph.map_node sym o apfst) (K (SOME univ)))) |
25101 | 523 |
end; |
24155 | 524 |
|
63164 | 525 |
fun compile_program { ctxt, program } = |
25101 | 526 |
let |
82454 | 527 |
fun add_stmts names (nbe_program, (max_idx, sym_tab)) = |
82375 | 528 |
if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names |
82454 | 529 |
then (nbe_program, (max_idx, sym_tab)) |
530 |
else (nbe_program, (max_idx, sym_tab)) |
|
82374 | 531 |
|> compile_stmts ctxt (map (fn sym => ((sym, Code_Symbol.Graph.get_node program sym), |
532 |
Code_Symbol.Graph.immediate_succs program sym)) names); |
|
28706 | 533 |
in |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
534 |
fold_rev add_stmts (Code_Symbol.Graph.strong_conn program) |
28706 | 535 |
end; |
24155 | 536 |
|
25944 | 537 |
|
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
538 |
(** normalization **) |
25944 | 539 |
|
63162 | 540 |
(* compilation and reconstruction of terms *) |
25944 | 541 |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
542 |
fun ad_hoc_eqn_of_term ((vs, _) : typscheme, t) = |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
543 |
(Code_Symbol.value, (vs, [([], t)])); |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
544 |
|
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
545 |
fun compile_term { ctxt, nbe_program, deps, tfrees, vs_ty_t = vs_ty_t as ((vs, _), _) } = |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
546 |
let |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
547 |
val tparams = map (fn (v, _) => TParam v) tfrees; |
25924 | 548 |
val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; |
549 |
in |
|
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
550 |
ad_hoc_eqn_of_term vs_ty_t |
55757 | 551 |
|> singleton (compile_eqnss ctxt nbe_program deps) |
25924 | 552 |
|> snd |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
553 |
|> (fn f => apps (f tparams) (rev dict_frees)) |
25924 | 554 |
end; |
24155 | 555 |
|
82454 | 556 |
fun reconstruct_term ctxt sym_tab tfrees = |
24155 | 557 |
let |
25101 | 558 |
fun take_until f [] = [] |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
559 |
| take_until f (x :: xs) = if f x then [] else x :: take_until f xs; |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
560 |
fun is_dict (Const ((idx, _), _)) = |
82454 | 561 |
(case Inttab.lookup sym_tab idx of |
55150 | 562 |
SOME (Constant _) => false |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
563 |
| _ => true) |
25101 | 564 |
| is_dict (DFree _) = true |
565 |
| is_dict _ = false; |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
55043
diff
changeset
|
566 |
fun const_of_idx idx = |
82454 | 567 |
case Inttab.lookup sym_tab idx of SOME (Constant const) => const; |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
568 |
fun reconstruct_type (Type (tyco, tys)) = Term.Type (tyco, map reconstruct_type tys) |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
569 |
| reconstruct_type (TParam v) = TFree (v, the (AList.lookup (op =) tfrees v)); |
24155 | 570 |
fun of_apps bounds (t, ts) = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
571 |
list_comb (t, rev (map (of_univ bounds) ts)) |
82454 | 572 |
and of_univ bounds (Const ((idx, tparams), us)) = |
24155 | 573 |
let |
82375 | 574 |
val const = const_of_idx idx; |
82454 | 575 |
val us' = take_until is_dict us; |
576 |
val T = Consts.instance (Proof_Context.consts_of ctxt) (const, map reconstruct_type tparams); |
|
577 |
in of_apps bounds (Term.Const (const, T), us') end |
|
578 |
| of_univ bounds (BVar (i, us)) = |
|
579 |
of_apps bounds (Bound (bounds - i - 1), us) |
|
580 |
| of_univ bounds (u as Abs _) = |
|
581 |
Term.Abs ("u", dummyT, of_univ (bounds + 1) (apps u [BVar (bounds, [])])) |
|
582 |
in of_univ 0 end; |
|
24155 | 583 |
|
82454 | 584 |
fun compile_and_reconstruct_term { ctxt, nbe_program, sym_tab, deps, tfrees, vs_ty_t } = |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
585 |
compile_term { ctxt = ctxt, nbe_program = nbe_program, deps = deps, tfrees = tfrees, vs_ty_t = vs_ty_t } |
82454 | 586 |
|> reconstruct_term ctxt sym_tab tfrees; |
35371 | 587 |
|
82443
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
588 |
fun retype_term ctxt t T = |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
589 |
let |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
590 |
val ctxt' = |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
591 |
ctxt |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
592 |
|> Variable.declare_typ T |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
593 |
|> Config.put Type_Infer.object_logic false |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
594 |
|> Config.put Type_Infer_Context.const_sorts false |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
595 |
in |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
596 |
singleton (Variable.export_terms ctxt' ctxt') (Syntax.check_term ctxt' (Type.constraint T t)) |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
597 |
end; |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
598 |
|
82454 | 599 |
fun normalize_term (nbe_program, sym_tab) raw_ctxt t_original vs_ty_t deps = |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
600 |
let |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
601 |
val T = fastype_of t_original; |
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
602 |
val tfrees = Term.add_tfrees t_original []; |
55757 | 603 |
val ctxt = Syntax.init_pretty_global (Proof_Context.theory_of raw_ctxt); |
82443
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
604 |
val string_of_term = |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
605 |
Syntax.string_of_term |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
606 |
(ctxt |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
607 |
|> Config.put show_types true |
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
608 |
|> Config.put show_sorts true); |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
609 |
fun retype t' = retype_term ctxt t' T; |
56969
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
610 |
fun check_tvars t' = |
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56925
diff
changeset
|
611 |
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
|
612 |
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
|
613 |
in |
63164 | 614 |
Code_Preproc.timed "computing NBE expression" #ctxt compile_and_reconstruct_term |
82454 | 615 |
{ ctxt = ctxt, nbe_program = nbe_program, sym_tab = sym_tab, deps = deps, |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
616 |
tfrees = tfrees, vs_ty_t = vs_ty_t } |
57435 | 617 |
|> traced ctxt (fn t => "Normalized:\n" ^ string_of_term t) |
82443
3e92066d2be7
restored type reconstruction for nbe: remaining type variables must remain schematic for checking
haftmann
parents:
82442
diff
changeset
|
618 |
|> retype |
57435 | 619 |
|> 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
|
620 |
|> check_tvars |
57435 | 621 |
|> traced ctxt (fn _ => "---\n") |
39392
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
622 |
end; |
7a0fcee7a2a3
more clear separation of static compilation and dynamic evaluation
haftmann
parents:
39388
diff
changeset
|
623 |
|
39436 | 624 |
|
25101 | 625 |
(* function store *) |
626 |
||
34173
458ced35abb8
reduced code generator cache to the baremost minimum
haftmann
parents:
33522
diff
changeset
|
627 |
structure Nbe_Functions = Code_Data |
25101 | 628 |
( |
82444
7a9164068583
represent constants with explicit typargs to avoid false positives for equality approximations like ‹card UNIV === card UNIV›
haftmann
parents:
82443
diff
changeset
|
629 |
type T = ((Type list -> 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
|
630 |
val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty)); |
25101 | 631 |
); |
632 |
||
55757 | 633 |
fun compile ignore_cache ctxt program = |
24155 | 634 |
let |
82454 | 635 |
val (nbe_program, (_, sym_tab)) = |
55757 | 636 |
Nbe_Functions.change (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt)) |
63164 | 637 |
(Code_Preproc.timed "compiling NBE program" #ctxt |
638 |
compile_program { ctxt = ctxt, program = program }); |
|
82454 | 639 |
in (nbe_program, sym_tab) end; |
24155 | 640 |
|
32544 | 641 |
|
47572 | 642 |
(* evaluation oracle *) |
24155 | 643 |
|
55757 | 644 |
fun mk_equals ctxt lhs raw_rhs = |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
645 |
let |
59586 | 646 |
val ty = Thm.typ_of_cterm lhs; |
74295 | 647 |
val eq = Thm.cterm_of ctxt \<^Const>\<open>Pure.eq ty\<close>; |
59642 | 648 |
val rhs = Thm.cterm_of ctxt raw_rhs; |
30947 | 649 |
in Thm.mk_binop eq lhs rhs end; |
26747
f32fa5f5bdd1
moved 'trivial classes' to foundation of code generator
haftmann
parents:
26739
diff
changeset
|
650 |
|
78795
f7e972d567f3
clarified signature: more concise variations on implicit theory setup;
wenzelm
parents:
77707
diff
changeset
|
651 |
val (_, raw_oracle) = |
f7e972d567f3
clarified signature: more concise variations on implicit theory setup;
wenzelm
parents:
77707
diff
changeset
|
652 |
Theory.setup_result (Thm.add_oracle (\<^binding>\<open>normalization_by_evaluation\<close>, |
82454 | 653 |
fn (nbe_program_sym_tab, ctxt, vs_ty_t, deps, ct) => |
654 |
mk_equals ctxt ct (normalize_term nbe_program_sym_tab ctxt (Thm.term_of ct) vs_ty_t deps))); |
|
55757 | 655 |
|
82454 | 656 |
fun oracle nbe_program_sym_tab ctxt vs_ty_t deps ct = |
657 |
raw_oracle (nbe_program_sym_tab, ctxt, vs_ty_t, deps, ct); |
|
31064 | 658 |
|
55757 | 659 |
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
|
660 |
(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
|
661 |
oracle (compile false ctxt program) ctxt')); |
30942
1e246776f876
diagnostic commands now in code_thingol; tuned code of funny continuations
haftmann
parents:
30672
diff
changeset
|
662 |
|
55757 | 663 |
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
|
664 |
(Code_Thingol.dynamic_value ctxt I (fn program => |
63162 | 665 |
normalize_term (compile false ctxt program) ctxt)); |
24155 | 666 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56972
diff
changeset
|
667 |
fun static_conv (ctxt_consts as { ctxt, ... }) = |
55757 | 668 |
let |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
669 |
val conv = Code_Thingol.static_conv_thingol ctxt_consts |
63164 | 670 |
(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
|
671 |
in fn ctxt' => lift_triv_classes_conv ctxt' conv end; |
39399 | 672 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56972
diff
changeset
|
673 |
fun static_value { ctxt, consts } = |
55757 | 674 |
let |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
675 |
val comp = Code_Thingol.static_value { ctxt = ctxt, lift_postproc = I, consts = consts } |
63164 | 676 |
(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
|
677 |
in fn ctxt' => lift_triv_classes_rew ctxt' (comp ctxt') end; |
41184 | 678 |
|
24155 | 679 |
end; |