| author | wenzelm | 
| Fri, 16 Mar 2012 18:20:12 +0100 | |
| changeset 46961 | 5c6955f487e5 | 
| parent 46949 | 94aa7b81bcf6 | 
| child 47609 | b3dab1892cda | 
| 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";  | 
| 44663 | 50  | 
val structure_generated = "Generated_Code";  | 
| 28054 | 51  | 
|
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
52  | 
datatype truth = Holds;  | 
| 39422 | 53  | 
|
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
54  | 
val _ = Context.>> (Context.map_theory  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
55  | 
(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
 | 
56  | 
  #> Code_Target.add_tyco_syntax target @{type_name prop}
 | 
| 
 
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
 
haftmann 
parents: 
39482 
diff
changeset
 | 
57  | 
(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
 | 
58  | 
  #> 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
 | 
59  | 
(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
 | 
60  | 
#> Code_Target.add_reserved target this  | 
| 
39485
 
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
 
haftmann 
parents: 
39482 
diff
changeset
 | 
61  | 
#> 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
 | 
62  | 
(*avoid further pervasive infix names*)  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
63  | 
|
| 40150 | 64  | 
val trace = Unsynchronized.ref false;  | 
65  | 
||
| 
39912
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
66  | 
fun exec verbose code =  | 
| 40150 | 67  | 
(if ! trace then tracing code else ();  | 
| 40257 | 68  | 
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
 | 
69  | 
|
| 
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
70  | 
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
 | 
71  | 
let  | 
| 
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
72  | 
val code = (prelude  | 
| 
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
73  | 
      ^ "\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
 | 
74  | 
^ " (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
 | 
75  | 
val ctxt' = ctxt  | 
| 
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
76  | 
      |> 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
 | 
77  | 
|> 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
 | 
78  | 
in get ctxt' () end;  | 
| 
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
79  | 
|
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
80  | 
|
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
81  | 
(* evaluation into target language values *)  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
82  | 
|
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
83  | 
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
 | 
84  | 
|
| 39605 | 85  | 
fun reject_vars thy t =  | 
86  | 
let  | 
|
| 42361 | 87  | 
val ctxt = Proof_Context.init_global thy;  | 
| 39605 | 88  | 
in ((Sign.no_frees ctxt o Sign.no_vars ctxt o map_types (K dummyT)) t; t) end;  | 
89  | 
||
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
90  | 
fun obtain_evaluator thy some_target = Code_Target.evaluator thy (the_default target some_target);  | 
| 
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  | 
| 42361 | 94  | 
val ctxt = Proof_Context.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  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41376 
diff
changeset
 | 
142  | 
(* FIXME avoid user error with non-user text *)  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
143  | 
fun init _ () = error "Truth_Result"  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
144  | 
);  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
145  | 
val put_truth = Truth_Result.put;  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
146  | 
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
 | 
147  | 
|
| 39605 | 148  | 
val reject_vars = fn thy => tap (reject_vars thy o Thm.term_of);  | 
149  | 
||
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
150  | 
local  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
151  | 
|
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
152  | 
fun check_holds thy evaluator vs_t deps ct =  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
153  | 
let  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
154  | 
val t = Thm.term_of ct;  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
155  | 
val _ = if fastype_of t <> propT  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
156  | 
      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
 | 
157  | 
else ();  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
158  | 
    val iff = Thm.cterm_of thy (Term.Const ("==", propT --> propT --> propT));
 | 
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
159  | 
val result = case partiality_as_none (evaluation truth_cookie thy evaluator vs_t [])  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
160  | 
of SOME Holds => true  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
161  | 
| _ => false;  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
162  | 
in  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
163  | 
    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
 | 
164  | 
end;  | 
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
165  | 
|
| 
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
166  | 
val (_, raw_check_holds_oracle) = Context.>>> (Context.map_theory_result  | 
| 
43619
 
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
 
wenzelm 
parents: 
43560 
diff
changeset
 | 
167  | 
  (Thm.add_oracle (@{binding holds_by_evaluation},
 | 
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
168  | 
fn (thy, evaluator, vs_t, deps, ct) => check_holds thy evaluator vs_t deps ct)));  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
169  | 
|
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
170  | 
fun check_holds_oracle thy evaluator ((_, vs_ty), t) deps ct =  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
171  | 
raw_check_holds_oracle (thy, evaluator, (vs_ty, t), deps, ct);  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
172  | 
|
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
173  | 
in  | 
| 
39485
 
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
 
haftmann 
parents: 
39482 
diff
changeset
 | 
174  | 
|
| 41247 | 175  | 
fun dynamic_holds_conv thy = Code_Thingol.dynamic_conv thy  | 
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
176  | 
(fn naming => fn program => fn vs_t => fn deps =>  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
177  | 
check_holds_oracle thy (obtain_evaluator thy NONE naming program deps) vs_t deps)  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
178  | 
o reject_vars thy;  | 
| 
39473
 
33638f4883ac
dynamic and static value computation; built-in evaluation of propositions
 
haftmann 
parents: 
39422 
diff
changeset
 | 
179  | 
|
| 
41349
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
180  | 
fun static_holds_conv thy consts = Code_Thingol.static_conv thy consts  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
181  | 
(fn naming => fn program => fn consts' =>  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
182  | 
check_holds_oracle thy (obtain_evaluator thy NONE naming program consts'))  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
183  | 
o reject_vars thy;  | 
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
184  | 
|
| 
 
0e2a3f22f018
evaluator separating static and dynamic operations
 
haftmann 
parents: 
41348 
diff
changeset
 | 
185  | 
end; (*local*)  | 
| 
39404
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
186  | 
|
| 
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
187  | 
|
| 
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
188  | 
(** instrumentalization **)  | 
| 
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
189  | 
|
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
190  | 
fun evaluation_code thy module_name tycos consts =  | 
| 33992 | 191  | 
let  | 
| 42361 | 192  | 
val ctxt = Proof_Context.init_global thy;  | 
| 
36271
 
2ef9dbddfcb8
optionally ignore errors during translation of equations
 
haftmann 
parents: 
35360 
diff
changeset
 | 
193  | 
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
 | 
194  | 
val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos;  | 
| 42359 | 195  | 
val (ml_code, target_names) =  | 
196  | 
Code_Target.produce_code_for thy  | 
|
197  | 
target NONE module_name [] naming program (consts' @ tycos');  | 
|
| 
34028
 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 
haftmann 
parents: 
33992 
diff
changeset
 | 
198  | 
val (consts'', tycos'') = chop (length consts') target_names;  | 
| 42359 | 199  | 
val consts_map = map2 (fn const =>  | 
200  | 
fn NONE =>  | 
|
201  | 
          error ("Constant " ^ (quote o Code.string_of_const thy) const ^
 | 
|
202  | 
"\nhas a user-defined serialization")  | 
|
203  | 
| SOME const'' => (const, const'')) consts consts''  | 
|
204  | 
val tycos_map = map2 (fn tyco =>  | 
|
205  | 
fn NONE =>  | 
|
| 42361 | 206  | 
          error ("Type " ^ quote (Proof_Context.extern_type ctxt tyco) ^
 | 
| 42359 | 207  | 
"\nhas a user-defined serialization")  | 
208  | 
| SOME tyco'' => (tyco, tyco'')) tycos tycos'';  | 
|
| 
34028
 
1e6206763036
split off evaluation mechanisms in separte module Code_Eval
 
haftmann 
parents: 
33992 
diff
changeset
 | 
209  | 
in (ml_code, (tycos_map, consts_map)) end;  | 
| 28054 | 210  | 
|
211  | 
||
| 
39404
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
212  | 
(* by antiquotation *)  | 
| 28054 | 213  | 
|
214  | 
local  | 
|
215  | 
||
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
216  | 
structure Code_Antiq_Data = Proof_Data  | 
| 28054 | 217  | 
(  | 
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
218  | 
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
 | 
219  | 
* (string * (string * string) list) lazy);  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
220  | 
  fun init _ = (([], []), (true, (Lazy.value ("", []))));
 | 
| 28054 | 221  | 
);  | 
222  | 
||
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
223  | 
val is_first_occ = fst o snd o Code_Antiq_Data.get;  | 
| 28054 | 224  | 
|
| 30962 | 225  | 
fun register_code new_tycos new_consts ctxt =  | 
| 28054 | 226  | 
let  | 
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
227  | 
val ((tycos, consts), _) = Code_Antiq_Data.get ctxt;  | 
| 30962 | 228  | 
val tycos' = fold (insert (op =)) new_tycos tycos;  | 
229  | 
val consts' = fold (insert (op =)) new_consts consts;  | 
|
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
230  | 
val acc_code = Lazy.lazy (fn () =>  | 
| 44663 | 231  | 
evaluation_code (Proof_Context.theory_of ctxt) structure_generated tycos' consts'  | 
| 
40422
 
d425d1ed82af
corrected slip: must keep constant map, not type map; tuned code
 
haftmann 
parents: 
40421 
diff
changeset
 | 
232  | 
|> apsnd snd);  | 
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
233  | 
in Code_Antiq_Data.put ((tycos', consts'), (false, acc_code)) ctxt end;  | 
| 30962 | 234  | 
|
235  | 
fun register_const const = register_code [] [const];  | 
|
| 28054 | 236  | 
|
| 
40421
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
237  | 
fun print_code is_first const ctxt =  | 
| 30962 | 238  | 
let  | 
| 
38930
 
072363be3fd9
modernized; avoid pointless tinkering with structure names
 
haftmann 
parents: 
38929 
diff
changeset
 | 
239  | 
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
 | 
240  | 
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
 | 
241  | 
val ml_code = if is_first then ml_code else "";  | 
| 
41486
 
82c1e348bc18
reverted 08240feb69c7 -- breaks positions of reports;
 
wenzelm 
parents: 
41472 
diff
changeset
 | 
242  | 
val body = "Isabelle." ^ the (AList.lookup (op =) consts_map const);  | 
| 
41376
 
08240feb69c7
more robust ML antiquotations -- allow original tokens without adjacent whitespace;
 
wenzelm 
parents: 
41349 
diff
changeset
 | 
243  | 
in (ml_code, body) end;  | 
| 28054 | 244  | 
|
245  | 
in  | 
|
246  | 
||
| 
35019
 
1ec0a3ff229e
simplified interface for ML antiquotations, struct_name is always "Isabelle";
 
wenzelm 
parents: 
34032 
diff
changeset
 | 
247  | 
fun ml_code_antiq raw_const background =  | 
| 28054 | 248  | 
let  | 
| 42361 | 249  | 
val const = Code.check_const (Proof_Context.theory_of background) raw_const;  | 
| 28054 | 250  | 
val is_first = is_first_occ background;  | 
251  | 
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
 | 
252  | 
in (print_code is_first const, background') end;  | 
| 30962 | 253  | 
|
| 28054 | 254  | 
end; (*local*)  | 
255  | 
||
256  | 
||
| 
39404
 
a8c337299bc1
Code_Runtime.value, corresponding to ML_Context.value; tuned
 
haftmann 
parents: 
39401 
diff
changeset
 | 
257  | 
(* reflection support *)  | 
| 36470 | 258  | 
|
| 
40421
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
259  | 
fun check_datatype thy tyco some_consts =  | 
| 36470 | 260  | 
let  | 
| 
40726
 
16dcfedc4eb7
keep type variable arguments of datatype constructors in bookkeeping
 
haftmann 
parents: 
40711 
diff
changeset
 | 
261  | 
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
 | 
262  | 
val _ = case some_consts  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
263  | 
of SOME consts =>  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
264  | 
let  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
265  | 
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
 | 
266  | 
val _ = if null missing_constrs then []  | 
| 45430 | 267  | 
              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
 | 
268  | 
^ " for datatype " ^ quote tyco);  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
269  | 
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
 | 
270  | 
val _ = if null false_constrs then []  | 
| 45430 | 271  | 
              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
 | 
272  | 
^ " for datatype " ^ quote tyco)  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
273  | 
in () end  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
274  | 
| NONE => ();  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
275  | 
in (tyco, constrs) end;  | 
| 36470 | 276  | 
|
277  | 
fun add_eval_tyco (tyco, tyco') thy =  | 
|
278  | 
let  | 
|
279  | 
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
 | 
280  | 
fun pr pr' _ [] = tyco'  | 
| 
 
b41aabb629ce
constructors to datatypes in code_reflect can be globbed; dropped unused code
 
haftmann 
parents: 
40320 
diff
changeset
 | 
281  | 
| pr pr' _ [ty] =  | 
| 36470 | 282  | 
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
 | 
283  | 
| pr pr' _ tys =  | 
| 36470 | 284  | 
          Code_Printer.concat [Code_Printer.enum "," "(" ")" (map (pr' Code_Printer.BR) tys), tyco']
 | 
285  | 
in  | 
|
286  | 
thy  | 
|
| 38923 | 287  | 
|> Code_Target.add_tyco_syntax target tyco (SOME (k, pr))  | 
| 36470 | 288  | 
end;  | 
289  | 
||
| 36514 | 290  | 
fun add_eval_constr (const, const') thy =  | 
291  | 
let  | 
|
292  | 
val k = Code.args_number thy const;  | 
|
293  | 
fun pr pr' fxy ts = Code_Printer.brackify fxy  | 
|
| 38922 | 294  | 
(const' :: the_list (Code_Printer.tuplify pr' Code_Printer.BR (map fst ts)));  | 
| 36514 | 295  | 
in  | 
296  | 
thy  | 
|
| 38923 | 297  | 
|> Code_Target.add_const_syntax target const (SOME (Code_Printer.simple_const_syntax (k, pr)))  | 
| 36514 | 298  | 
end;  | 
299  | 
||
| 38923 | 300  | 
fun add_eval_const (const, const') = Code_Target.add_const_syntax target  | 
| 36470 | 301  | 
const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const')));  | 
302  | 
||
| 
39912
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
303  | 
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
 | 
304  | 
thy  | 
| 
 
2cf3d8305b47
corrected misbehaved additional qualification of generated names
 
haftmann 
parents: 
38933 
diff
changeset
 | 
305  | 
|> 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
 | 
306  | 
|> Context.theory_map (exec true code)  | 
| 
38935
 
2cf3d8305b47
corrected misbehaved additional qualification of generated names
 
haftmann 
parents: 
38933 
diff
changeset
 | 
307  | 
|> fold (add_eval_tyco o apsnd Code_Printer.str) tyco_map  | 
| 
 
2cf3d8305b47
corrected misbehaved additional qualification of generated names
 
haftmann 
parents: 
38933 
diff
changeset
 | 
308  | 
|> fold (add_eval_constr o apsnd Code_Printer.str) constr_map  | 
| 
 
2cf3d8305b47
corrected misbehaved additional qualification of generated names
 
haftmann 
parents: 
38933 
diff
changeset
 | 
309  | 
|> 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
 | 
310  | 
| process_reflection (code, _) _ (SOME file_name) thy =  | 
| 36470 | 311  | 
let  | 
| 
37950
 
bc285d91041e
moved basic thy file name operations from Thy_Load to Thy_Header;
 
wenzelm 
parents: 
37744 
diff
changeset
 | 
312  | 
val preamble =  | 
| 46737 | 313  | 
"(* Generated from " ^  | 
314  | 
Path.implode (Thy_Load.thy_path (Path.basic (Context.theory_name thy))) ^  | 
|
315  | 
"; DO NOT EDIT! *)";  | 
|
| 
39912
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
316  | 
val _ = File.write (Path.explode file_name) (preamble ^ "\n\n" ^ code);  | 
| 36470 | 317  | 
in  | 
318  | 
thy  | 
|
319  | 
end;  | 
|
320  | 
||
321  | 
fun gen_code_reflect prep_type prep_const raw_datatypes raw_functions module_name some_file thy =  | 
|
322  | 
let  | 
|
323  | 
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
 | 
324  | 
(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
 | 
325  | 
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
 | 
326  | 
|> apsnd flat;  | 
| 36470 | 327  | 
val functions = map (prep_const thy) raw_functions;  | 
| 36514 | 328  | 
val result = evaluation_code thy module_name tycos (constrs @ functions)  | 
329  | 
|> (apsnd o apsnd) (chop (length constrs));  | 
|
| 36470 | 330  | 
in  | 
331  | 
thy  | 
|
| 
39485
 
f7270a5e2550
closures preserve static serializer context for static evaluation; tuned
 
haftmann 
parents: 
39482 
diff
changeset
 | 
332  | 
|> process_reflection result module_name some_file  | 
| 36470 | 333  | 
end;  | 
334  | 
||
| 39422 | 335  | 
val code_reflect = gen_code_reflect Code_Target.cert_tyco (K I);  | 
| 36470 | 336  | 
val code_reflect_cmd = gen_code_reflect Code_Target.read_tyco Code.read_const;  | 
337  | 
||
338  | 
||
| 28054 | 339  | 
(** Isar setup **)  | 
340  | 
||
| 
43560
 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
341  | 
val _ =  | 
| 
 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
342  | 
Context.>> (Context.map_theory  | 
| 
43619
 
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
 
wenzelm 
parents: 
43560 
diff
changeset
 | 
343  | 
    (ML_Context.add_antiq @{binding code} (fn _ => Args.term >> ml_code_antiq)));
 | 
| 28054 | 344  | 
|
| 36470 | 345  | 
local  | 
346  | 
||
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36610 
diff
changeset
 | 
347  | 
val parse_datatype =  | 
| 46949 | 348  | 
  Parse.name --| @{keyword "="} --
 | 
| 
40711
 
81bc73585eec
globbing constant expressions use more idiomatic underscore rather than star
 
haftmann 
parents: 
40422 
diff
changeset
 | 
349  | 
(((Parse.sym_ident || Parse.string) >> (fn "_" => NONE | _ => Scan.fail ()))  | 
| 46949 | 350  | 
    || ((Parse.term ::: (Scan.repeat (@{keyword "|"} |-- Parse.term))) >> SOME));
 | 
| 36470 | 351  | 
|
352  | 
in  | 
|
353  | 
||
354  | 
val _ =  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
355  | 
  Outer_Syntax.command @{command_spec "code_reflect"}
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
356  | 
"enrich runtime environment with generated code"  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
357  | 
    (Parse.name -- Scan.optional (@{keyword "datatypes"} |-- (parse_datatype
 | 
| 46949 | 358  | 
      ::: Scan.repeat (@{keyword "and"} |-- parse_datatype))) []
 | 
359  | 
    -- Scan.optional (@{keyword "functions"} |-- Scan.repeat1 Parse.name) []
 | 
|
360  | 
    -- Scan.option (@{keyword "file"} |-- Parse.name)
 | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
361  | 
>> (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
 | 
362  | 
(code_reflect_cmd raw_datatypes raw_functions module_name some_file)));  | 
| 36470 | 363  | 
|
364  | 
end; (*local*)  | 
|
365  | 
||
| 
39816
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
366  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
367  | 
(** 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
 | 
368  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
369  | 
local  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
370  | 
|
| 39820 | 371  | 
structure Loaded_Values = Theory_Data  | 
372  | 
(  | 
|
| 
39816
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
373  | 
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
 | 
374  | 
val empty = []  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41376 
diff
changeset
 | 
375  | 
val extend = I  | 
| 39820 | 376  | 
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
 | 
377  | 
);  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
378  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
379  | 
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
 | 
380  | 
let  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
381  | 
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
 | 
382  | 
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
 | 
383  | 
in () end;  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
384  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
385  | 
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
 | 
386  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
387  | 
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
 | 
388  | 
 {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
 | 
389  | 
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  | 
   {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
 | 
391  | 
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
 | 
392  | 
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
 | 
393  | 
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
 | 
394  | 
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
 | 
395  | 
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
 | 
396  | 
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
 | 
397  | 
enterType = abort,  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
398  | 
enterFix = abort,  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
399  | 
enterStruct = abort,  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
400  | 
enterSig = abort,  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
401  | 
enterFunct = abort,  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
402  | 
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
 | 
403  | 
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
 | 
404  | 
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
 | 
405  | 
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
 | 
406  | 
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
 | 
407  | 
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
 | 
408  | 
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
 | 
409  | 
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
 | 
410  | 
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
 | 
411  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
412  | 
in  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
413  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
414  | 
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
 | 
415  | 
let  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
416  | 
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
 | 
417  | 
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
 | 
418  | 
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
 | 
419  | 
(0, Path.implode filepath) false (File.read filepath);  | 
| 45268 | 420  | 
val thy'' = Context.the_theory (Context.the_thread_data ());  | 
| 
39912
 
2ffe7060ca75
avoid antiquotation processing for code_reflect; moved ML_Context.value to Code_Runtime
 
haftmann 
parents: 
39816 
diff
changeset
 | 
421  | 
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
 | 
422  | 
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
 | 
423  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
424  | 
end;  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
425  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
426  | 
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
 | 
427  | 
thy  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
428  | 
|> 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
 | 
429  | 
|> 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
 | 
430  | 
|-> (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
 | 
431  | 
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
 | 
432  | 
(SOME (Code_Printer.simple_const_syntax (0, (K o K o K o Code_Printer.str) ml_name)))  | 
| 44663 | 433  | 
#> tap (fn thy => Code_Target.produce_code thy [const] target NONE structure_generated []));  | 
| 
39816
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
434  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
435  | 
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
 | 
436  | 
let  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
437  | 
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
 | 
438  | 
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
 | 
439  | 
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
 | 
440  | 
      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
 | 
441  | 
^ " 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
 | 
442  | 
^ " 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
 | 
443  | 
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
 | 
444  | 
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
 | 
445  | 
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
 | 
446  | 
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
 | 
447  | 
|
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
448  | 
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
 | 
449  | 
let  | 
| 
 
c7cec137c48a
added polyml_as_definition -- using external SML files as substitute for proper definitions -- only for polyml!
 
haftmann 
parents: 
39605 
diff
changeset
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
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
 | 
453  | 
      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
 | 
454  | 
^ " 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
 | 
455  | 
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
 | 
456  | 
|
| 28054 | 457  | 
end; (*struct*)  |