| author | wenzelm | 
| Thu, 08 Sep 2022 20:46:22 +0200 | |
| changeset 76090 | f8eff19a3825 | 
| parent 74561 | 8e6c973003c8 | 
| 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  | 
| 53147 | 17  | 
val trace: bool Config.T  | 
| 37442 | 18  | 
end;  | 
19  | 
||
20  | 
structure Code_Simp : CODE_SIMP =  | 
|
21  | 
struct  | 
|
22  | 
||
23  | 
(* dedicated simpset *)  | 
|
24  | 
||
| 
38759
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38671 
diff
changeset
 | 
25  | 
structure Simpset = Theory_Data  | 
| 
 
37a9092de102
simplification/standardization of some theory data;
 
wenzelm 
parents: 
38671 
diff
changeset
 | 
26  | 
(  | 
| 37442 | 27  | 
type T = simpset;  | 
28  | 
val empty = empty_ss;  | 
|
29  | 
val merge = merge_ss;  | 
|
30  | 
);  | 
|
31  | 
||
| 63164 | 32  | 
fun map_ss f thy =  | 
33  | 
Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy;  | 
|
| 37442 | 34  | 
|
| 63164 | 35  | 
fun simpset_default ctxt some_ss =  | 
36  | 
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  | 
|
| 69593 | 41  | 
val trace = Attrib.setup_config_bool \<^binding>\<open>code_simp_trace\<close> (K false);  | 
| 53147 | 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  | 
|
| 63164 | 64  | 
val simpset_program =  | 
65  | 
Code_Preproc.timed "building simpset" #ctxt  | 
|
66  | 
  (fn { ctxt, some_ss, program } => simpset_map ctxt (add_program program) (simpset_default ctxt some_ss));
 | 
|
| 37461 | 67  | 
|
| 55757 | 68  | 
fun rewrite_modulo ctxt some_ss program =  | 
69  | 
let  | 
|
| 63164 | 70  | 
val ss = simpset_program  | 
71  | 
      { ctxt = ctxt, some_ss = some_ss, program = program };
 | 
|
72  | 
in fn ctxt =>  | 
|
73  | 
Code_Preproc.timed_conv "simplifying"  | 
|
74  | 
Simplifier.full_rewrite (ctxt |> put_simpset ss |> set_trace)  | 
|
75  | 
end;  | 
|
| 
52736
 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 
haftmann 
parents: 
51717 
diff
changeset
 | 
76  | 
|
| 55757 | 77  | 
fun conclude_tac ctxt some_ss =  | 
78  | 
let  | 
|
79  | 
val ss = simpset_default ctxt some_ss  | 
|
80  | 
in fn ctxt => Simplifier.full_simp_tac (ctxt |> put_simpset ss) end;  | 
|
| 37442 | 81  | 
|
82  | 
||
| 39486 | 83  | 
(* evaluation with dynamic code context *)  | 
| 37442 | 84  | 
|
| 55757 | 85  | 
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
 | 
86  | 
(fn program => fn _ => fn _ => rewrite_modulo ctxt NONE program ctxt);  | 
| 37442 | 87  | 
|
| 55757 | 88  | 
fun dynamic_tac ctxt = CONVERSION (dynamic_conv ctxt)  | 
89  | 
THEN' conclude_tac ctxt NONE ctxt;  | 
|
| 37442 | 90  | 
|
| 55757 | 91  | 
fun dynamic_value ctxt =  | 
| 59633 | 92  | 
snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.cterm_of ctxt;  | 
| 37444 | 93  | 
|
| 59323 | 94  | 
val _ = Theory.setup  | 
| 69593 | 95  | 
(Method.setup \<^binding>\<open>code_simp\<close>  | 
| 55757 | 96  | 
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac)))  | 
| 59323 | 97  | 
"simplification with code equations");  | 
| 37442 | 98  | 
|
99  | 
||
| 39486 | 100  | 
(* evaluation with static code context *)  | 
| 37442 | 101  | 
|
| 
56973
 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 
haftmann 
parents: 
56969 
diff
changeset
 | 
102  | 
fun static_conv { ctxt, simpset, consts } =
 | 
| 63156 | 103  | 
  Code_Thingol.static_conv_isa { ctxt = ctxt, consts = consts }
 | 
| 63164 | 104  | 
(fn program => let  | 
105  | 
val conv = rewrite_modulo ctxt simpset program  | 
|
106  | 
in fn ctxt => fn _ => conv ctxt end);  | 
|
| 37442 | 107  | 
|
| 
56973
 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 
haftmann 
parents: 
56969 
diff
changeset
 | 
108  | 
fun static_tac { ctxt, simpset, consts } =
 | 
| 
52736
 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 
haftmann 
parents: 
51717 
diff
changeset
 | 
109  | 
let  | 
| 
56973
 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 
haftmann 
parents: 
56969 
diff
changeset
 | 
110  | 
    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
 | 
111  | 
val tac = conclude_tac ctxt simpset;  | 
| 55757 | 112  | 
in fn ctxt' => CONVERSION (conv ctxt') THEN' (tac ctxt') end;  | 
| 37442 | 113  | 
|
114  | 
end;  |