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