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