author | nipkow |
Wed, 13 Dec 2017 11:44:11 +0100 | |
changeset 67169 | 1fabca1c2199 |
parent 65062 | dc746d43f40e |
child 67327 | 89be5a4f514b |
permissions | -rw-r--r-- |
39401 | 1 |
(* Title: Tools/Code/code_runtime.ML |
28054 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
4 |
Runtime services building on code generation into implementation language SML. |
28054 | 5 |
*) |
6 |
||
39401 | 7 |
signature CODE_RUNTIME = |
28054 | 8 |
sig |
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
9 |
val target: string |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
10 |
val value: Proof.context -> |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
11 |
(Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string -> |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
12 |
string * string -> 'a |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
13 |
type 'a cookie = (Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string |
55757 | 14 |
val dynamic_value: 'a cookie -> Proof.context -> string option |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
15 |
-> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a option |
55757 | 16 |
val dynamic_value_strict: 'a cookie -> Proof.context -> string option |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
17 |
-> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a |
55757 | 18 |
val dynamic_value_exn: 'a cookie -> Proof.context -> string option |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
19 |
-> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a Exn.result |
55757 | 20 |
val dynamic_holds_conv: Proof.context -> conv |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
21 |
val code_reflect: (string * string list option) list -> string list -> string |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
22 |
-> string option -> theory -> theory |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
23 |
datatype truth = Holds |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
24 |
val put_truth: (unit -> truth) -> Proof.context -> Proof.context |
64987 | 25 |
val mount_computation: Proof.context -> (string * typ) list -> typ |
64988 | 26 |
-> (term -> 'ml) -> ((term -> term) -> 'ml option -> 'a) -> Proof.context -> term -> 'a |
64989 | 27 |
val mount_computation_conv: Proof.context -> (string * typ) list -> typ |
65043
fd753468786f
explicit dynamic context for gap-bridging function
haftmann
parents:
65034
diff
changeset
|
28 |
-> (term -> 'ml) -> (Proof.context -> 'ml -> conv) -> Proof.context -> conv |
64989 | 29 |
val mount_computation_check: Proof.context -> (string * typ) list |
30 |
-> (term -> truth) -> Proof.context -> conv |
|
64987 | 31 |
val polyml_as_definition: (binding * typ) list -> Path.T list -> theory -> theory |
57435 | 32 |
val trace: bool Config.T |
28054 | 33 |
end; |
34 |
||
39401 | 35 |
structure Code_Runtime : CODE_RUNTIME = |
28054 | 36 |
struct |
37 |
||
55150 | 38 |
open Basic_Code_Symbol; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
39 |
|
64957 | 40 |
(** ML compiler as evaluation environment **) |
33992 | 41 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
42 |
(* technical prerequisites *) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
43 |
|
64959 | 44 |
val thisN = "Code_Runtime"; |
64987 | 45 |
val prefix_this = Long_Name.append thisN; |
46 |
val truthN = prefix_this "truth"; |
|
47 |
val HoldsN = prefix_this "Holds"; |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
48 |
|
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
49 |
val target = "Eval"; |
28054 | 50 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
51 |
datatype truth = Holds; |
39422 | 52 |
|
53171 | 53 |
val _ = Theory.setup |
59104 | 54 |
(Code_Target.add_derived_target (target, [(Code_ML.target_SML, I)]) |
55150 | 55 |
#> Code_Target.set_printings (Type_Constructor (@{type_name prop}, |
64959 | 56 |
[(target, SOME (0, (K o K o K) (Code_Printer.str truthN)))])) |
55150 | 57 |
#> Code_Target.set_printings (Constant (@{const_name Code_Generator.holds}, |
64959 | 58 |
[(target, SOME (Code_Printer.plain_const_syntax HoldsN))])) |
59 |
#> Code_Target.add_reserved target thisN |
|
53171 | 60 |
#> fold (Code_Target.add_reserved target) ["oo", "ooo", "oooo", "upto", "downto", "orf", "andf"]); |
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
61 |
(*avoid further pervasive infix names*) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
62 |
|
57435 | 63 |
val trace = Attrib.setup_config_bool @{binding "code_runtime_trace"} (K false); |
40150 | 64 |
|
63163 | 65 |
fun compile_ML verbose code context = |
66 |
(if Config.get_generic context trace then tracing code else (); |
|
63164 | 67 |
Code_Preproc.timed "compiling ML" Context.proof_of |
68 |
(ML_Context.exec (fn () => ML_Compiler0.ML ML_Env.context |
|
63163 | 69 |
{line = 0, file = "generated code", verbose = verbose, |
63164 | 70 |
debug = false} code)) context); |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
71 |
|
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
72 |
fun value ctxt (get, put, put_ml) (prelude, value) = |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
73 |
let |
62889 | 74 |
val code = |
75 |
prelude ^ "\nval _ = Context.put_generic_context (SOME (Context.map_proof (" ^ |
|
76 |
put_ml ^ " (fn () => " ^ value ^ ")) (Context.the_generic_context ())))"; |
|
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
77 |
val ctxt' = ctxt |
64957 | 78 |
|> put (fn () => error ("Bad compilation for " ^ quote put_ml)) |
63163 | 79 |
|> Context.proof_map (compile_ML false code); |
63164 | 80 |
val computator = get ctxt'; |
81 |
in Code_Preproc.timed_exec "running ML" computator ctxt' end; |
|
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
82 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
83 |
|
64957 | 84 |
(* evaluation into ML language values *) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
85 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
86 |
type 'a cookie = (Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
87 |
|
55757 | 88 |
fun reject_vars ctxt t = |
89 |
((Sign.no_frees ctxt o Sign.no_vars ctxt o map_types (K dummyT)) t; t); |
|
39605 | 90 |
|
64957 | 91 |
fun build_compilation_text ctxt some_target program consts = |
92 |
Code_Target.compilation_text ctxt (the_default target some_target) program consts false |
|
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
93 |
#>> (fn ml_modules => space_implode "\n\n" (map snd ml_modules)); |
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
94 |
|
64957 | 95 |
fun run_compilation_text cookie ctxt comp vs_t args = |
55757 | 96 |
let |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
97 |
val (program_code, value_name) = comp vs_t; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
98 |
val value_code = space_implode " " |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
99 |
(value_name :: "()" :: map (enclose "(" ")") args); |
40235
87998864284e
use Exn.interruptible_capture to keep user-code interruptible (Exn.capture not immediately followed by Exn.release here);
wenzelm
parents:
40150
diff
changeset
|
100 |
in Exn.interruptible_capture (value ctxt cookie) (program_code, value_code) end; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
101 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
102 |
fun partiality_as_none e = SOME (Exn.release e) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
103 |
handle General.Match => NONE |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
104 |
| General.Bind => NONE |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
105 |
| General.Fail _ => NONE; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
106 |
|
55757 | 107 |
fun dynamic_value_exn cookie ctxt some_target postproc t args = |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
108 |
let |
55757 | 109 |
val _ = reject_vars ctxt t; |
57435 | 110 |
val _ = if Config.get ctxt trace |
55757 | 111 |
then tracing ("Evaluation of term " ^ quote (Syntax.string_of_term ctxt t)) |
41099 | 112 |
else () |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
113 |
fun comp program _ vs_ty_t deps = |
64957 | 114 |
run_compilation_text cookie ctxt (build_compilation_text ctxt some_target program deps) vs_ty_t args; |
63157
65a81a4ef7f8
clarified naming conventions and code for code evaluation sandwiches
haftmann
parents:
63156
diff
changeset
|
115 |
in Code_Thingol.dynamic_value ctxt (Exn.map_res o postproc) comp t end; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
116 |
|
55757 | 117 |
fun dynamic_value_strict cookie ctxt some_target postproc t args = |
118 |
Exn.release (dynamic_value_exn cookie ctxt some_target postproc t args); |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
119 |
|
55757 | 120 |
fun dynamic_value cookie ctxt some_target postproc t args = |
121 |
partiality_as_none (dynamic_value_exn cookie ctxt some_target postproc t args); |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
122 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
123 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
124 |
(* evaluation for truth or nothing *) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
125 |
|
39820 | 126 |
structure Truth_Result = Proof_Data |
127 |
( |
|
59155 | 128 |
type T = unit -> truth; |
129 |
val empty: T = fn () => raise Fail "Truth_Result"; |
|
130 |
fun init _ = empty; |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
131 |
); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
132 |
val put_truth = Truth_Result.put; |
64987 | 133 |
val truth_cookie = (Truth_Result.get, put_truth, prefix_this "put_truth"); |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
134 |
|
58558
30cc7ee5ac10
basic support for fully static evaluator generation without dynamic compiler invocation
haftmann
parents:
58557
diff
changeset
|
135 |
local |
39605 | 136 |
|
58558
30cc7ee5ac10
basic support for fully static evaluator generation without dynamic compiler invocation
haftmann
parents:
58557
diff
changeset
|
137 |
val reject_vars = fn ctxt => tap (reject_vars ctxt o Thm.term_of); |
41349
0e2a3f22f018
evaluator separating static and dynamic operations
haftmann
parents:
41348
diff
changeset
|
138 |
|
55757 | 139 |
fun check_holds ctxt evaluator vs_t ct = |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
140 |
let |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
141 |
val t = Thm.term_of ct; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
142 |
val _ = if fastype_of t <> propT |
59633 | 143 |
then error ("Not a proposition: " ^ Syntax.string_of_term ctxt t) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
144 |
else (); |
59633 | 145 |
val iff = Thm.cterm_of ctxt (Term.Const (@{const_name Pure.eq}, propT --> propT --> propT)); |
64957 | 146 |
val result = case partiality_as_none (run_compilation_text truth_cookie ctxt evaluator vs_t []) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
147 |
of SOME Holds => true |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
148 |
| _ => false; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
149 |
in |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
150 |
Thm.mk_binop iff ct (if result then @{cprop "PROP Code_Generator.holds"} else ct) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
151 |
end; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
152 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
153 |
val (_, raw_check_holds_oracle) = Context.>>> (Context.map_theory_result |
43619
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
43560
diff
changeset
|
154 |
(Thm.add_oracle (@{binding holds_by_evaluation}, |
55757 | 155 |
fn (ctxt, evaluator, vs_t, ct) => check_holds ctxt evaluator vs_t ct))); |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
156 |
|
56920
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
haftmann
parents:
56618
diff
changeset
|
157 |
fun check_holds_oracle ctxt evaluator vs_ty_t ct = |
d651b944c67e
normalizing of type variables before evaluation with explicit resubstitution function: make nbe work with funny type variables like \<AA>;
haftmann
parents:
56618
diff
changeset
|
158 |
raw_check_holds_oracle (ctxt, evaluator, vs_ty_t, ct); |
41349
0e2a3f22f018
evaluator separating static and dynamic operations
haftmann
parents:
41348
diff
changeset
|
159 |
|
0e2a3f22f018
evaluator separating static and dynamic operations
haftmann
parents:
41348
diff
changeset
|
160 |
in |
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
161 |
|
55757 | 162 |
fun dynamic_holds_conv ctxt = Code_Thingol.dynamic_conv ctxt |
56969
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56920
diff
changeset
|
163 |
(fn program => fn vs_t => fn deps => |
64957 | 164 |
check_holds_oracle ctxt (build_compilation_text ctxt NONE program deps) vs_t) |
55757 | 165 |
o reject_vars ctxt; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
166 |
|
41349
0e2a3f22f018
evaluator separating static and dynamic operations
haftmann
parents:
41348
diff
changeset
|
167 |
end; (*local*) |
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
168 |
|
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
169 |
|
64957 | 170 |
(** generator for computations -- partial implementations of the universal morphism from Isabelle to ML terms **) |
171 |
||
172 |
(* auxiliary *) |
|
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
173 |
|
64957 | 174 |
val generated_computationN = "Generated_Computation"; |
175 |
||
176 |
||
177 |
(* possible type signatures of constants *) |
|
64943
c5618df67c2a
explicit a-priori detection of unsuitable terms for computations
haftmann
parents:
64940
diff
changeset
|
178 |
|
64990 | 179 |
fun typ_signatures' T = |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
180 |
let |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
181 |
val (Ts, T') = strip_type T; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
182 |
in map_range (fn n => (drop n Ts ---> T', take n Ts)) (length Ts + 1) end; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
183 |
|
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
184 |
fun typ_signatures cTs = |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
185 |
let |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
186 |
fun add (c, T) = |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
187 |
fold (fn (T, Ts) => Typtab.map_default (T, []) (cons (c, Ts))) |
64990 | 188 |
(typ_signatures' T); |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
189 |
in |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
190 |
Typtab.empty |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
191 |
|> fold add cTs |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
192 |
|> Typtab.lookup_list |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
193 |
end; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
194 |
|
64957 | 195 |
|
196 |
(* name mangling *) |
|
197 |
||
198 |
local |
|
199 |
||
200 |
fun tycos_of (Type (tyco, Ts)) = maps tycos_of Ts @ [tyco] |
|
201 |
| tycos_of _ = []; |
|
202 |
||
203 |
val ml_name_of = Name.desymbolize NONE o Long_Name.base_name; |
|
204 |
||
205 |
in |
|
206 |
||
64959 | 207 |
val covered_constsN = "covered_consts"; |
208 |
||
64957 | 209 |
fun of_term_for_typ Ts = |
210 |
let |
|
211 |
val names = Ts |
|
212 |
|> map (suffix "_of_term" o space_implode "_" o map ml_name_of o tycos_of) |
|
213 |
|> Name.variant_list []; |
|
214 |
in the o AList.lookup (op =) (Ts ~~ names) end; |
|
215 |
||
216 |
fun eval_for_const ctxt cTs = |
|
217 |
let |
|
218 |
fun symbol_list (c, T) = c :: maps tycos_of (Sign.const_typargs (Proof_Context.theory_of ctxt) (c, T)) |
|
219 |
val names = cTs |
|
220 |
|> map (prefix "eval_" o space_implode "_" o map ml_name_of o symbol_list) |
|
221 |
|> Name.variant_list []; |
|
222 |
in the o AList.lookup (op =) (cTs ~~ names) end; |
|
223 |
||
224 |
end; |
|
225 |
||
226 |
||
227 |
(* checks for input terms *) |
|
228 |
||
229 |
fun monomorphic T = fold_atyps ((K o K) false) T true; |
|
230 |
||
231 |
fun check_typ ctxt T t = |
|
232 |
Syntax.check_term ctxt (Type.constraint T t); |
|
233 |
||
234 |
fun check_computation_input ctxt cTs t = |
|
235 |
let |
|
236 |
fun check t = check_comb (strip_comb t) |
|
237 |
and check_comb (t as Abs _, _) = |
|
238 |
error ("Bad term, contains abstraction: " ^ Syntax.string_of_term ctxt t) |
|
239 |
| check_comb (t as Const (cT as (c, T)), ts) = |
|
240 |
let |
|
241 |
val _ = if not (member (op =) cTs cT) |
|
242 |
then error ("Bad term, computation cannot proceed on constant " ^ Syntax.string_of_term ctxt t) |
|
243 |
else (); |
|
244 |
val _ = if not (monomorphic T) |
|
245 |
then error ("Bad term, contains polymorphic constant " ^ Syntax.string_of_term ctxt t) |
|
246 |
else (); |
|
247 |
val _ = map check ts; |
|
248 |
in () end; |
|
249 |
val _ = check t; |
|
250 |
in t end; |
|
251 |
||
252 |
||
253 |
(* code generation for of the universal morphism *) |
|
254 |
||
64959 | 255 |
val print_const = ML_Syntax.print_pair ML_Syntax.print_string ML_Syntax.print_typ; |
256 |
||
64990 | 257 |
fun print_of_term_funs { typ_signatures_for, eval_for_const, of_term_for_typ } Ts = |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
258 |
let |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
259 |
val var_names = map_range (fn n => "t" ^ string_of_int (n + 1)); |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
260 |
fun print_lhs c xs = "Const (" ^ quote c ^ ", _)" |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
261 |
|> fold (fn x => fn s => s ^ " $ " ^ x) xs |
64943
c5618df67c2a
explicit a-priori detection of unsuitable terms for computations
haftmann
parents:
64940
diff
changeset
|
262 |
|> enclose "(" ")"; |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
263 |
fun print_rhs c Ts T xs = eval_for_const (c, Ts ---> T) |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
264 |
|> fold2 (fn T' => fn x => fn s => |
64943
c5618df67c2a
explicit a-priori detection of unsuitable terms for computations
haftmann
parents:
64940
diff
changeset
|
265 |
s ^ (" (" ^ of_term_for_typ T' ^ " " ^ x ^ ")")) Ts xs |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
266 |
fun print_eq T (c, Ts) = |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
267 |
let |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
268 |
val xs = var_names (length Ts); |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
269 |
in print_lhs c xs ^ " = " ^ print_rhs c Ts T xs end; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
270 |
fun print_eqs T = |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
271 |
let |
64990 | 272 |
val typ_signs = typ_signatures_for T; |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
273 |
val name = of_term_for_typ T; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
274 |
in |
64943
c5618df67c2a
explicit a-priori detection of unsuitable terms for computations
haftmann
parents:
64940
diff
changeset
|
275 |
map (print_eq T) typ_signs |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
276 |
|> map (prefix (name ^ " ")) |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
277 |
|> space_implode "\n | " |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
278 |
end; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
279 |
in |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
280 |
map print_eqs Ts |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
281 |
|> space_implode "\nand " |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
282 |
|> prefix "fun " |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
283 |
end; |
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
284 |
|
64990 | 285 |
fun print_computation_code ctxt compiled_value [] requested_Ts = |
286 |
if null requested_Ts then ("", []) |
|
287 |
else error ("No equation available for requested type " |
|
288 |
^ Syntax.string_of_typ ctxt (hd requested_Ts)) |
|
64959 | 289 |
| print_computation_code ctxt compiled_value cTs requested_Ts = |
290 |
let |
|
291 |
val proper_cTs = map_filter I cTs; |
|
64990 | 292 |
val typ_signatures_for = typ_signatures proper_cTs; |
64959 | 293 |
fun add_typ T Ts = |
294 |
if member (op =) Ts T |
|
295 |
then Ts |
|
64990 | 296 |
else case typ_signatures_for T of |
297 |
[] => error ("No equation available for requested type " |
|
298 |
^ Syntax.string_of_typ ctxt T) |
|
299 |
| typ_signs => |
|
300 |
Ts |
|
301 |
|> cons T |
|
302 |
|> fold (fold add_typ o snd) typ_signs; |
|
64959 | 303 |
val required_Ts = fold add_typ requested_Ts []; |
304 |
val of_term_for_typ' = of_term_for_typ required_Ts; |
|
305 |
val eval_for_const' = eval_for_const ctxt proper_cTs; |
|
306 |
val eval_for_const'' = the_default "_" o Option.map eval_for_const'; |
|
307 |
val eval_tuple = enclose "(" ")" (commas (map eval_for_const' proper_cTs)); |
|
308 |
fun mk_abs s = "fn " ^ s ^ " => "; |
|
309 |
val eval_abs = space_implode "" |
|
310 |
(map (mk_abs o eval_for_const'') cTs); |
|
311 |
val of_term_code = print_of_term_funs { |
|
64990 | 312 |
typ_signatures_for = typ_signatures_for, |
64959 | 313 |
eval_for_const = eval_for_const', |
314 |
of_term_for_typ = of_term_for_typ' } required_Ts; |
|
315 |
val of_term_names = map (Long_Name.append generated_computationN |
|
316 |
o of_term_for_typ') requested_Ts; |
|
317 |
in |
|
318 |
cat_lines [ |
|
319 |
"structure " ^ generated_computationN ^ " =", |
|
320 |
"struct", |
|
321 |
"", |
|
322 |
"val " ^ covered_constsN ^ " = " ^ ML_Syntax.print_list print_const proper_cTs ^ ";", |
|
323 |
"", |
|
324 |
"val " ^ eval_tuple ^ " = " ^ compiled_value ^ " ()", |
|
325 |
" (" ^ eval_abs, |
|
326 |
" " ^ eval_tuple ^ ");", |
|
327 |
"", |
|
328 |
of_term_code, |
|
329 |
"", |
|
330 |
"end" |
|
331 |
] |> rpair of_term_names |
|
332 |
end; |
|
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
333 |
|
64993 | 334 |
|
335 |
(* dedicated preprocessor for computations *) |
|
336 |
||
337 |
structure Computation_Preproc_Data = Theory_Data |
|
338 |
( |
|
339 |
type T = thm list; |
|
340 |
val empty = []; |
|
341 |
val extend = I; |
|
342 |
val merge = Library.merge Thm.eq_thm_prop; |
|
343 |
); |
|
344 |
||
345 |
local |
|
346 |
||
347 |
fun add thm thy = |
|
348 |
let |
|
349 |
val thms = Simplifier.mksimps (Proof_Context.init_global thy) thm; |
|
350 |
in |
|
351 |
thy |
|
352 |
|> Computation_Preproc_Data.map (union Thm.eq_thm_prop thms) |
|
353 |
end; |
|
354 |
||
355 |
fun get ctxt = Computation_Preproc_Data.get (Proof_Context.theory_of ctxt); |
|
356 |
||
357 |
in |
|
358 |
||
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
359 |
fun preprocess_conv { ctxt } = |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
360 |
let |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
361 |
val rules = get ctxt; |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
362 |
in fn ctxt' => Raw_Simplifier.rewrite ctxt' false rules end; |
64993 | 363 |
|
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
364 |
fun preprocess_term { ctxt } = |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
365 |
let |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
366 |
val rules = map (Logic.dest_equals o Thm.plain_prop_of) (get ctxt); |
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
367 |
in fn ctxt' => Pattern.rewrite_term (Proof_Context.theory_of ctxt') rules [] end; |
64993 | 368 |
|
369 |
val _ = Theory.setup |
|
370 |
(Attrib.setup @{binding code_computation_unfold} |
|
371 |
(Scan.succeed (Thm.declaration_attribute (fn thm => Context.mapping (add thm) I))) |
|
372 |
"preprocessing equations for computation"); |
|
373 |
||
374 |
end; |
|
375 |
||
376 |
||
377 |
(* mounting computations *) |
|
378 |
||
64991
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
379 |
fun prechecked_computation T raw_computation ctxt = |
64987 | 380 |
check_typ ctxt T |
381 |
#> reject_vars ctxt |
|
64991
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
382 |
#> raw_computation ctxt; |
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
383 |
|
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
384 |
fun prechecked_conv T raw_conv ctxt = |
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
385 |
tap (check_typ ctxt T o reject_vars ctxt o Thm.term_of) |
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
386 |
#> raw_conv ctxt; |
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
387 |
|
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
388 |
fun checked_computation cTs raw_computation ctxt = |
d2c79b16e133
variables and type must be checked before entering evaluation sandwich
haftmann
parents:
64990
diff
changeset
|
389 |
check_computation_input ctxt cTs |
65005 | 390 |
#> Exn.interruptible_capture raw_computation |
64988 | 391 |
#> partiality_as_none; |
64987 | 392 |
|
64957 | 393 |
fun mount_computation ctxt cTs T raw_computation lift_postproc = |
64993 | 394 |
let |
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
395 |
val preprocess = preprocess_term { ctxt = ctxt }; |
64993 | 396 |
val computation = prechecked_computation T (Code_Preproc.static_value |
397 |
{ ctxt = ctxt, lift_postproc = lift_postproc, consts = [] } |
|
398 |
(K (checked_computation cTs raw_computation))); |
|
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
399 |
in fn ctxt' => preprocess ctxt' #> computation ctxt' end; |
64943
c5618df67c2a
explicit a-priori detection of unsuitable terms for computations
haftmann
parents:
64940
diff
changeset
|
400 |
|
64989 | 401 |
fun mount_computation_conv ctxt cTs T raw_computation conv = |
64993 | 402 |
let |
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
403 |
val preprocess = preprocess_conv { ctxt = ctxt }; |
64993 | 404 |
val computation_conv = prechecked_conv T (Code_Preproc.static_conv |
405 |
{ ctxt = ctxt, consts = [] } |
|
406 |
(K (fn ctxt' => fn t => |
|
407 |
case checked_computation cTs raw_computation ctxt' t of |
|
65043
fd753468786f
explicit dynamic context for gap-bridging function
haftmann
parents:
65034
diff
changeset
|
408 |
SOME x => conv ctxt' x |
64993 | 409 |
| NONE => Conv.all_conv))); |
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
410 |
in fn ctxt' => preprocess ctxt' then_conv computation_conv ctxt' end; |
64989 | 411 |
|
412 |
local |
|
413 |
||
414 |
fun holds ct = Thm.mk_binop @{cterm "Pure.eq :: prop \<Rightarrow> prop \<Rightarrow> prop"} |
|
415 |
ct @{cprop "PROP Code_Generator.holds"}; |
|
416 |
||
417 |
val (_, holds_oracle) = Context.>>> (Context.map_theory_result |
|
418 |
(Thm.add_oracle (@{binding holds}, holds))); |
|
419 |
||
420 |
in |
|
421 |
||
422 |
fun mount_computation_check ctxt cTs raw_computation = |
|
423 |
mount_computation_conv ctxt cTs @{typ prop} raw_computation |
|
65043
fd753468786f
explicit dynamic context for gap-bridging function
haftmann
parents:
65034
diff
changeset
|
424 |
((K o K) holds_oracle); |
64989 | 425 |
|
426 |
end; |
|
427 |
||
64959 | 428 |
|
429 |
(** variants of universal runtime code generation **) |
|
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
430 |
|
64959 | 431 |
(*FIXME consolidate variants*) |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
432 |
|
64959 | 433 |
fun runtime_code'' ctxt module_name program tycos consts = |
33992 | 434 |
let |
55757 | 435 |
val thy = Proof_Context.theory_of ctxt; |
48568 | 436 |
val (ml_modules, target_names) = |
55757 | 437 |
Code_Target.produce_code_for ctxt |
55683 | 438 |
target NONE module_name [] program false (map Constant consts @ map Type_Constructor tycos); |
48568 | 439 |
val ml_code = space_implode "\n\n" (map snd ml_modules); |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
53171
diff
changeset
|
440 |
val (consts', tycos') = chop (length consts) target_names; |
42359 | 441 |
val consts_map = map2 (fn const => |
442 |
fn NONE => |
|
443 |
error ("Constant " ^ (quote o Code.string_of_const thy) const ^ |
|
444 |
"\nhas a user-defined serialization") |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
53171
diff
changeset
|
445 |
| SOME const' => (const, const')) consts consts' |
42359 | 446 |
val tycos_map = map2 (fn tyco => |
447 |
fn NONE => |
|
55304 | 448 |
error ("Type " ^ quote (Proof_Context.markup_type ctxt tyco) ^ |
42359 | 449 |
"\nhas a user-defined serialization") |
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
53171
diff
changeset
|
450 |
| SOME tyco' => (tyco, tyco')) tycos tycos'; |
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
451 |
in (ml_code, (tycos_map, consts_map)) end; |
28054 | 452 |
|
64959 | 453 |
fun runtime_code' ctxt some_module_name named_tycos named_consts computation_Ts program evals vs_ty_evals deps = |
454 |
let |
|
455 |
val thy = Proof_Context.theory_of ctxt; |
|
456 |
fun the_const (Const cT) = cT |
|
457 |
| the_const t = error ("No constant after preprocessing: " ^ Syntax.string_of_term ctxt t) |
|
458 |
val raw_computation_cTs = case evals of |
|
459 |
Abs (_, _, t) => (map the_const o snd o strip_comb) t |
|
460 |
| _ => error ("Bad term after preprocessing: " ^ Syntax.string_of_term ctxt evals); |
|
461 |
val computation_cTs = fold_rev (fn cT => fn cTs => |
|
462 |
(if member (op =) cTs (SOME cT) then NONE else SOME cT) :: cTs) raw_computation_cTs []; |
|
463 |
val consts' = fold (fn NONE => I | SOME (c, _) => insert (op =) c) |
|
464 |
computation_cTs named_consts; |
|
465 |
val program' = Code_Thingol.consts_program ctxt consts'; |
|
466 |
(*FIXME insufficient interfaces require double invocation of code generator*) |
|
65034
1846c4551153
more complete program generation in presence of dictionaries
haftmann
parents:
65005
diff
changeset
|
467 |
val program'' = Code_Symbol.Graph.merge (K true) (program, program'); |
64959 | 468 |
val ((ml_modules, compiled_value), deresolve) = |
65034
1846c4551153
more complete program generation in presence of dictionaries
haftmann
parents:
65005
diff
changeset
|
469 |
Code_Target.compilation_text' ctxt target some_module_name program'' |
64959 | 470 |
(map Code_Symbol.Type_Constructor named_tycos @ map Code_Symbol.Constant consts' @ deps) true vs_ty_evals; |
471 |
(*FIXME constrain signature*) |
|
472 |
fun deresolve_const c = case (deresolve o Code_Symbol.Constant) c of |
|
473 |
NONE => error ("Constant " ^ (quote o Code.string_of_const thy) c ^ |
|
474 |
"\nhas a user-defined serialization") |
|
475 |
| SOME c' => c'; |
|
476 |
fun deresolve_tyco tyco = case (deresolve o Code_Symbol.Type_Constructor) tyco of |
|
477 |
NONE => error ("Type " ^ quote (Proof_Context.markup_type ctxt tyco) ^ |
|
478 |
"\nhas a user-defined serialization") |
|
479 |
| SOME c' => c'; |
|
480 |
val tyco_names = map deresolve_const named_tycos; |
|
481 |
val const_names = map deresolve_const named_consts; |
|
482 |
val generated_code = space_implode "\n\n" (map snd ml_modules); |
|
483 |
val (of_term_code, of_term_names) = |
|
484 |
print_computation_code ctxt compiled_value computation_cTs computation_Ts; |
|
485 |
val compiled_computation = generated_code ^ "\n" ^ of_term_code; |
|
486 |
in |
|
487 |
compiled_computation |
|
488 |
|> rpair { tyco_map = named_tycos ~~ tyco_names, |
|
489 |
const_map = named_consts ~~ const_names, |
|
490 |
of_term_map = computation_Ts ~~ of_term_names } |
|
491 |
end; |
|
492 |
||
493 |
fun funs_of_maps { tyco_map, const_map, of_term_map } = |
|
494 |
{ name_for_tyco = the o AList.lookup (op =) tyco_map, |
|
495 |
name_for_const = the o AList.lookup (op =) const_map, |
|
496 |
of_term_for_typ = the o AList.lookup (op =) of_term_map }; |
|
497 |
||
498 |
fun runtime_code ctxt some_module_name named_tycos named_consts computation_Ts program evals vs_ty_evals deps = |
|
499 |
runtime_code' ctxt some_module_name named_tycos named_consts computation_Ts program evals vs_ty_evals deps |
|
500 |
||> funs_of_maps; |
|
501 |
||
502 |
||
503 |
(** code and computation antiquotations **) |
|
504 |
||
64989 | 505 |
local |
64959 | 506 |
|
64989 | 507 |
val mount_computationN = prefix_this "mount_computation"; |
508 |
val mount_computation_convN = prefix_this "mount_computation_conv"; |
|
509 |
val mount_computation_checkN = prefix_this "mount_computation_check"; |
|
28054 | 510 |
|
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
511 |
structure Code_Antiq_Data = Proof_Data |
28054 | 512 |
( |
64955 | 513 |
type T = { named_consts: string list, |
64959 | 514 |
computation_Ts: typ list, computation_cTs: (string * typ) list, |
515 |
position_index: int, |
|
516 |
generated_code: (string * { |
|
517 |
name_for_tyco: string -> string, |
|
518 |
name_for_const: string -> string, |
|
519 |
of_term_for_typ: typ -> string |
|
520 |
}) lazy |
|
64955 | 521 |
}; |
522 |
val empty: T = { named_consts = [], |
|
64959 | 523 |
computation_Ts = [], computation_cTs = [], |
524 |
position_index = 0, |
|
525 |
generated_code = Lazy.lazy (fn () => raise Fail "empty") |
|
64955 | 526 |
}; |
59150 | 527 |
fun init _ = empty; |
28054 | 528 |
); |
529 |
||
64959 | 530 |
val current_position_index = #position_index o Code_Antiq_Data.get; |
64955 | 531 |
|
64959 | 532 |
fun register { named_consts, computation_Ts, computation_cTs } ctxt = |
64955 | 533 |
let |
64959 | 534 |
val data = Code_Antiq_Data.get ctxt; |
535 |
val named_consts' = union (op =) named_consts (#named_consts data); |
|
536 |
val computation_Ts' = union (op =) computation_Ts (#computation_Ts data); |
|
537 |
val computation_cTs' = union (op =) computation_cTs (#computation_cTs data); |
|
538 |
val position_index' = #position_index data + 1; |
|
539 |
fun generated_code' () = |
|
540 |
let |
|
541 |
val evals = Abs ("eval", map snd computation_cTs' ---> |
|
64993 | 542 |
TFree (Name.aT, []), list_comb (Bound 0, map Const computation_cTs')) |
64995
a7af4045f873
preprocessing rules must always stem from static context
haftmann
parents:
64993
diff
changeset
|
543 |
|> preprocess_term { ctxt = ctxt } ctxt |
64959 | 544 |
in Code_Thingol.dynamic_value ctxt |
545 |
(K I) (runtime_code ctxt NONE [] named_consts' computation_Ts') evals |
|
546 |
end; |
|
64955 | 547 |
in |
548 |
ctxt |
|
549 |
|> Code_Antiq_Data.put { |
|
64959 | 550 |
named_consts = named_consts', |
551 |
computation_Ts = computation_Ts', |
|
552 |
computation_cTs = computation_cTs', |
|
553 |
position_index = position_index', |
|
554 |
generated_code = Lazy.lazy generated_code' |
|
64955 | 555 |
} |
556 |
end; |
|
30962 | 557 |
|
64959 | 558 |
fun register_const const = |
559 |
register { named_consts = [const], |
|
560 |
computation_Ts = [], |
|
561 |
computation_cTs = [] }; |
|
562 |
||
563 |
fun register_computation cTs T = |
|
564 |
register { named_consts = [], |
|
565 |
computation_Ts = [T], |
|
566 |
computation_cTs = cTs }; |
|
567 |
||
568 |
fun print body_code_for ctxt ctxt' = |
|
30962 | 569 |
let |
64959 | 570 |
val position_index = current_position_index ctxt; |
571 |
val (code, name_ofs) = (Lazy.force o #generated_code o Code_Antiq_Data.get) ctxt'; |
|
572 |
val context_code = if position_index = 0 then code else ""; |
|
64989 | 573 |
val body_code = body_code_for name_ofs (ML_Context.struct_name ctxt'); |
64958 | 574 |
in (context_code, body_code) end; |
28054 | 575 |
|
64959 | 576 |
fun print_code ctxt const = |
64989 | 577 |
print (fn { name_for_const, ... } => fn prfx => |
64959 | 578 |
Long_Name.append prfx (name_for_const const)) ctxt; |
579 |
||
64989 | 580 |
fun print_computation kind ctxt T = |
581 |
print (fn { of_term_for_typ, ... } => fn prfx => |
|
64959 | 582 |
space_implode " " [ |
64989 | 583 |
kind, |
64959 | 584 |
"(Context.proof_of (Context.the_generic_context ()))", |
585 |
Long_Name.implode [prfx, generated_computationN, covered_constsN], |
|
586 |
(ML_Syntax.atomic o ML_Syntax.print_typ) T, |
|
587 |
Long_Name.append prfx (of_term_for_typ T) |
|
588 |
]) ctxt; |
|
589 |
||
64989 | 590 |
fun print_computation_check ctxt = |
591 |
print (fn { of_term_for_typ, ... } => fn prfx => |
|
592 |
space_implode " " [ |
|
593 |
mount_computation_checkN, |
|
594 |
"(Context.proof_of (Context.the_generic_context ()))", |
|
595 |
Long_Name.implode [prfx, generated_computationN, covered_constsN], |
|
596 |
Long_Name.append prfx (of_term_for_typ @{typ prop}) |
|
597 |
]) ctxt; |
|
598 |
||
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
599 |
|
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
600 |
fun add_all_constrs ctxt (dT as Type (tyco, Ts)) = |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
601 |
let |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
602 |
val ((vs, constrs), _) = Code.get_type (Proof_Context.theory_of ctxt) tyco; |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
603 |
val subst_TFree = the o AList.lookup (op =) (map fst vs ~~ Ts); |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
604 |
val cs = map (fn (c, (_, Ts')) => |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
605 |
(c, (map o map_atyps) (fn TFree (v, _) => subst_TFree v) Ts' |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
606 |
---> dT)) constrs; |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
607 |
in |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
608 |
union (op =) cs |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
609 |
#> fold (add_all_constrs ctxt) Ts |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
610 |
end; |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
611 |
|
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
612 |
fun prep_spec ctxt (raw_ts, raw_dTs) = |
64989 | 613 |
let |
614 |
val ts = map (Syntax.check_term ctxt) raw_ts; |
|
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
615 |
val dTs = map (Syntax.check_typ ctxt) raw_dTs; |
64989 | 616 |
in |
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
617 |
[] |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
618 |
|> (fold o fold_aterms) |
64989 | 619 |
(fn (t as Const (cT as (_, T))) => |
620 |
if not (monomorphic T) then error ("Polymorphic constant: " ^ Syntax.string_of_term ctxt t) |
|
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
621 |
else insert (op =) cT | _ => I) ts |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
622 |
|> fold (fn dT => |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
623 |
if not (monomorphic dT) then error ("Polymorphic datatype: " ^ Syntax.string_of_typ ctxt dT) |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
624 |
else add_all_constrs ctxt dT) dTs |
64989 | 625 |
end; |
626 |
||
28054 | 627 |
in |
628 |
||
56069
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
wenzelm
parents:
55757
diff
changeset
|
629 |
fun ml_code_antiq raw_const ctxt = |
28054 | 630 |
let |
53169
e2d31807cbf6
clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents:
52435
diff
changeset
|
631 |
val thy = Proof_Context.theory_of ctxt; |
e2d31807cbf6
clarified type ML_Context.antiq: context parser maintains compilation context, declaration is applied to final context;
wenzelm
parents:
52435
diff
changeset
|
632 |
val const = Code.check_const thy raw_const; |
64959 | 633 |
in (print_code ctxt const, register_const const ctxt) end; |
634 |
||
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
635 |
fun gen_ml_computation_antiq kind (raw_T, raw_spec) ctxt = |
64959 | 636 |
let |
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
637 |
val cTs = prep_spec ctxt raw_spec; |
64959 | 638 |
val T = Syntax.check_typ ctxt raw_T; |
639 |
val _ = if not (monomorphic T) |
|
640 |
then error ("Polymorphic type: " ^ Syntax.string_of_typ ctxt T) |
|
641 |
else (); |
|
64989 | 642 |
in (print_computation kind ctxt T, register_computation cTs T ctxt) end; |
643 |
||
644 |
val ml_computation_antiq = gen_ml_computation_antiq mount_computationN; |
|
645 |
||
646 |
val ml_computation_conv_antiq = gen_ml_computation_antiq mount_computation_convN; |
|
647 |
||
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
648 |
fun ml_computation_check_antiq raw_spec ctxt = |
64989 | 649 |
let |
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
650 |
val cTs = insert (op =) (dest_Const @{const holds}) (prep_spec ctxt raw_spec); |
64989 | 651 |
in (print_computation_check ctxt, register_computation cTs @{typ prop} ctxt) end; |
30962 | 652 |
|
28054 | 653 |
end; (*local*) |
654 |
||
655 |
||
58558
30cc7ee5ac10
basic support for fully static evaluator generation without dynamic compiler invocation
haftmann
parents:
58557
diff
changeset
|
656 |
(** reflection support **) |
36470 | 657 |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
658 |
fun check_datatype thy tyco some_consts = |
36470 | 659 |
let |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
660 |
val declared_constrs = (map fst o snd o fst o Code.get_type thy) tyco; |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
661 |
val constrs = case some_consts |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
662 |
of SOME [] => [] |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
663 |
| SOME consts => |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
664 |
let |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
665 |
val missing_constrs = subtract (op =) consts declared_constrs; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
666 |
val _ = if null missing_constrs then [] |
45430 | 667 |
else error ("Missing constructor(s) " ^ commas_quote missing_constrs |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
668 |
^ " for datatype " ^ quote tyco); |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
669 |
val false_constrs = subtract (op =) declared_constrs consts; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
670 |
val _ = if null false_constrs then [] |
45430 | 671 |
else error ("Non-constructor(s) " ^ commas_quote false_constrs |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
672 |
^ " for datatype " ^ quote tyco) |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
673 |
in consts end |
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
674 |
| NONE => declared_constrs; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
675 |
in (tyco, constrs) end; |
36470 | 676 |
|
677 |
fun add_eval_tyco (tyco, tyco') thy = |
|
678 |
let |
|
679 |
val k = Sign.arity_number thy tyco; |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
680 |
fun pr pr' _ [] = tyco' |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
681 |
| pr pr' _ [ty] = |
36470 | 682 |
Code_Printer.concat [pr' Code_Printer.BR ty, tyco'] |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
683 |
| pr pr' _ tys = |
36470 | 684 |
Code_Printer.concat [Code_Printer.enum "," "(" ")" (map (pr' Code_Printer.BR) tys), tyco'] |
685 |
in |
|
686 |
thy |
|
55150 | 687 |
|> Code_Target.set_printings (Type_Constructor (tyco, [(target, SOME (k, pr))])) |
36470 | 688 |
end; |
689 |
||
36514 | 690 |
fun add_eval_constr (const, const') thy = |
691 |
let |
|
692 |
val k = Code.args_number thy const; |
|
693 |
fun pr pr' fxy ts = Code_Printer.brackify fxy |
|
38922 | 694 |
(const' :: the_list (Code_Printer.tuplify pr' Code_Printer.BR (map fst ts))); |
36514 | 695 |
in |
696 |
thy |
|
55150 | 697 |
|> Code_Target.set_printings (Constant (const, |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52434
diff
changeset
|
698 |
[(target, SOME (Code_Printer.simple_const_syntax (k, pr)))])) |
36514 | 699 |
end; |
700 |
||
55150 | 701 |
fun add_eval_const (const, const') = Code_Target.set_printings (Constant |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52434
diff
changeset
|
702 |
(const, [(target, SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const')))])); |
36470 | 703 |
|
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
704 |
fun process_reflection (code, (tyco_map, (constr_map, const_map))) module_name NONE thy = |
38935
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
705 |
thy |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
706 |
|> Code_Target.add_reserved target module_name |
63163 | 707 |
|> Context.theory_map (compile_ML true code) |
38935
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
708 |
|> fold (add_eval_tyco o apsnd Code_Printer.str) tyco_map |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
709 |
|> fold (add_eval_constr o apsnd Code_Printer.str) constr_map |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
710 |
|> fold (add_eval_const o apsnd Code_Printer.str) const_map |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
711 |
| process_reflection (code, _) _ (SOME file_name) thy = |
36470 | 712 |
let |
37950
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
wenzelm
parents:
37744
diff
changeset
|
713 |
val preamble = |
46737 | 714 |
"(* Generated from " ^ |
56208 | 715 |
Path.implode (Resources.thy_path (Path.basic (Context.theory_name thy))) ^ |
46737 | 716 |
"; DO NOT EDIT! *)"; |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
717 |
val _ = File.write (Path.explode file_name) (preamble ^ "\n\n" ^ code); |
36470 | 718 |
in |
719 |
thy |
|
720 |
end; |
|
721 |
||
722 |
fun gen_code_reflect prep_type prep_const raw_datatypes raw_functions module_name some_file thy = |
|
723 |
let |
|
55757 | 724 |
val ctxt = Proof_Context.init_global thy; |
36470 | 725 |
val datatypes = map (fn (raw_tyco, raw_cos) => |
55757 | 726 |
(prep_type ctxt raw_tyco, (Option.map o map) (prep_const thy) raw_cos)) raw_datatypes; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
727 |
val (tycos, constrs) = map_split (uncurry (check_datatype thy)) datatypes |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
728 |
|> apsnd flat; |
36470 | 729 |
val functions = map (prep_const thy) raw_functions; |
58557 | 730 |
val consts = constrs @ functions; |
63159 | 731 |
val program = Code_Thingol.consts_program ctxt consts; |
64959 | 732 |
val result = runtime_code'' ctxt module_name program tycos consts |
36514 | 733 |
|> (apsnd o apsnd) (chop (length constrs)); |
36470 | 734 |
in |
735 |
thy |
|
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
736 |
|> process_reflection result module_name some_file |
36470 | 737 |
end; |
738 |
||
39422 | 739 |
val code_reflect = gen_code_reflect Code_Target.cert_tyco (K I); |
36470 | 740 |
val code_reflect_cmd = gen_code_reflect Code_Target.read_tyco Code.read_const; |
741 |
||
742 |
||
28054 | 743 |
(** Isar setup **) |
744 |
||
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
745 |
local |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
746 |
|
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
747 |
val parse_consts_spec = |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
748 |
Scan.optional (Scan.lift (Args.$$$ "terms" -- Args.colon) |-- Scan.repeat1 Args.term) [] |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
749 |
-- Scan.optional (Scan.lift (Args.$$$ "datatypes" -- Args.colon) |-- Scan.repeat1 Args.typ) []; |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
750 |
|
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
751 |
in |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
752 |
|
64993 | 753 |
val _ = Theory.setup |
754 |
(ML_Antiquotation.declaration @{binding code} |
|
755 |
Args.term (K ml_code_antiq) |
|
756 |
#> ML_Antiquotation.declaration @{binding computation} |
|
757 |
(Args.typ -- parse_consts_spec) (K ml_computation_antiq) |
|
758 |
#> ML_Antiquotation.declaration @{binding computation_conv} |
|
759 |
(Args.typ -- parse_consts_spec) (K ml_computation_conv_antiq) |
|
760 |
#> ML_Antiquotation.declaration @{binding computation_check} |
|
761 |
parse_consts_spec (K ml_computation_check_antiq)); |
|
64989 | 762 |
|
64992
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
763 |
end; |
41e2c3617582
extended syntax allows to include datatype constructors directly in computations
haftmann
parents:
64991
diff
changeset
|
764 |
|
36470 | 765 |
local |
766 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
767 |
val parse_datatype = |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
768 |
Parse.name -- Scan.optional (@{keyword "="} |-- |
40711
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
haftmann
parents:
40422
diff
changeset
|
769 |
(((Parse.sym_ident || Parse.string) >> (fn "_" => NONE | _ => Scan.fail ())) |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
770 |
|| ((Parse.term ::: (Scan.repeat (@{keyword "|"} |-- Parse.term))) >> SOME))) (SOME []); |
36470 | 771 |
|
772 |
in |
|
773 |
||
774 |
val _ = |
|
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59633
diff
changeset
|
775 |
Outer_Syntax.command @{command_keyword code_reflect} |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
776 |
"enrich runtime environment with generated code" |
63174
57c0d60e491c
do not export abstract constructors in code_reflect
haftmann
parents:
63164
diff
changeset
|
777 |
(Parse.name -- Scan.optional (@{keyword "datatypes"} |-- Parse.!!! (parse_datatype |
46949 | 778 |
::: Scan.repeat (@{keyword "and"} |-- parse_datatype))) [] |
52434 | 779 |
-- Scan.optional (@{keyword "functions"} |-- Parse.!!! (Scan.repeat1 Parse.name)) [] |
780 |
-- Scan.option (@{keyword "file"} |-- Parse.!!! Parse.name) |
|
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
781 |
>> (fn (((module_name, raw_datatypes), raw_functions), some_file) => Toplevel.theory |
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
782 |
(code_reflect_cmd raw_datatypes raw_functions module_name some_file))); |
36470 | 783 |
|
784 |
end; (*local*) |
|
785 |
||
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
786 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
787 |
(** using external SML files as substitute for proper definitions -- only for polyml! **) |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
788 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
789 |
local |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
790 |
|
39820 | 791 |
structure Loaded_Values = Theory_Data |
792 |
( |
|
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
793 |
type T = string list |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
794 |
val empty = [] |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41376
diff
changeset
|
795 |
val extend = I |
39820 | 796 |
fun merge data : T = Library.merge (op =) data |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
797 |
); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
798 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
799 |
fun notify_val (string, value) = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
800 |
let |
62495 | 801 |
val _ = #enterVal ML_Env.name_space (string, value); |
53171 | 802 |
val _ = Theory.setup (Loaded_Values.map (insert (op =) string)); |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
803 |
in () end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
804 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
805 |
fun abort _ = error "Only value bindings allowed."; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
806 |
|
62494 | 807 |
val notifying_context : ML_Compiler0.context = |
62354 | 808 |
{name_space = |
62495 | 809 |
{lookupVal = #lookupVal ML_Env.name_space, |
810 |
lookupType = #lookupType ML_Env.name_space, |
|
811 |
lookupFix = #lookupFix ML_Env.name_space, |
|
812 |
lookupStruct = #lookupStruct ML_Env.name_space, |
|
813 |
lookupSig = #lookupSig ML_Env.name_space, |
|
814 |
lookupFunct = #lookupFunct ML_Env.name_space, |
|
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
815 |
enterVal = notify_val, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
816 |
enterType = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
817 |
enterFix = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
818 |
enterStruct = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
819 |
enterSig = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
820 |
enterFunct = abort, |
62495 | 821 |
allVal = #allVal ML_Env.name_space, |
822 |
allType = #allType ML_Env.name_space, |
|
823 |
allFix = #allFix ML_Env.name_space, |
|
824 |
allStruct = #allStruct ML_Env.name_space, |
|
825 |
allSig = #allSig ML_Env.name_space, |
|
826 |
allFunct = #allFunct ML_Env.name_space}, |
|
62716
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
wenzelm
parents:
62607
diff
changeset
|
827 |
print_depth = NONE, |
62495 | 828 |
here = #here ML_Env.context, |
829 |
print = #print ML_Env.context, |
|
830 |
error = #error ML_Env.context}; |
|
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
831 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
832 |
in |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
833 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
834 |
fun use_file filepath thy = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
835 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
836 |
val thy' = Loaded_Values.put [] thy; |
62889 | 837 |
val _ = Context.put_generic_context ((SOME o Context.Theory) thy'); |
60956 | 838 |
val _ = |
62902
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
wenzelm
parents:
62889
diff
changeset
|
839 |
ML_Compiler0.ML notifying_context |
60956 | 840 |
{line = 0, file = Path.implode filepath, verbose = false, debug = false} |
841 |
(File.read filepath); |
|
62876 | 842 |
val thy'' = Context.the_global_context (); |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
843 |
val names = Loaded_Values.get thy''; |
40320
abc52faa7761
polyml_as_definition does not require explicit dependencies on external ML files
haftmann
parents:
40257
diff
changeset
|
844 |
in (names, thy'') end; |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
845 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
846 |
end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
847 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
848 |
fun add_definiendum (ml_name, (b, T)) thy = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
849 |
thy |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
850 |
|> Code_Target.add_reserved target ml_name |
63178 | 851 |
|> Specification.axiomatization [(b, SOME T, NoSyn)] [] [] [] |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
852 |
|-> (fn ([Const (const, _)], _) => |
55150 | 853 |
Code_Target.set_printings (Constant (const, |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52434
diff
changeset
|
854 |
[(target, SOME (Code_Printer.simple_const_syntax (0, (K o K o K o Code_Printer.str) ml_name)))])) |
64940
19ca3644ec46
experimental computations: use arbitrary generated code for RHSs, not just constants
haftmann
parents:
64928
diff
changeset
|
855 |
#> tap (fn thy => Code_Target.produce_code (Proof_Context.init_global thy) false [const] target NONE Code_Target.generatedN [])); |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
856 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
857 |
fun process_file filepath (definienda, thy) = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
858 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
859 |
val (ml_names, thy') = use_file filepath thy; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
860 |
val superfluous = subtract (fn ((name1, _), name2) => name1 = name2) definienda ml_names; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
861 |
val _ = if null superfluous then () |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
862 |
else error ("Value binding(s) " ^ commas_quote superfluous |
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41486
diff
changeset
|
863 |
^ " found in external file " ^ Path.print filepath |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
864 |
^ " not present among the given contants binding(s)."); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
865 |
val these_definienda = AList.make (the o AList.lookup (op =) definienda) ml_names; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
866 |
val thy'' = fold add_definiendum these_definienda thy'; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
867 |
val definienda' = fold (AList.delete (op =)) ml_names definienda; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
868 |
in (definienda', thy'') end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
869 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
870 |
fun polyml_as_definition bTs filepaths thy = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
871 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
872 |
val definienda = map (fn bT => ((Binding.name_of o fst) bT, bT)) bTs; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
873 |
val (remaining, thy') = fold process_file filepaths (definienda, thy); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
874 |
val _ = if null remaining then () |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
875 |
else error ("Constant binding(s) " ^ commas_quote (map fst remaining) |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
876 |
^ " not present in external file(s)."); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
877 |
in thy' end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
878 |
|
28054 | 879 |
end; (*struct*) |