author | haftmann |
Tue, 21 Dec 2010 09:19:37 +0100 | |
changeset 41348 | 692c3fbde3c9 |
parent 41347 | 064133cb4ef6 |
child 41349 | 0e2a3f22f018 |
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 |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
14 |
val dynamic_value: 'a cookie -> theory -> string option |
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 |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
16 |
val dynamic_value_strict: 'a cookie -> theory -> string option |
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 |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
18 |
val dynamic_value_exn: 'a cookie -> theory -> string option |
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 |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
20 |
val static_value: 'a cookie -> theory -> string option |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
21 |
-> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a option |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
22 |
val static_value_strict: 'a cookie -> theory -> string option |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
23 |
-> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
24 |
val static_value_exn: 'a cookie -> theory -> string option |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
25 |
-> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a Exn.result |
41247 | 26 |
val dynamic_holds_conv: theory -> conv |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
27 |
val static_holds_conv: theory -> string list -> conv |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
28 |
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
|
29 |
-> string option -> theory -> theory |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
30 |
datatype truth = Holds |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
31 |
val put_truth: (unit -> truth) -> Proof.context -> Proof.context |
40150 | 32 |
val trace: bool Unsynchronized.ref |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
33 |
val polyml_as_definition: (binding * typ) list -> Path.T list -> theory -> theory |
28054 | 34 |
end; |
35 |
||
39401 | 36 |
structure Code_Runtime : CODE_RUNTIME = |
28054 | 37 |
struct |
38 |
||
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
39 |
open Basic_Code_Thingol; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
40 |
|
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
41 |
(** evaluation **) |
33992 | 42 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
43 |
(* technical prerequisites *) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
44 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
45 |
val this = "Code_Runtime"; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
46 |
val s_truth = Long_Name.append this "truth"; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
47 |
val s_Holds = Long_Name.append this "Holds"; |
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 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
53 |
val _ = Context.>> (Context.map_theory |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
54 |
(Code_Target.extend_target (target, (Code_ML.target_SML, K I)) |
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
55 |
#> Code_Target.add_tyco_syntax target @{type_name prop} |
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
56 |
(SOME (0, (K o K o K) (Code_Printer.str s_truth))) |
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
57 |
#> Code_Target.add_const_syntax target @{const_name Code_Generator.holds} |
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
58 |
(SOME (Code_Printer.plain_const_syntax s_Holds)) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
59 |
#> Code_Target.add_reserved target this |
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
60 |
#> fold (Code_Target.add_reserved target) ["oo", "ooo", "oooo", "upto", "downto", "orf", "andf"])); |
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 |
|
40150 | 63 |
val trace = Unsynchronized.ref false; |
64 |
||
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
65 |
fun exec verbose code = |
40150 | 66 |
(if ! trace then tracing code else (); |
40257 | 67 |
ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code)); |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
68 |
|
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
69 |
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
|
70 |
let |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
71 |
val code = (prelude |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
72 |
^ "\nval _ = Context.set_thread_data (SOME (Context.map_proof (" ^ put_ml |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
73 |
^ " (fn () => " ^ value ^ ")) (ML_Context.the_generic_context ())))"); |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
74 |
val ctxt' = ctxt |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
75 |
|> put (fn () => error ("Bad evaluation for " ^ quote put_ml)) |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
76 |
|> Context.proof_map (exec false code); |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
77 |
in get ctxt' () end; |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
78 |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
79 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
80 |
(* evaluation into target language values *) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
81 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
82 |
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
|
83 |
|
39605 | 84 |
fun reject_vars thy t = |
85 |
let |
|
86 |
val ctxt = ProofContext.init_global thy; |
|
87 |
in ((Sign.no_frees ctxt o Sign.no_vars ctxt o map_types (K dummyT)) t; t) end; |
|
88 |
||
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
89 |
fun obtain_evaluator thy some_target naming program deps = |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
90 |
Code_Target.evaluator thy (the_default target some_target) naming program deps; |
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
91 |
|
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
92 |
fun evaluation cookie thy evaluator vs_t args = |
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
93 |
let |
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
94 |
val ctxt = ProofContext.init_global thy; |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
95 |
val (program_code, value_name) = evaluator vs_t; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
96 |
val value_code = space_implode " " |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
97 |
(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
|
98 |
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
|
99 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
100 |
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
|
101 |
handle General.Match => NONE |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
102 |
| General.Bind => NONE |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
103 |
| General.Fail _ => NONE; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
104 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
105 |
fun dynamic_value_exn cookie thy some_target postproc t args = |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
106 |
let |
39605 | 107 |
val _ = reject_vars thy t; |
41099 | 108 |
val _ = if ! trace |
109 |
then tracing ("Evaluation of term " ^ quote (Syntax.string_of_term_global thy t)) |
|
110 |
else () |
|
39482 | 111 |
fun evaluator naming program ((_, vs_ty), t) deps = |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
112 |
evaluation cookie thy (obtain_evaluator thy some_target naming program deps) (vs_ty, t) args; |
41184 | 113 |
in Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t end; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
114 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
115 |
fun dynamic_value_strict cookie thy some_target postproc t args = |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
116 |
Exn.release (dynamic_value_exn cookie thy some_target postproc t args); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
117 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
118 |
fun dynamic_value cookie thy some_target postproc t args = |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
119 |
partiality_as_none (dynamic_value_exn cookie thy some_target postproc t args); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
120 |
|
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
121 |
fun static_evaluator cookie thy some_target naming program consts' = |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
122 |
let |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
123 |
val evaluator = obtain_evaluator thy some_target naming program consts' |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
124 |
in fn ((_, vs_ty), t) => fn _ => evaluation cookie thy evaluator (vs_ty, t) [] end; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
125 |
|
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
126 |
fun static_value_exn cookie thy some_target postproc consts = |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
127 |
Code_Thingol.static_value thy (Exn.map_result o postproc) consts |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
128 |
(static_evaluator cookie thy some_target) o reject_vars thy; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
129 |
|
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
130 |
fun static_value_strict cookie thy some_target postproc consts = |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
131 |
Exn.release o static_value_exn cookie thy some_target postproc consts; |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
132 |
|
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
133 |
fun static_value cookie thy some_target postproc consts = |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
134 |
partiality_as_none o static_value_exn cookie thy some_target postproc consts; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
135 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
136 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
137 |
(* evaluation for truth or nothing *) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
138 |
|
39820 | 139 |
structure Truth_Result = Proof_Data |
140 |
( |
|
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
141 |
type T = unit -> truth |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
142 |
fun init _ () = error "Truth_Result" |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
143 |
); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
144 |
val put_truth = Truth_Result.put; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
145 |
val truth_cookie = (Truth_Result.get, put_truth, Long_Name.append this "put_truth"); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
146 |
|
39605 | 147 |
val reject_vars = fn thy => tap (reject_vars thy o Thm.term_of); |
148 |
||
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
149 |
fun check_holds some_target naming thy program vs_t deps ct = |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
150 |
let |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
151 |
val t = Thm.term_of ct; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
152 |
val _ = if fastype_of t <> propT |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
153 |
then error ("Not a proposition: " ^ Syntax.string_of_term_global thy t) |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
154 |
else (); |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
155 |
val iff = Thm.cterm_of thy (Term.Const ("==", propT --> propT --> propT)); |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
156 |
val result = case partiality_as_none |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
157 |
(evaluation truth_cookie thy (obtain_evaluator thy some_target naming program deps) vs_t []) |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
158 |
of SOME Holds => true |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
159 |
| _ => false; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
160 |
in |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
161 |
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
|
162 |
end; |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
163 |
|
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
164 |
val (_, raw_check_holds_oracle) = Context.>>> (Context.map_theory_result |
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
165 |
(Thm.add_oracle (Binding.name "holds_by_evaluation", |
41348 | 166 |
fn (some_target, naming, thy, program, vs_t, deps, ct) => check_holds some_target naming thy program vs_t deps ct))); |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
167 |
|
41348 | 168 |
fun check_holds_oracle some_target naming thy program ((_, vs_ty), t) deps ct = |
169 |
raw_check_holds_oracle (some_target, naming, thy, program, (vs_ty, t), deps, ct); |
|
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
170 |
|
41247 | 171 |
fun dynamic_holds_conv thy = Code_Thingol.dynamic_conv thy |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
172 |
(fn naming => check_holds_oracle NONE naming thy) |
41247 | 173 |
o reject_vars thy; |
39473
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
haftmann
parents:
39422
diff
changeset
|
174 |
|
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
175 |
fun static_holds_conv thy consts = |
41347
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
176 |
Code_Thingol.static_conv thy consts |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
177 |
(fn naming => fn program => fn consts' => check_holds_oracle NONE naming thy program) |
064133cb4ef6
evaluator separating static and dynamic operations
haftmann
parents:
41343
diff
changeset
|
178 |
o reject_vars thy; |
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
179 |
|
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
180 |
|
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
181 |
(** instrumentalization **) |
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
182 |
|
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
183 |
fun evaluation_code thy module_name tycos consts = |
33992 | 184 |
let |
36271
2ef9dbddfcb8
optionally ignore errors during translation of equations
haftmann
parents:
35360
diff
changeset
|
185 |
val (consts', (naming, program)) = Code_Thingol.consts_program thy false consts; |
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
186 |
val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; |
38921 | 187 |
val (ml_code, target_names) = Code_Target.produce_code_for thy |
38933
bd77e092f67c
avoid strange special treatment of empty module names
haftmann
parents:
38931
diff
changeset
|
188 |
target NONE module_name [] naming program (consts' @ tycos'); |
34028
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
189 |
val (consts'', tycos'') = chop (length consts') target_names; |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
190 |
val consts_map = map2 (fn const => fn NONE => |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
191 |
error ("Constant " ^ (quote o Code.string_of_const thy) const |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
192 |
^ "\nhas a user-defined serialization") |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
193 |
| SOME const'' => (const, const'')) consts consts'' |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
194 |
val tycos_map = map2 (fn tyco => fn NONE => |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
195 |
error ("Type " ^ (quote o Sign.extern_type thy) tyco |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
196 |
^ "\nhas a user-defined serialization") |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
197 |
| SOME tyco'' => (tyco, tyco'')) tycos tycos''; |
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
haftmann
parents:
33992
diff
changeset
|
198 |
in (ml_code, (tycos_map, consts_map)) end; |
28054 | 199 |
|
200 |
||
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
201 |
(* by antiquotation *) |
28054 | 202 |
|
203 |
local |
|
204 |
||
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
205 |
structure Code_Antiq_Data = Proof_Data |
28054 | 206 |
( |
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
207 |
type T = (string list * string list) * (bool |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
208 |
* (string * (string * string) list) lazy); |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
209 |
fun init _ = (([], []), (true, (Lazy.value ("", [])))); |
28054 | 210 |
); |
211 |
||
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
212 |
val is_first_occ = fst o snd o Code_Antiq_Data.get; |
28054 | 213 |
|
30962 | 214 |
fun register_code new_tycos new_consts ctxt = |
28054 | 215 |
let |
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
216 |
val ((tycos, consts), _) = Code_Antiq_Data.get ctxt; |
30962 | 217 |
val tycos' = fold (insert (op =)) new_tycos tycos; |
218 |
val consts' = fold (insert (op =)) new_consts consts; |
|
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
219 |
val acc_code = Lazy.lazy (fn () => |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
220 |
evaluation_code (ProofContext.theory_of ctxt) "Code" tycos' consts' |
40422
d425d1ed82af
corrected slip: must keep constant map, not type map; tuned code
haftmann
parents:
40421
diff
changeset
|
221 |
|> apsnd snd); |
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
222 |
in Code_Antiq_Data.put ((tycos', consts'), (false, acc_code)) ctxt end; |
30962 | 223 |
|
224 |
fun register_const const = register_code [] [const]; |
|
28054 | 225 |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
226 |
fun print_const const all_struct_name consts_map = |
30962 | 227 |
(Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; |
228 |
||
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
229 |
fun print_code is_first const ctxt = |
30962 | 230 |
let |
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
231 |
val (_, (_, acc_code)) = Code_Antiq_Data.get ctxt; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
232 |
val (ml_code, consts_map) = Lazy.force acc_code; |
40422
d425d1ed82af
corrected slip: must keep constant map, not type map; tuned code
haftmann
parents:
40421
diff
changeset
|
233 |
val ml_code = if is_first then ml_code else ""; |
38930
072363be3fd9
modernized; avoid pointless tinkering with structure names
haftmann
parents:
38929
diff
changeset
|
234 |
val all_struct_name = "Isabelle"; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
235 |
in (ml_code, print_const const all_struct_name consts_map) end; |
28054 | 236 |
|
237 |
in |
|
238 |
||
35019
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
wenzelm
parents:
34032
diff
changeset
|
239 |
fun ml_code_antiq raw_const background = |
28054 | 240 |
let |
31156 | 241 |
val const = Code.check_const (ProofContext.theory_of background) raw_const; |
28054 | 242 |
val is_first = is_first_occ background; |
243 |
val background' = register_const const background; |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
244 |
in (print_code is_first const, background') end; |
30962 | 245 |
|
28054 | 246 |
end; (*local*) |
247 |
||
248 |
||
39404
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
haftmann
parents:
39401
diff
changeset
|
249 |
(* reflection support *) |
36470 | 250 |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
251 |
fun check_datatype thy tyco some_consts = |
36470 | 252 |
let |
40726
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents:
40711
diff
changeset
|
253 |
val constrs = (map fst o snd o fst o Code.get_type thy) tyco; |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
254 |
val _ = case some_consts |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
255 |
of SOME consts => |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
256 |
let |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
257 |
val missing_constrs = subtract (op =) consts constrs; |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
258 |
val _ = if null missing_constrs then [] |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
259 |
else error ("Missing constructor(s) " ^ commas (map quote missing_constrs) |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
260 |
^ " for datatype " ^ quote tyco); |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
261 |
val false_constrs = subtract (op =) constrs consts; |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
262 |
val _ = if null false_constrs then [] |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
263 |
else error ("Non-constructor(s) " ^ commas (map quote false_constrs) |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
264 |
^ " for datatype " ^ quote tyco) |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
265 |
in () end |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
266 |
| NONE => (); |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
267 |
in (tyco, constrs) end; |
36470 | 268 |
|
269 |
fun add_eval_tyco (tyco, tyco') thy = |
|
270 |
let |
|
271 |
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
|
272 |
fun pr pr' _ [] = tyco' |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
273 |
| pr pr' _ [ty] = |
36470 | 274 |
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
|
275 |
| pr pr' _ tys = |
36470 | 276 |
Code_Printer.concat [Code_Printer.enum "," "(" ")" (map (pr' Code_Printer.BR) tys), tyco'] |
277 |
in |
|
278 |
thy |
|
38923 | 279 |
|> Code_Target.add_tyco_syntax target tyco (SOME (k, pr)) |
36470 | 280 |
end; |
281 |
||
36514 | 282 |
fun add_eval_constr (const, const') thy = |
283 |
let |
|
284 |
val k = Code.args_number thy const; |
|
285 |
fun pr pr' fxy ts = Code_Printer.brackify fxy |
|
38922 | 286 |
(const' :: the_list (Code_Printer.tuplify pr' Code_Printer.BR (map fst ts))); |
36514 | 287 |
in |
288 |
thy |
|
38923 | 289 |
|> Code_Target.add_const_syntax target const (SOME (Code_Printer.simple_const_syntax (k, pr))) |
36514 | 290 |
end; |
291 |
||
38923 | 292 |
fun add_eval_const (const, const') = Code_Target.add_const_syntax target |
36470 | 293 |
const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const'))); |
294 |
||
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
295 |
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
|
296 |
thy |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
297 |
|> Code_Target.add_reserved target module_name |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
298 |
|> Context.theory_map (exec true code) |
38935
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
299 |
|> fold (add_eval_tyco o apsnd Code_Printer.str) tyco_map |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
300 |
|> fold (add_eval_constr o apsnd Code_Printer.str) constr_map |
2cf3d8305b47
corrected misbehaved additional qualification of generated names
haftmann
parents:
38933
diff
changeset
|
301 |
|> 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
|
302 |
| process_reflection (code, _) _ (SOME file_name) thy = |
36470 | 303 |
let |
37950
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
wenzelm
parents:
37744
diff
changeset
|
304 |
val preamble = |
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
wenzelm
parents:
37744
diff
changeset
|
305 |
"(* Generated from " ^ Path.implode (Thy_Header.thy_path (Context.theory_name thy)) |
36470 | 306 |
^ "; DO NOT EDIT! *)"; |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
307 |
val _ = File.write (Path.explode file_name) (preamble ^ "\n\n" ^ code); |
36470 | 308 |
in |
309 |
thy |
|
310 |
end; |
|
311 |
||
312 |
fun gen_code_reflect prep_type prep_const raw_datatypes raw_functions module_name some_file thy = |
|
313 |
let |
|
314 |
val datatypes = map (fn (raw_tyco, raw_cos) => |
|
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
315 |
(prep_type thy raw_tyco, (Option.map o map) (prep_const thy) raw_cos)) raw_datatypes; |
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
316 |
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
|
317 |
|> apsnd flat; |
36470 | 318 |
val functions = map (prep_const thy) raw_functions; |
36514 | 319 |
val result = evaluation_code thy module_name tycos (constrs @ functions) |
320 |
|> (apsnd o apsnd) (chop (length constrs)); |
|
36470 | 321 |
in |
322 |
thy |
|
39485
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
haftmann
parents:
39482
diff
changeset
|
323 |
|> process_reflection result module_name some_file |
36470 | 324 |
end; |
325 |
||
39422 | 326 |
val code_reflect = gen_code_reflect Code_Target.cert_tyco (K I); |
36470 | 327 |
val code_reflect_cmd = gen_code_reflect Code_Target.read_tyco Code.read_const; |
328 |
||
329 |
||
28054 | 330 |
(** Isar setup **) |
331 |
||
332 |
val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); |
|
333 |
||
36470 | 334 |
local |
335 |
||
336 |
val datatypesK = "datatypes"; |
|
337 |
val functionsK = "functions"; |
|
338 |
val fileK = "file"; |
|
339 |
val andK = "and" |
|
340 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
341 |
val _ = List.app Keyword.keyword [datatypesK, functionsK]; |
36470 | 342 |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
343 |
val parse_datatype = |
40711
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
haftmann
parents:
40422
diff
changeset
|
344 |
Parse.name --| Parse.$$$ "=" -- |
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
haftmann
parents:
40422
diff
changeset
|
345 |
(((Parse.sym_ident || Parse.string) >> (fn "_" => NONE | _ => Scan.fail ())) |
40421
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
haftmann
parents:
40320
diff
changeset
|
346 |
|| ((Parse.term ::: (Scan.repeat (Parse.$$$ "|" |-- Parse.term))) >> SOME)); |
36470 | 347 |
|
348 |
in |
|
349 |
||
350 |
val _ = |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
351 |
Outer_Syntax.command "code_reflect" "enrich runtime environment with generated code" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
352 |
Keyword.thy_decl (Parse.name -- Scan.optional (Parse.$$$ datatypesK |-- (parse_datatype |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
353 |
::: Scan.repeat (Parse.$$$ andK |-- parse_datatype))) [] |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
354 |
-- Scan.optional (Parse.$$$ functionsK |-- Scan.repeat1 Parse.name) [] |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
355 |
-- Scan.option (Parse.$$$ fileK |-- Parse.name) |
36534 | 356 |
>> (fn (((module_name, raw_datatypes), raw_functions), some_file) => Toplevel.theory |
36470 | 357 |
(code_reflect_cmd raw_datatypes raw_functions module_name some_file))); |
358 |
||
359 |
end; (*local*) |
|
360 |
||
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
361 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
362 |
(** 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
|
363 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
364 |
local |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
365 |
|
39820 | 366 |
structure Loaded_Values = Theory_Data |
367 |
( |
|
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
368 |
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
|
369 |
val empty = [] |
39820 | 370 |
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
|
371 |
val extend = I |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
372 |
); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
373 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
374 |
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
|
375 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
376 |
val _ = #enterVal ML_Env.name_space (string, value); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
377 |
val _ = Context.>> ((Context.map_theory o Loaded_Values.map) (insert (op =) string)); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
378 |
in () end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
379 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
380 |
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
|
381 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
382 |
val notifying_context : use_context = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
383 |
{tune_source = #tune_source ML_Env.local_context, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
384 |
name_space = |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
385 |
{lookupVal = #lookupVal ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
386 |
lookupType = #lookupType ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
387 |
lookupFix = #lookupFix ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
388 |
lookupStruct = #lookupStruct ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
389 |
lookupSig = #lookupSig ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
390 |
lookupFunct = #lookupFunct ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
391 |
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
|
392 |
enterType = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
393 |
enterFix = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
394 |
enterStruct = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
395 |
enterSig = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
396 |
enterFunct = abort, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
397 |
allVal = #allVal ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
398 |
allType = #allType ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
399 |
allFix = #allFix ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
400 |
allStruct = #allStruct ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
401 |
allSig = #allSig ML_Env.name_space, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
402 |
allFunct = #allFunct ML_Env.name_space}, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
403 |
str_of_pos = #str_of_pos ML_Env.local_context, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
404 |
print = #print ML_Env.local_context, |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
405 |
error = #error ML_Env.local_context}; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
406 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
407 |
in |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
408 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
409 |
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
|
410 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
411 |
val thy' = Loaded_Values.put [] thy; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
412 |
val _ = Context.set_thread_data ((SOME o Context.Theory) thy'); |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
413 |
val _ = Secure.use_text notifying_context |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
414 |
(0, Path.implode filepath) false (File.read filepath); |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
415 |
val thy'' = (Context.the_theory o the) (Context.thread_data ()); |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
416 |
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
|
417 |
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
|
418 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
419 |
end; |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
420 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
421 |
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
|
422 |
thy |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
423 |
|> Code_Target.add_reserved target ml_name |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
424 |
|> Specification.axiomatization [(b, SOME T, NoSyn)] [] |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
425 |
|-> (fn ([Const (const, _)], _) => |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
426 |
Code_Target.add_const_syntax target const |
39912
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
427 |
(SOME (Code_Printer.simple_const_syntax (0, (K o K o K o Code_Printer.str) ml_name))) |
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
haftmann
parents:
39816
diff
changeset
|
428 |
#> tap (fn thy => Code_Target.produce_code thy [const] target NONE "Code" [])); |
39816
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
429 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
430 |
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
|
431 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
432 |
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
|
433 |
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
|
434 |
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
|
435 |
else error ("Value binding(s) " ^ commas_quote superfluous |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
436 |
^ " found in external file " ^ quote (Path.implode filepath) |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
437 |
^ " 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
|
438 |
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
|
439 |
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
|
440 |
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
|
441 |
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
|
442 |
|
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
443 |
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
|
444 |
let |
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
haftmann
parents:
39605
diff
changeset
|
445 |
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
|
446 |
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
|
447 |
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
|
448 |
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
|
449 |
^ " 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
|
450 |
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
|
451 |
|
28054 | 452 |
end; (*struct*) |