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