author | wenzelm |
Mon, 21 Jul 2014 16:58:12 +0200 | |
changeset 57593 | 2f7d91242b99 |
parent 56973 | 62da80041afd |
child 59323 | 468bd3aedfa1 |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: Tools/Code/code_simp.ML |
37442 | 2 |
Author: Florian Haftmann, TU Muenchen |
3 |
||
4 |
Connecting the simplifier and the code generator. |
|
5 |
*) |
|
6 |
||
7 |
signature CODE_SIMP = |
|
8 |
sig |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48072
diff
changeset
|
9 |
val map_ss: (Proof.context -> Proof.context) -> theory -> theory |
55757 | 10 |
val dynamic_conv: Proof.context -> conv |
11 |
val dynamic_tac: Proof.context -> int -> tactic |
|
12 |
val dynamic_value: Proof.context -> term -> term |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
13 |
val static_conv: { ctxt: Proof.context, simpset: simpset option, consts: string list } |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
14 |
-> Proof.context -> conv |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
15 |
val static_tac: { ctxt: Proof.context, simpset: simpset option, consts: string list } |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
16 |
-> Proof.context -> int -> tactic |
37442 | 17 |
val setup: theory -> theory |
53147 | 18 |
val trace: bool Config.T |
37442 | 19 |
end; |
20 |
||
21 |
structure Code_Simp : CODE_SIMP = |
|
22 |
struct |
|
23 |
||
24 |
(* dedicated simpset *) |
|
25 |
||
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38671
diff
changeset
|
26 |
structure Simpset = Theory_Data |
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38671
diff
changeset
|
27 |
( |
37442 | 28 |
type T = simpset; |
29 |
val empty = empty_ss; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48072
diff
changeset
|
30 |
val extend = I; |
37442 | 31 |
val merge = merge_ss; |
32 |
); |
|
33 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48072
diff
changeset
|
34 |
fun map_ss f thy = Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy; |
37442 | 35 |
|
55757 | 36 |
fun simpset_default ctxt some_ss = the_default (Simpset.get (Proof_Context.theory_of ctxt)) some_ss; |
54929
f1ded3cea58d
proper context for simplifier invocations in code generation stack
haftmann
parents:
54928
diff
changeset
|
37 |
|
37442 | 38 |
|
54928 | 39 |
(* diagnostic *) |
53147 | 40 |
|
41 |
val trace = Attrib.setup_config_bool @{binding "code_simp_trace"} (K false); |
|
42 |
||
43 |
fun set_trace ctxt = |
|
44 |
let |
|
45 |
val global = Config.get ctxt trace; |
|
46 |
in |
|
47 |
ctxt |> Config.map Simplifier.simp_trace (fn b => b orelse global) |
|
48 |
end; |
|
49 |
||
50 |
||
55757 | 51 |
(* build simpset context and conversion from program *) |
37442 | 52 |
|
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
54929
diff
changeset
|
53 |
fun add_stmt (Code_Thingol.Fun ((_, eqs), some_cong)) ss = |
52736
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
54 |
ss |
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
55 |
|> fold Simplifier.add_simp ((map_filter (fst o snd)) eqs) |
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
43619
diff
changeset
|
56 |
|> fold Simplifier.add_cong (the_list some_cong) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47576
diff
changeset
|
57 |
| add_stmt (Code_Thingol.Classinst { inst_params, ... }) ss = |
52736
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
58 |
ss |
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
59 |
|> fold Simplifier.add_simp (map (fst o snd) inst_params) |
37442 | 60 |
| add_stmt _ ss = ss; |
61 |
||
55147
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
haftmann
parents:
54929
diff
changeset
|
62 |
val add_program = Code_Symbol.Graph.fold (add_stmt o fst o snd); |
37442 | 63 |
|
55757 | 64 |
fun simpset_program ctxt some_ss program = |
65 |
simpset_map ctxt (add_program program) (simpset_default ctxt some_ss); |
|
37461 | 66 |
|
55757 | 67 |
fun rewrite_modulo ctxt some_ss program = |
68 |
let |
|
69 |
val ss = simpset_program ctxt some_ss program; |
|
70 |
in fn ctxt => Simplifier.full_rewrite (ctxt |> put_simpset ss |> set_trace) end; |
|
52736
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
71 |
|
55757 | 72 |
fun conclude_tac ctxt some_ss = |
73 |
let |
|
74 |
val ss = simpset_default ctxt some_ss |
|
75 |
in fn ctxt => Simplifier.full_simp_tac (ctxt |> put_simpset ss) end; |
|
37442 | 76 |
|
77 |
||
39486 | 78 |
(* evaluation with dynamic code context *) |
37442 | 79 |
|
55757 | 80 |
fun dynamic_conv ctxt = Code_Thingol.dynamic_conv ctxt |
56969
7491932da574
dropped obsolete hand-waving adjustion of type variables: safely done in preprocessor
haftmann
parents:
56924
diff
changeset
|
81 |
(fn program => fn _ => fn _ => rewrite_modulo ctxt NONE program ctxt); |
37442 | 82 |
|
55757 | 83 |
fun dynamic_tac ctxt = CONVERSION (dynamic_conv ctxt) |
84 |
THEN' conclude_tac ctxt NONE ctxt; |
|
37442 | 85 |
|
55757 | 86 |
fun dynamic_value ctxt = |
87 |
snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.cterm_of (Proof_Context.theory_of ctxt); |
|
37444 | 88 |
|
43619
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
89 |
val setup = |
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
90 |
Method.setup @{binding code_simp} |
55757 | 91 |
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac))) |
56924
2f94c9a50f06
dimiss simplified as evaluator due to little practical relevance
haftmann
parents:
56920
diff
changeset
|
92 |
"simplification with code equations"; |
37442 | 93 |
|
94 |
||
39486 | 95 |
(* evaluation with static code context *) |
37442 | 96 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
97 |
fun static_conv { ctxt, simpset, consts } = |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
98 |
Code_Thingol.static_conv_simple { ctxt = ctxt, consts = consts } |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
99 |
(K oo rewrite_modulo ctxt simpset); |
37442 | 100 |
|
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
101 |
fun static_tac { ctxt, simpset, consts } = |
52736
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
haftmann
parents:
51717
diff
changeset
|
102 |
let |
56973
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
103 |
val conv = static_conv { ctxt = ctxt, simpset = simpset, consts = consts }; |
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
haftmann
parents:
56969
diff
changeset
|
104 |
val tac = conclude_tac ctxt simpset; |
55757 | 105 |
in fn ctxt' => CONVERSION (conv ctxt') THEN' (tac ctxt') end; |
37442 | 106 |
|
107 |
end; |