author | haftmann |
Fri, 30 May 2008 08:02:19 +0200 | |
changeset 27024 | fcab2dd46872 |
parent 27014 | a5f53d9d2b60 |
permissions | -rw-r--r-- |
24219 | 1 |
(* Title: Tools/code/code_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
24918 | 5 |
Code generator interfaces and Isar setup. |
24219 | 6 |
*) |
7 |
||
8 |
signature CODE_PACKAGE = |
|
9 |
sig |
|
27014
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
10 |
val source_of: theory -> string -> string -> string list -> string; |
26011 | 11 |
val evaluate_conv: theory |
26740 | 12 |
-> (term -> term * (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm |
13 |
-> string list -> thm)) |
|
24283 | 14 |
-> cterm -> thm; |
26011 | 15 |
val evaluate_term: theory |
26740 | 16 |
-> (term -> term * (CodeThingol.code -> CodeThingol.typscheme * CodeThingol.iterm |
17 |
-> string list -> 'a)) |
|
24835
8c26128f8997
clarified relationship of code generator conversions and evaluations
haftmann
parents:
24811
diff
changeset
|
18 |
-> term -> 'a; |
26011 | 19 |
val eval_conv: string * (unit -> thm) option ref |
20 |
-> theory -> cterm -> string list -> thm; |
|
21 |
val eval_term: string * (unit -> 'a) option ref |
|
22 |
-> theory -> term -> string list -> 'a; |
|
23 |
val satisfies: theory -> term -> string list -> bool; |
|
24585 | 24 |
val satisfies_ref: (unit -> bool) option ref; |
25611 | 25 |
val codegen_shell_command: string (*theory name*) -> string (*cg expr*) -> unit; |
24219 | 26 |
end; |
27 |
||
28 |
structure CodePackage : CODE_PACKAGE = |
|
29 |
struct |
|
30 |
||
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24835
diff
changeset
|
31 |
(** code theorems **) |
24219 | 32 |
|
24283 | 33 |
fun code_depgr thy [] = CodeFuncgr.make thy [] |
24219 | 34 |
| code_depgr thy consts = |
35 |
let |
|
24283 | 36 |
val gr = CodeFuncgr.make thy consts; |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
37 |
val select = Graph.all_succs gr consts; |
24219 | 38 |
in |
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
39 |
gr |
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
40 |
|> Graph.subgraph (member (op =) select) |
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25485
diff
changeset
|
41 |
|> Graph.map_nodes ((apsnd o map) (AxClass.overload thy)) |
24219 | 42 |
end; |
43 |
||
44 |
fun code_thms thy = |
|
45 |
Pretty.writeln o CodeFuncgr.pretty thy o code_depgr thy; |
|
46 |
||
47 |
fun code_deps thy consts = |
|
48 |
let |
|
49 |
val gr = code_depgr thy consts; |
|
50 |
fun mk_entry (const, (_, (_, parents))) = |
|
51 |
let |
|
52 |
val name = CodeUnit.string_of_const thy const; |
|
53 |
val nameparents = map (CodeUnit.string_of_const thy) parents; |
|
54 |
in { name = name, ID = name, dir = "", unfold = true, |
|
55 |
path = "", parents = nameparents } |
|
56 |
end; |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
57 |
val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; |
24219 | 58 |
in Present.display_graph prgr end; |
59 |
||
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24835
diff
changeset
|
60 |
|
24918 | 61 |
(** code generation interfaces **) |
62 |
||
63 |
(* code data *) |
|
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24835
diff
changeset
|
64 |
|
24219 | 65 |
structure Program = CodeDataFun |
66 |
( |
|
67 |
type T = CodeThingol.code; |
|
68 |
val empty = CodeThingol.empty_code; |
|
69 |
fun merge _ = CodeThingol.merge_code; |
|
70 |
fun purge _ NONE _ = CodeThingol.empty_code |
|
71 |
| purge NONE _ _ = CodeThingol.empty_code |
|
72 |
| purge (SOME thy) (SOME cs) code = |
|
73 |
let |
|
74 |
val cs_exisiting = |
|
75 |
map_filter (CodeName.const_rev thy) (Graph.keys code); |
|
76 |
val dels = (Graph.all_preds code |
|
77 |
o map (CodeName.const thy) |
|
24423
ae9cd0e92423
overloaded definitions accompanied by explicit constants
haftmann
parents:
24381
diff
changeset
|
78 |
o filter (member (op =) cs_exisiting) |
24219 | 79 |
) cs; |
80 |
in Graph.del_nodes dels code end; |
|
81 |
); |
|
82 |
||
24918 | 83 |
(* generic generation combinators *) |
24811 | 84 |
|
24918 | 85 |
val ensure_const = CodeThingol.ensure_const; |
24219 | 86 |
|
24918 | 87 |
fun perhaps_const thy algbr funcgr c trns = |
88 |
case try (CodeThingol.ensure_const thy algbr funcgr c) trns |
|
24219 | 89 |
of SOME (c, trns) => (SOME c, trns) |
90 |
| NONE => (NONE, trns); |
|
91 |
||
25969 | 92 |
fun generate thy funcgr f x = |
93 |
Program.change_yield thy (CodeThingol.transact thy funcgr |
|
94 |
(fn thy => fn funcgr => fn algbr => f thy funcgr algbr x)); |
|
24219 | 95 |
|
26113 | 96 |
(* export_code functionality *) |
97 |
||
24436 | 98 |
fun code thy permissive cs seris = |
99 |
let |
|
100 |
val code = Program.get thy; |
|
27000 | 101 |
fun mk_seri_dest file = case file |
102 |
of NONE => CodeTarget.compile |
|
27001 | 103 |
| SOME "-" => CodeTarget.write |
27000 | 104 |
| SOME f => CodeTarget.file (Path.explode f) |
105 |
val _ = map (fn (((target, module), file), args) => |
|
106 |
(mk_seri_dest file (CodeTarget.serialize thy target permissive module args code cs))) seris; |
|
107 |
in () end; |
|
24436 | 108 |
|
27014
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
109 |
(* code retrieval *) |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
110 |
|
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
111 |
fun source_of thy target module_name cs = |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
112 |
let |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
113 |
val (cs', _) = generate thy (CodeFuncgr.make thy cs) |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
114 |
(fold_map ooo ensure_const) cs; |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
115 |
val code = Program.get thy; |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
116 |
in |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
117 |
CodeTarget.string |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
118 |
(CodeTarget.serialize thy target false (SOME module_name) [] code (SOME cs')) |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
119 |
end; |
a5f53d9d2b60
yet another attempt to circumvent printmode problems
haftmann
parents:
27001
diff
changeset
|
120 |
|
26113 | 121 |
(* evaluation machinery *) |
122 |
||
26740 | 123 |
fun evaluate eval_kind thy evaluator = |
24250 | 124 |
let |
26740 | 125 |
fun evaluator'' evaluator''' funcgr t = |
126 |
let |
|
127 |
val ((code, (vs_ty_t, deps)), _) = |
|
128 |
generate thy funcgr CodeThingol.ensure_value t; |
|
129 |
in evaluator''' code vs_ty_t deps end; |
|
130 |
fun evaluator' t = |
|
131 |
let |
|
132 |
val (t', evaluator''') = evaluator t; |
|
133 |
in (t', evaluator'' evaluator''') end; |
|
134 |
in eval_kind thy evaluator' end |
|
26011 | 135 |
|
26740 | 136 |
fun evaluate_conv thy = evaluate CodeFuncgr.eval_conv thy; |
137 |
fun evaluate_term thy = evaluate CodeFuncgr.eval_term thy; |
|
26011 | 138 |
|
26740 | 139 |
fun eval_ml reff args thy code ((vs, ty), t) deps = |
26011 | 140 |
CodeTarget.eval thy reff code (t, ty) args; |
24219 | 141 |
|
26011 | 142 |
fun eval evaluate term_of reff thy ct args = |
143 |
let |
|
144 |
val _ = if null (term_frees (term_of ct)) then () else error ("Term " |
|
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26740
diff
changeset
|
145 |
^ quote (Syntax.string_of_term_global thy (term_of ct)) |
26740 | 146 |
^ " to be evaluated contains free variables"); |
147 |
in evaluate thy (fn t => (t, eval_ml reff args thy)) ct end; |
|
26011 | 148 |
|
149 |
fun eval_conv reff = eval evaluate_conv Thm.term_of reff; |
|
150 |
fun eval_term reff = eval evaluate_term I reff; |
|
24219 | 151 |
|
24585 | 152 |
val satisfies_ref : (unit -> bool) option ref = ref NONE; |
24219 | 153 |
|
26011 | 154 |
val satisfies = eval_term ("CodePackage.satisfies_ref", satisfies_ref); |
24219 | 155 |
|
26113 | 156 |
(* code antiquotation *) |
157 |
||
158 |
fun code_antiq (ctxt, args) = |
|
159 |
let |
|
160 |
val thy = Context.theory_of ctxt; |
|
161 |
val (ts, (ctxt', args')) = Scan.repeat1 Args.term (ctxt, args); |
|
162 |
val cs = map (CodeUnit.check_const thy) ts; |
|
163 |
val (cs', code') = generate thy (CodeFuncgr.make thy cs) |
|
164 |
(fold_map ooo ensure_const) cs; |
|
27000 | 165 |
val code'' = CodeTarget.sml_of thy code' cs' ^ " ()"; |
26113 | 166 |
in (("codevals", code''), (ctxt', args')) end; |
167 |
||
168 |
||
169 |
(* const expressions *) |
|
170 |
||
26615
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
171 |
local |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
172 |
|
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
173 |
fun consts_of thy some_thyname = |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
174 |
let |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
175 |
val this_thy = Option.map ThyInfo.get_theory some_thyname |> the_default thy; |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
176 |
val raw_cs = Symtab.fold (fn (c, (_, NONE)) => cons c | _ => I) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
177 |
((snd o #constants o Consts.dest o #consts o Sign.rep_sg) this_thy) []; |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
178 |
val cs = map (CodeUnit.subst_alias thy) raw_cs; |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
179 |
fun belongs_here thyname c = |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
180 |
not (exists (fn thy' => Sign.declared_const thy' c) (Theory.parents_of this_thy)) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
181 |
in case some_thyname |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
182 |
of NONE => cs |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
183 |
| SOME thyname => filter (belongs_here thyname) cs |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
184 |
end; |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
185 |
|
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
186 |
fun read_const_expr thy "*" = ([], consts_of thy NONE) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
187 |
| read_const_expr thy s = if String.isSuffix ".*" s |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
188 |
then ([], consts_of thy (SOME (unsuffix ".*" s))) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
189 |
else ([CodeUnit.read_const thy s], []); |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
190 |
|
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
191 |
in |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
192 |
|
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
193 |
fun read_const_exprs thy select exprs = |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
194 |
case (pairself flat o split_list o map (read_const_expr thy)) exprs |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
195 |
of (consts, []) => (false, consts) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
196 |
| (consts, consts') => (true, consts @ select consts'); |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
197 |
|
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
198 |
end; (*local*) |
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
199 |
|
24219 | 200 |
fun filter_generatable thy consts = |
201 |
let |
|
24283 | 202 |
val (consts', funcgr) = CodeFuncgr.make_consts thy consts; |
24918 | 203 |
val (consts'', _) = generate thy funcgr (fold_map ooo perhaps_const) consts'; |
24219 | 204 |
val consts''' = map_filter (fn (const, SOME _) => SOME const | (_, NONE) => NONE) |
205 |
(consts' ~~ consts''); |
|
206 |
in consts''' end; |
|
207 |
||
24436 | 208 |
fun generate_const_exprs thy raw_cs = |
209 |
let |
|
26615
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
210 |
val (perm1, cs) = read_const_exprs thy |
24436 | 211 |
(filter_generatable thy) raw_cs; |
212 |
val (perm2, cs') = case generate thy (CodeFuncgr.make thy cs) |
|
24918 | 213 |
(fold_map ooo ensure_const) cs |
214 |
of ([], _) => (true, NONE) |
|
215 |
| (cs, _) => (false, SOME cs); |
|
24436 | 216 |
in (perm1 orelse perm2, cs') end; |
217 |
||
218 |
||
219 |
(** code properties **) |
|
220 |
||
221 |
fun mk_codeprops thy all_cs sel_cs = |
|
222 |
let |
|
24976
821628d16552
moved Drule.unvarify to Thm.unvarify (cf. more_thm.ML);
wenzelm
parents:
24971
diff
changeset
|
223 |
fun select (thmref, thm) = case try (Thm.unvarify o Drule.zero_var_indexes) thm |
24436 | 224 |
of NONE => NONE |
225 |
| SOME thm => let |
|
226 |
val t = (ObjectLogic.drop_judgment thy o Thm.prop_of) thm; |
|
227 |
val cs = fold_aterms (fn Const (c, ty) => |
|
25597
34860182b250
moved instance parameter management from class.ML to axclass.ML
haftmann
parents:
25485
diff
changeset
|
228 |
cons (AxClass.unoverload_const thy (c, ty)) | _ => I) t []; |
24436 | 229 |
in if exists (member (op =) sel_cs) cs |
230 |
andalso forall (member (op =) all_cs) cs |
|
231 |
then SOME (thmref, thm) else NONE end; |
|
232 |
fun mk_codeprop (thmref, thm) = |
|
233 |
let |
|
234 |
val t = ObjectLogic.drop_judgment thy (Thm.prop_of thm); |
|
235 |
val ty_judg = fastype_of t; |
|
236 |
val tfrees1 = fold_aterms (fn Const (c, ty) => |
|
237 |
Term.add_tfreesT ty | _ => I) t []; |
|
238 |
val vars = Term.add_frees t []; |
|
239 |
val tfrees2 = fold (Term.add_tfreesT o snd) vars []; |
|
240 |
val tfrees' = subtract (op =) tfrees2 tfrees1 |> map TFree; |
|
241 |
val ty = map Term.itselfT tfrees' @ map snd vars ---> ty_judg; |
|
242 |
val tfree_vars = map Logic.mk_type tfrees'; |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26285
diff
changeset
|
243 |
val c = Facts.string_of_ref thmref |
24436 | 244 |
|> NameSpace.explode |
245 |
|> (fn [x] => [x] | (x::xs) => xs) |
|
246 |
|> space_implode "_" |
|
247 |
val propdef = (((c, ty), tfree_vars @ map Free vars), t); |
|
248 |
in if c = "" then NONE else SOME (thmref, propdef) end; |
|
249 |
in |
|
26690 | 250 |
Facts.dest_static (PureThy.facts_of thy) |
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26285
diff
changeset
|
251 |
|> maps Facts.selections |
24436 | 252 |
|> map_filter select |
253 |
|> map_filter mk_codeprop |
|
254 |
end; |
|
255 |
||
256 |
fun add_codeprops all_cs sel_cs thy = |
|
257 |
let |
|
258 |
val codeprops = mk_codeprops thy all_cs sel_cs; |
|
259 |
fun lift_name_yield f x = (Name.context, x) |> f ||> snd; |
|
260 |
fun add (thmref, (((raw_c, ty), ts), t)) (names, thy) = |
|
261 |
let |
|
26336
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents:
26285
diff
changeset
|
262 |
val _ = warning ("Adding theorem " ^ quote (Facts.string_of_ref thmref) |
24436 | 263 |
^ " as code property " ^ quote raw_c); |
264 |
val ([raw_c'], names') = Name.variants [raw_c] names; |
|
24971
4d006b03aa4a
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24918
diff
changeset
|
265 |
val (const as Const (c, _), thy') = thy |> Sign.declare_const [] (raw_c', ty, NoSyn); |
4d006b03aa4a
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24918
diff
changeset
|
266 |
val eq = Logic.mk_equals (list_comb (const, ts), t); |
4d006b03aa4a
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24918
diff
changeset
|
267 |
val ([def], thy'') = thy' |> PureThy.add_defs_i false [((Thm.def_name raw_c', eq), [])]; |
4d006b03aa4a
replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents:
24918
diff
changeset
|
268 |
in ((c, def), (names', thy'')) end; |
24436 | 269 |
in |
270 |
thy |
|
271 |
|> Sign.sticky_prefix "codeprop" |
|
272 |
|> lift_name_yield (fold_map add codeprops) |
|
273 |
||> Sign.restore_naming thy |
|
24621 | 274 |
|-> (fn c_thms => fold (Code.add_func o snd) c_thms #> pair c_thms) |
24436 | 275 |
end; |
276 |
||
277 |
||
24918 | 278 |
(** interfaces and Isar setup **) |
24219 | 279 |
|
280 |
local |
|
281 |
||
282 |
structure P = OuterParse |
|
283 |
and K = OuterKeyword |
|
284 |
||
24436 | 285 |
fun code_cmd raw_cs seris thy = |
24219 | 286 |
let |
24436 | 287 |
val (permissive, cs) = generate_const_exprs thy raw_cs; |
288 |
val _ = code thy permissive cs seris; |
|
289 |
in () end; |
|
24219 | 290 |
|
291 |
fun code_thms_cmd thy = |
|
26615
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
292 |
code_thms thy o snd o read_const_exprs thy (fst o CodeFuncgr.make_consts thy); |
24219 | 293 |
|
294 |
fun code_deps_cmd thy = |
|
26615
8a9d3eebd534
added read_const_exprs (from Pure/Isar/code_unit.ML);
wenzelm
parents:
26604
diff
changeset
|
295 |
code_deps thy o snd o read_const_exprs thy (fst o CodeFuncgr.make_consts thy); |
24219 | 296 |
|
24436 | 297 |
fun code_props_cmd raw_cs seris thy = |
298 |
let |
|
299 |
val (_, all_cs) = generate_const_exprs thy ["*"]; |
|
300 |
val (permissive, cs) = generate_const_exprs thy raw_cs; |
|
301 |
val (c_thms, thy') = add_codeprops (map (the o CodeName.const_rev thy) (these all_cs)) |
|
302 |
(map (the o CodeName.const_rev thy) (these cs)) thy; |
|
303 |
val prop_cs = (filter_generatable thy' o map fst) c_thms; |
|
24918 | 304 |
val _ = if null seris then () else (generate thy' (CodeFuncgr.make thy' prop_cs) |
305 |
(fold_map ooo ensure_const) prop_cs; ()); |
|
24436 | 306 |
val _ = if null seris then () else code thy' permissive |
307 |
(SOME (map (CodeName.const thy') prop_cs)) seris; |
|
308 |
in thy' end; |
|
309 |
||
24250 | 310 |
val (inK, module_nameK, fileK) = ("in", "module_name", "file"); |
24219 | 311 |
|
24436 | 312 |
fun code_exprP cmd = |
24219 | 313 |
(Scan.repeat P.term |
314 |
-- Scan.repeat (P.$$$ inK |-- P.name |
|
24250 | 315 |
-- Scan.option (P.$$$ module_nameK |-- P.name) |
24219 | 316 |
-- Scan.option (P.$$$ fileK |-- P.name) |
317 |
-- Scan.optional (P.$$$ "(" |-- P.arguments --| P.$$$ ")") [] |
|
24436 | 318 |
) >> (fn (raw_cs, seris) => cmd raw_cs seris)); |
24219 | 319 |
|
24867 | 320 |
val _ = OuterSyntax.keywords [inK, module_nameK, fileK]; |
24219 | 321 |
|
24436 | 322 |
val (codeK, code_thmsK, code_depsK, code_propsK) = |
323 |
("export_code", "code_thms", "code_deps", "code_props"); |
|
24219 | 324 |
|
325 |
in |
|
326 |
||
24867 | 327 |
val _ = |
25110 | 328 |
OuterSyntax.command codeK "generate executable code for constants" |
24436 | 329 |
K.diag (P.!!! (code_exprP code_cmd) >> (fn f => Toplevel.keep (f o Toplevel.theory_of))); |
24219 | 330 |
|
26604 | 331 |
fun codegen_shell_command thyname cmd = Toplevel.program (fn _ => |
25611 | 332 |
(use_thy thyname; case Scan.read OuterLex.stopper (P.!!! (code_exprP code_cmd)) ((filter OuterLex.is_proper o OuterSyntax.scan) cmd) |
333 |
of SOME f => (writeln "Now generating code..."; f (theory thyname)) |
|
334 |
| NONE => error ("Bad directive " ^ quote cmd))) |
|
335 |
handle TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure; |
|
24219 | 336 |
|
24867 | 337 |
val _ = |
24219 | 338 |
OuterSyntax.improper_command code_thmsK "print system of defining equations for code" OuterKeyword.diag |
339 |
(Scan.repeat P.term |
|
340 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
341 |
o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); |
|
342 |
||
24867 | 343 |
val _ = |
24219 | 344 |
OuterSyntax.improper_command code_depsK "visualize dependencies of defining equations for code" OuterKeyword.diag |
345 |
(Scan.repeat P.term |
|
346 |
>> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory |
|
347 |
o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); |
|
348 |
||
24867 | 349 |
val _ = |
24436 | 350 |
OuterSyntax.command code_propsK "generate characteristic properties for executable constants" |
351 |
K.thy_decl (P.!!! (code_exprP code_props_cmd) >> Toplevel.theory); |
|
24219 | 352 |
|
26113 | 353 |
val _ = ML_Context.value_antiq "code" code_antiq; |
354 |
||
24436 | 355 |
end; (*local*) |
356 |
||
357 |
end; (*struct*) |