| author | wenzelm | 
| Sun, 29 Mar 2020 21:32:20 +0200 | |
| changeset 71622 | ab5009192ebb | 
| parent 69593 | 3dda49e08b9d | 
| child 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: 
48072diff
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: 
56969diff
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: 
56969diff
changeset | 14 | -> Proof.context -> conv | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
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: 
56969diff
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: 
38671diff
changeset | 25 | structure Simpset = Theory_Data | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38671diff
changeset | 26 | ( | 
| 37442 | 27 | type T = simpset; | 
| 28 | val empty = empty_ss; | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
48072diff
changeset | 29 | val extend = I; | 
| 37442 | 30 | val merge = merge_ss; | 
| 31 | ); | |
| 32 | ||
| 63164 | 33 | fun map_ss f thy = | 
| 34 | Simpset.map (simpset_map (Proof_Context.init_global thy) f) thy; | |
| 37442 | 35 | |
| 63164 | 36 | fun simpset_default ctxt some_ss = | 
| 37 | the_default (Simpset.get (Proof_Context.theory_of ctxt)) some_ss; | |
| 54929 
f1ded3cea58d
proper context for simplifier invocations in code generation stack
 haftmann parents: 
54928diff
changeset | 38 | |
| 37442 | 39 | |
| 54928 | 40 | (* diagnostic *) | 
| 53147 | 41 | |
| 69593 | 42 | val trace = Attrib.setup_config_bool \<^binding>\<open>code_simp_trace\<close> (K false); | 
| 53147 | 43 | |
| 44 | fun set_trace ctxt = | |
| 45 | let | |
| 46 | val global = Config.get ctxt trace; | |
| 47 | in | |
| 48 | ctxt |> Config.map Simplifier.simp_trace (fn b => b orelse global) | |
| 49 | end; | |
| 50 | ||
| 51 | ||
| 55757 | 52 | (* build simpset context and conversion from program *) | 
| 37442 | 53 | |
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54929diff
changeset | 54 | 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: 
51717diff
changeset | 55 | ss | 
| 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 haftmann parents: 
51717diff
changeset | 56 | |> 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: 
43619diff
changeset | 57 | |> fold Simplifier.add_cong (the_list some_cong) | 
| 48072 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 haftmann parents: 
47576diff
changeset | 58 |   | add_stmt (Code_Thingol.Classinst { inst_params, ... }) ss =
 | 
| 52736 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 haftmann parents: 
51717diff
changeset | 59 | ss | 
| 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 haftmann parents: 
51717diff
changeset | 60 | |> fold Simplifier.add_simp (map (fst o snd) inst_params) | 
| 37442 | 61 | | add_stmt _ ss = ss; | 
| 62 | ||
| 55147 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 haftmann parents: 
54929diff
changeset | 63 | val add_program = Code_Symbol.Graph.fold (add_stmt o fst o snd); | 
| 37442 | 64 | |
| 63164 | 65 | val simpset_program = | 
| 66 | Code_Preproc.timed "building simpset" #ctxt | |
| 67 |   (fn { ctxt, some_ss, program } => simpset_map ctxt (add_program program) (simpset_default ctxt some_ss));
 | |
| 37461 | 68 | |
| 55757 | 69 | fun rewrite_modulo ctxt some_ss program = | 
| 70 | let | |
| 63164 | 71 | val ss = simpset_program | 
| 72 |       { ctxt = ctxt, some_ss = some_ss, program = program };
 | |
| 73 | in fn ctxt => | |
| 74 | Code_Preproc.timed_conv "simplifying" | |
| 75 | Simplifier.full_rewrite (ctxt |> put_simpset ss |> set_trace) | |
| 76 | end; | |
| 52736 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 haftmann parents: 
51717diff
changeset | 77 | |
| 55757 | 78 | fun conclude_tac ctxt some_ss = | 
| 79 | let | |
| 80 | val ss = simpset_default ctxt some_ss | |
| 81 | in fn ctxt => Simplifier.full_simp_tac (ctxt |> put_simpset ss) end; | |
| 37442 | 82 | |
| 83 | ||
| 39486 | 84 | (* evaluation with dynamic code context *) | 
| 37442 | 85 | |
| 55757 | 86 | 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: 
56924diff
changeset | 87 | (fn program => fn _ => fn _ => rewrite_modulo ctxt NONE program ctxt); | 
| 37442 | 88 | |
| 55757 | 89 | fun dynamic_tac ctxt = CONVERSION (dynamic_conv ctxt) | 
| 90 | THEN' conclude_tac ctxt NONE ctxt; | |
| 37442 | 91 | |
| 55757 | 92 | fun dynamic_value ctxt = | 
| 59633 | 93 | snd o Logic.dest_equals o Thm.prop_of o dynamic_conv ctxt o Thm.cterm_of ctxt; | 
| 37444 | 94 | |
| 59323 | 95 | val _ = Theory.setup | 
| 69593 | 96 | (Method.setup \<^binding>\<open>code_simp\<close> | 
| 55757 | 97 | (Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo dynamic_tac))) | 
| 59323 | 98 | "simplification with code equations"); | 
| 37442 | 99 | |
| 100 | ||
| 39486 | 101 | (* evaluation with static code context *) | 
| 37442 | 102 | |
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 103 | fun static_conv { ctxt, simpset, consts } =
 | 
| 63156 | 104 |   Code_Thingol.static_conv_isa { ctxt = ctxt, consts = consts }
 | 
| 63164 | 105 | (fn program => let | 
| 106 | val conv = rewrite_modulo ctxt simpset program | |
| 107 | in fn ctxt => fn _ => conv ctxt end); | |
| 37442 | 108 | |
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 109 | fun static_tac { ctxt, simpset, consts } =
 | 
| 52736 
317663b422bb
more correct context for dynamic invocations of static code conversions etc.
 haftmann parents: 
51717diff
changeset | 110 | let | 
| 56973 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 111 |     val conv = static_conv { ctxt = ctxt, simpset = simpset, consts = consts };
 | 
| 
62da80041afd
syntactic means to prevent accidental mixup of static and dynamic context
 haftmann parents: 
56969diff
changeset | 112 | val tac = conclude_tac ctxt simpset; | 
| 55757 | 113 | in fn ctxt' => CONVERSION (conv ctxt') THEN' (tac ctxt') end; | 
| 37442 | 114 | |
| 115 | end; |