author | blanchet |
Mon, 18 Jun 2012 17:50:06 +0200 | |
changeset 48104 | d2173ff80c57 |
parent 48072 | ace701efe203 |
child 51717 | 9e7d1c139569 |
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 |
|
9 |
val map_ss: (simpset -> simpset) -> theory -> theory |
|
41247 | 10 |
val dynamic_conv: theory -> conv |
41188 | 11 |
val dynamic_tac: theory -> int -> tactic |
41184 | 12 |
val dynamic_value: theory -> term -> term |
13 |
val static_conv: theory -> simpset option -> string list -> conv |
|
41188 | 14 |
val static_tac: theory -> simpset option -> string list -> int -> tactic |
37442 | 15 |
val setup: theory -> theory |
16 |
end; |
|
17 |
||
18 |
structure Code_Simp : CODE_SIMP = |
|
19 |
struct |
|
20 |
||
21 |
(* dedicated simpset *) |
|
22 |
||
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38671
diff
changeset
|
23 |
structure Simpset = Theory_Data |
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38671
diff
changeset
|
24 |
( |
37442 | 25 |
type T = simpset; |
26 |
val empty = empty_ss; |
|
41225 | 27 |
fun extend ss = Simplifier.inherit_context empty_ss ss; |
37442 | 28 |
val merge = merge_ss; |
29 |
); |
|
30 |
||
31 |
val map_ss = Simpset.map; |
|
32 |
||
37461 | 33 |
fun simpset_default thy = Simplifier.global_context thy o the_default (Simpset.get thy); |
34 |
||
37442 | 35 |
|
36 |
(* build simpset and conversion from program *) |
|
37 |
||
38 |
fun add_stmt (Code_Thingol.Fun (_, ((_, eqs), some_cong))) ss = |
|
45620
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
43619
diff
changeset
|
39 |
ss addsimps (map_filter (fst o snd)) eqs |
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
43619
diff
changeset
|
40 |
|> fold Simplifier.add_cong (the_list some_cong) |
48072
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47576
diff
changeset
|
41 |
| add_stmt (Code_Thingol.Classinst { inst_params, ... }) ss = |
ace701efe203
prefer records with speaking labels over deeply nested tuples
haftmann
parents:
47576
diff
changeset
|
42 |
ss addsimps (map (fst o snd) inst_params) |
37442 | 43 |
| add_stmt _ ss = ss; |
44 |
||
37839 | 45 |
val add_program = Graph.fold (add_stmt o fst o snd); |
37442 | 46 |
|
47 |
fun rewrite_modulo thy some_ss program = Simplifier.full_rewrite |
|
37461 | 48 |
(add_program program (simpset_default thy some_ss)); |
49 |
||
50 |
fun conclude_tac thy some_ss = Simplifier.full_simp_tac (simpset_default thy some_ss); |
|
37442 | 51 |
|
52 |
||
39486 | 53 |
(* evaluation with dynamic code context *) |
37442 | 54 |
|
41247 | 55 |
fun dynamic_conv thy = Code_Thingol.dynamic_conv thy |
47576 | 56 |
(fn _ => fn program => fn _ => fn _ => rewrite_modulo thy NONE program); |
37442 | 57 |
|
41247 | 58 |
fun dynamic_tac thy = CONVERSION (dynamic_conv thy) THEN' conclude_tac thy NONE; |
37442 | 59 |
|
41247 | 60 |
fun dynamic_value thy = snd o Logic.dest_equals o Thm.prop_of o dynamic_conv thy o Thm.cterm_of thy; |
37444 | 61 |
|
43619
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
62 |
val setup = |
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
63 |
Method.setup @{binding code_simp} |
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
64 |
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac o Proof_Context.theory_of))) |
3803869014aa
proper @{binding} antiquotations (relevant for formal references);
wenzelm
parents:
42361
diff
changeset
|
65 |
"simplification with code equations" |
42361 | 66 |
#> Value.add_evaluator ("simp", dynamic_value o Proof_Context.theory_of); |
37442 | 67 |
|
68 |
||
39486 | 69 |
(* evaluation with static code context *) |
37442 | 70 |
|
41184 | 71 |
fun static_conv thy some_ss consts = |
72 |
Code_Thingol.static_conv_simple thy consts |
|
41346 | 73 |
(fn program => fn _ => fn _ => rewrite_modulo thy some_ss program); |
37442 | 74 |
|
41188 | 75 |
fun static_tac thy some_ss consts = CONVERSION (static_conv thy some_ss consts) |
37461 | 76 |
THEN' conclude_tac thy some_ss; |
37442 | 77 |
|
78 |
end; |