| author | wenzelm | 
| Sat, 30 Mar 2013 18:24:33 +0100 | |
| changeset 51592 | c3a7d6592e3f | 
| parent 51590 | c52891242de2 | 
| child 51688 | 27ecd33d3366 | 
| permissions | -rw-r--r-- | 
| 16014 | 1 | (* Title: Pure/simplifier.ML | 
| 2 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 3 | ||
| 4 | Generic simplifier, suitable for most logics (see also | |
| 41386 | 5 | raw_simplifier.ML for the actual meta-level rewriting engine). | 
| 16014 | 6 | *) | 
| 7 | ||
| 8 | signature BASIC_SIMPLIFIER = | |
| 9 | sig | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 10 | include BASIC_RAW_SIMPLIFIER | 
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 11 | val map_simpset: (simpset -> simpset) -> Proof.context -> Proof.context | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 12 | val simpset_of: Proof.context -> simpset | 
| 32148 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
32091diff
changeset | 13 | val global_simpset_of: theory -> simpset | 
| 16014 | 14 | val Addsimprocs: simproc list -> unit | 
| 15 | val Delsimprocs: simproc list -> unit | |
| 35613 | 16 | val simp_tac: simpset -> int -> tactic | 
| 17 | val asm_simp_tac: simpset -> int -> tactic | |
| 18 | val full_simp_tac: simpset -> int -> tactic | |
| 19 | val asm_lr_simp_tac: simpset -> int -> tactic | |
| 20 | val asm_full_simp_tac: simpset -> int -> tactic | |
| 50107 | 21 | val safe_simp_tac: simpset -> int -> tactic | 
| 22 | val safe_asm_simp_tac: simpset -> int -> tactic | |
| 23 | val safe_full_simp_tac: simpset -> int -> tactic | |
| 24 | val safe_asm_lr_simp_tac: simpset -> int -> tactic | |
| 25 | val safe_asm_full_simp_tac: simpset -> int -> tactic | |
| 35613 | 26 | val simplify: simpset -> thm -> thm | 
| 27 | val asm_simplify: simpset -> thm -> thm | |
| 28 | val full_simplify: simpset -> thm -> thm | |
| 29 | val asm_lr_simplify: simpset -> thm -> thm | |
| 16014 | 30 | val asm_full_simplify: simpset -> thm -> thm | 
| 31 | end; | |
| 32 | ||
| 33 | signature SIMPLIFIER = | |
| 34 | sig | |
| 35 | include BASIC_SIMPLIFIER | |
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 36 | val map_simpset_global: (simpset -> simpset) -> theory -> theory | 
| 30356 | 37 | val pretty_ss: Proof.context -> simpset -> Pretty.T | 
| 17004 | 38 | val clear_ss: simpset -> simpset | 
| 47239 | 39 | val default_mk_sym: simpset -> thm -> thm option | 
| 32738 | 40 | val debug_bounds: bool Unsynchronized.ref | 
| 43597 | 41 | val prems_of: simpset -> thm list | 
| 41226 | 42 | val add_simp: thm -> simpset -> simpset | 
| 43 | val del_simp: thm -> simpset -> simpset | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 44 | val add_eqcong: thm -> simpset -> simpset | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 45 | val del_eqcong: thm -> simpset -> simpset | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 46 | val add_cong: thm -> simpset -> simpset | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 47 | val del_cong: thm -> simpset -> simpset | 
| 41226 | 48 | val add_prems: thm list -> simpset -> simpset | 
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 49 | val mksimps: simpset -> thm -> thm list | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 50 | val set_mksimps: (simpset -> thm -> thm list) -> simpset -> simpset | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 51 | val set_mkcong: (simpset -> thm -> thm) -> simpset -> simpset | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 52 | val set_mksym: (simpset -> thm -> thm option) -> simpset -> simpset | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 53 | val set_mkeqTrue: (simpset -> thm -> thm option) -> simpset -> simpset | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 54 | val set_termless: (term * term -> bool) -> simpset -> simpset | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 55 | val set_subgoaler: (simpset -> int -> tactic) -> simpset -> simpset | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 56 | val inherit_context: simpset -> simpset -> simpset | 
| 17898 | 57 | val the_context: simpset -> Proof.context | 
| 58 | val context: Proof.context -> simpset -> simpset | |
| 37441 | 59 | val global_context: theory -> simpset -> simpset | 
| 36600 | 60 | val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset | 
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 61 | val simproc_global_i: theory -> string -> term list -> | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 62 | (theory -> simpset -> term -> thm option) -> simproc | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 63 | val simproc_global: theory -> string -> string list -> | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 64 | (theory -> simpset -> term -> thm option) -> simproc | 
| 35613 | 65 | val rewrite: simpset -> conv | 
| 66 | val asm_rewrite: simpset -> conv | |
| 67 | val full_rewrite: simpset -> conv | |
| 68 | val asm_lr_rewrite: simpset -> conv | |
| 23598 | 69 | val asm_full_rewrite: simpset -> conv | 
| 22379 | 70 | val get_ss: Context.generic -> simpset | 
| 71 | val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 72 | val attrib: (thm -> simpset -> simpset) -> attribute | 
| 18728 | 73 | val simp_add: attribute | 
| 74 | val simp_del: attribute | |
| 75 | val cong_add: attribute | |
| 76 | val cong_del: attribute | |
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 77 | val check_simproc: Proof.context -> xstring * Position.T -> string | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 78 | val the_simproc: Proof.context -> string -> simproc | 
| 42464 | 79 |   val def_simproc: {name: binding, lhss: term list,
 | 
| 22236 | 80 | proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> | 
| 22201 | 81 | local_theory -> local_theory | 
| 42464 | 82 |   val def_simproc_cmd: {name: binding, lhss: string list,
 | 
| 22236 | 83 | proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> | 
| 22201 | 84 | local_theory -> local_theory | 
| 30513 | 85 | val cong_modifiers: Method.modifier parser list | 
| 86 | val simp_modifiers': Method.modifier parser list | |
| 87 | val simp_modifiers: Method.modifier parser list | |
| 88 | val method_setup: Method.modifier parser list -> theory -> theory | |
| 18708 | 89 | val easy_setup: thm -> thm list -> theory -> theory | 
| 16014 | 90 | end; | 
| 91 | ||
| 92 | structure Simplifier: SIMPLIFIER = | |
| 93 | struct | |
| 94 | ||
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 95 | open Raw_Simplifier; | 
| 21708 | 96 | |
| 97 | ||
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 98 | (** data **) | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 99 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 100 | structure Data = Generic_Data | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 101 | ( | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 102 | type T = simpset * simproc Name_Space.table; | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 103 | val empty : T = (empty_ss, Name_Space.empty_table "simproc"); | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 104 | fun extend (ss, tab) = (Raw_Simplifier.inherit_context empty_ss ss, tab); | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 105 | fun merge ((ss1, tab1), (ss2, tab2)) = | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 106 | (merge_ss (ss1, ss2), Name_Space.merge_tables (tab1, tab2)); | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 107 | ); | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 108 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 109 | val get_ss = fst o Data.get; | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 110 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 111 | fun map_ss f context = | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 112 | Data.map (apfst ((Raw_Simplifier.with_context (Context.proof_of context) f))) context; | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 113 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 114 | val get_simprocs = snd o Data.get o Context.Proof; | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 115 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 116 | |
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 117 | |
| 30356 | 118 | (** pretty printing **) | 
| 119 | ||
| 120 | fun pretty_ss ctxt ss = | |
| 121 | let | |
| 51590 | 122 | val pretty_term = Syntax.pretty_term ctxt; | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31300diff
changeset | 123 | val pretty_thm = Display.pretty_thm ctxt; | 
| 51584 | 124 | val pretty_thm_item = Display.pretty_thm_item ctxt; | 
| 125 | ||
| 51580 
64ef8260dc60
Pretty.item markup for improved readability of lists of items;
 wenzelm parents: 
50107diff
changeset | 126 | fun pretty_proc (name, lhss) = | 
| 51590 | 127 | Pretty.big_list (name ^ ":") (map (Pretty.item o single o pretty_term o Thm.term_of) lhss); | 
| 128 | ||
| 129 | fun pretty_cong_name (const, name) = | |
| 130 | pretty_term ((if const then Const else Free) (name, dummyT)); | |
| 30356 | 131 | fun pretty_cong (name, thm) = | 
| 51590 | 132 | Pretty.block [pretty_cong_name name, Pretty.str ":", Pretty.brk 1, pretty_thm thm]; | 
| 30356 | 133 | |
| 134 |     val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss;
 | |
| 135 | in | |
| 51584 | 136 | [Pretty.big_list "simplification rules:" (map (pretty_thm_item o #2) simps), | 
| 30356 | 137 | Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)), | 
| 138 | Pretty.big_list "congruences:" (map pretty_cong congs), | |
| 139 |       Pretty.strs ("loopers:" :: map quote loopers),
 | |
| 140 |       Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
 | |
| 141 |       Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
 | |
| 142 | |> Pretty.chunks | |
| 143 | end; | |
| 144 | ||
| 145 | ||
| 146 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 147 | (** simpset data **) | 
| 16014 | 148 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 149 | (* attributes *) | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 150 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 151 | fun attrib f = Thm.declaration_attribute (map_ss o f); | 
| 16014 | 152 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 153 | val simp_add = attrib add_simp; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 154 | val simp_del = attrib del_simp; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 155 | val cong_add = attrib add_cong; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 156 | val cong_del = attrib del_cong; | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 157 | |
| 16014 | 158 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 159 | (* local simpset *) | 
| 16014 | 160 | |
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 161 | fun map_simpset f = Context.proof_map (map_ss f); | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 162 | fun simpset_of ctxt = Raw_Simplifier.context ctxt (get_ss (Context.Proof ctxt)); | 
| 16014 | 163 | |
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 164 | val _ = Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 165 | (ML_Antiquote.value (Binding.name "simpset") | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47468diff
changeset | 166 | (Scan.succeed "Simplifier.simpset_of ML_context"))); | 
| 22132 | 167 | |
| 16014 | 168 | |
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 169 | (* global simpset *) | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 170 | |
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 171 | fun map_simpset_global f = Context.theory_map (map_ss f); | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 172 | fun global_simpset_of thy = | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 173 | Raw_Simplifier.context (Proof_Context.init_global thy) (get_ss (Context.Theory thy)); | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 174 | |
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 175 | fun Addsimprocs args = Context.>> (map_ss (fn ss => ss addsimprocs args)); | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 176 | fun Delsimprocs args = Context.>> (map_ss (fn ss => ss delsimprocs args)); | 
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 177 | |
| 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 178 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 179 | |
| 22201 | 180 | (** named simprocs **) | 
| 181 | ||
| 22204 | 182 | (* get simprocs *) | 
| 183 | ||
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
47001diff
changeset | 184 | fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1; | 
| 42466 | 185 | val the_simproc = Name_Space.get o get_simprocs; | 
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 186 | |
| 42464 | 187 | val _ = | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 188 | Context.>> (Context.map_theory | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 189 | (ML_Antiquote.value (Binding.name "simproc") | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 190 | (Args.context -- Scan.lift (Parse.position Args.name) | 
| 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 191 | >> (fn (ctxt, name) => | 
| 48776 
37cd53e69840
faster compilation of ML with antiquotations: static ML_context is bound once in auxiliary structure Isabelle;
 wenzelm parents: 
47468diff
changeset | 192 | "Simplifier.the_simproc ML_context " ^ | 
| 43560 
d1650e3720fd
ML antiquotations are managed as theory data, with proper name space and entity markup;
 wenzelm parents: 
42795diff
changeset | 193 | ML_Syntax.print_string (check_simproc ctxt name))))); | 
| 22204 | 194 | |
| 195 | ||
| 196 | (* define simprocs *) | |
| 22201 | 197 | |
| 198 | local | |
| 199 | ||
| 42464 | 200 | fun gen_simproc prep {name = b, lhss, proc, identifier} lthy =
 | 
| 22201 | 201 | let | 
| 22236 | 202 | val simproc = make_simproc | 
| 47001 
a0e370d3d149
proper naming of simprocs according to actual target context;
 wenzelm parents: 
46776diff
changeset | 203 |       {name = Local_Theory.full_name lthy b,
 | 
| 22236 | 204 | lhss = | 
| 205 | let | |
| 206 | val lhss' = prep lthy lhss; | |
| 45326 
8fa859aebc0d
tuned -- Variable.declare_term is already part of Variable.auto_fixes;
 wenzelm parents: 
45291diff
changeset | 207 | val ctxt' = fold Variable.auto_fixes lhss' lthy; | 
| 22236 | 208 | in Variable.export_terms ctxt' lthy lhss' end | 
| 42360 | 209 | |> map (Thm.cterm_of (Proof_Context.theory_of lthy)), | 
| 22236 | 210 | proc = proc, | 
| 33551 
c40ced05b10a
define simprocs: do not apply target_morphism prematurely, this is already done in LocalTheory.declaration;
 wenzelm parents: 
33519diff
changeset | 211 | identifier = identifier}; | 
| 22201 | 212 | in | 
| 47001 
a0e370d3d149
proper naming of simprocs according to actual target context;
 wenzelm parents: 
46776diff
changeset | 213 |     lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context =>
 | 
| 22201 | 214 | let | 
| 28991 | 215 | val b' = Morphism.binding phi b; | 
| 45290 | 216 | val simproc' = transform_simproc phi simproc; | 
| 22201 | 217 | in | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 218 | context | 
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 219 | |> Data.map (fn (ss, tab) => | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
47001diff
changeset | 220 | (ss addsimprocs [simproc'], #2 (Name_Space.define context true (b', simproc') tab))) | 
| 22201 | 221 | end) | 
| 222 | end; | |
| 223 | ||
| 224 | in | |
| 225 | ||
| 42464 | 226 | val def_simproc = gen_simproc Syntax.check_terms; | 
| 227 | val def_simproc_cmd = gen_simproc Syntax.read_terms; | |
| 22201 | 228 | |
| 229 | end; | |
| 230 | ||
| 231 | ||
| 232 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 233 | (** simplification tactics and rules **) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 234 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 235 | fun solve_all_tac solvers ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 236 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 237 |     val (_, {subgoal_tac, ...}) = Raw_Simplifier.internal_ss ss;
 | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 238 | val solve_tac = subgoal_tac (Raw_Simplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 239 | in DEPTH_SOLVE (solve_tac 1) end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 240 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 241 | (*NOTE: may instantiate unknowns that appear also in other subgoals*) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 242 | fun generic_simp_tac safe mode ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 243 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 244 |     val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = Raw_Simplifier.internal_ss ss;
 | 
| 21286 
b5e7b80caa6a
introduces canonical AList functions for loop_tacs
 haftmann parents: 
20872diff
changeset | 245 | val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs)); | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 246 | val solve_tac = FIRST' (map (Raw_Simplifier.solver ss) | 
| 22717 | 247 | (rev (if safe then solvers else unsafe_solvers))); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 248 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 249 | fun simp_loop_tac i = | 
| 46465 | 250 | Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 251 | (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); | 
| 47468 
402b753d8383
outermost SELECT_GOAL potentially improves performance;
 wenzelm parents: 
47239diff
changeset | 252 | in SELECT_GOAL (simp_loop_tac 1) end; | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 253 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 254 | local | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 255 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 256 | fun simp rew mode ss thm = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 257 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 258 |     val (_, {solvers = (unsafe_solvers, _), ...}) = Raw_Simplifier.internal_ss ss;
 | 
| 22717 | 259 | val tacf = solve_all_tac (rev unsafe_solvers); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 260 | fun prover s th = Option.map #1 (Seq.pull (tacf s th)); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 261 | in rew mode prover ss thm end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 262 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 263 | in | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 264 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 265 | val simp_thm = simp Raw_Simplifier.rewrite_thm; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 266 | val simp_cterm = simp Raw_Simplifier.rewrite_cterm; | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 267 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 268 | end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 269 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 270 | |
| 16806 | 271 | (* tactics *) | 
| 272 | ||
| 16014 | 273 | val simp_tac = generic_simp_tac false (false, false, false); | 
| 274 | val asm_simp_tac = generic_simp_tac false (false, true, false); | |
| 275 | val full_simp_tac = generic_simp_tac false (true, false, false); | |
| 276 | val asm_lr_simp_tac = generic_simp_tac false (true, true, false); | |
| 277 | val asm_full_simp_tac = generic_simp_tac false (true, true, true); | |
| 50107 | 278 | |
| 279 | (*not totally safe: may instantiate unknowns that appear also in other subgoals*) | |
| 280 | val safe_simp_tac = generic_simp_tac true (false, false, false); | |
| 281 | val safe_asm_simp_tac = generic_simp_tac true (false, true, false); | |
| 282 | val safe_full_simp_tac = generic_simp_tac true (true, false, false); | |
| 283 | val safe_asm_lr_simp_tac = generic_simp_tac true (true, true, false); | |
| 16014 | 284 | val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); | 
| 285 | ||
| 16806 | 286 | |
| 287 | (* conversions *) | |
| 288 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 289 | val simplify = simp_thm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 290 | val asm_simplify = simp_thm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 291 | val full_simplify = simp_thm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 292 | val asm_lr_simplify = simp_thm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 293 | val asm_full_simplify = simp_thm (true, true, true); | 
| 16014 | 294 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 295 | val rewrite = simp_cterm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 296 | val asm_rewrite = simp_cterm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 297 | val full_rewrite = simp_cterm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 298 | val asm_lr_rewrite = simp_cterm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 299 | val asm_full_rewrite = simp_cterm (true, true, true); | 
| 16014 | 300 | |
| 301 | ||
| 302 | ||
| 303 | (** concrete syntax of attributes **) | |
| 304 | ||
| 305 | (* add / del *) | |
| 306 | ||
| 307 | val simpN = "simp"; | |
| 308 | val congN = "cong"; | |
| 309 | val onlyN = "only"; | |
| 310 | val no_asmN = "no_asm"; | |
| 311 | val no_asm_useN = "no_asm_use"; | |
| 312 | val no_asm_simpN = "no_asm_simp"; | |
| 313 | val asm_lrN = "asm_lr"; | |
| 314 | ||
| 315 | ||
| 24024 | 316 | (* simprocs *) | 
| 317 | ||
| 318 | local | |
| 319 | ||
| 320 | val add_del = | |
| 321 | (Args.del -- Args.colon >> K (op delsimprocs) || | |
| 322 | Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) | |
| 323 | >> (fn f => fn simproc => fn phi => Thm.declaration_attribute | |
| 45290 | 324 | (K (map_ss (fn ss => f (ss, [transform_simproc phi simproc]))))); | 
| 24024 | 325 | |
| 326 | in | |
| 327 | ||
| 30528 | 328 | val simproc_att = | 
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 329 | (Args.context -- Scan.lift add_del) :|-- (fn (ctxt, decl) => | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 330 | Scan.repeat1 (Scan.lift (Args.named_attribute (decl o the_simproc ctxt o check_simproc ctxt)))) | 
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
45326diff
changeset | 331 | >> (fn atts => Thm.declaration_attribute (fn th => | 
| 46776 | 332 | fold (fn att => Thm.attribute_declaration (Morphism.form att) th) atts)); | 
| 24024 | 333 | |
| 334 | end; | |
| 24124 
4399175e3014
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24024diff
changeset | 335 | |
| 24024 | 336 | |
| 16014 | 337 | (* conversions *) | 
| 338 | ||
| 339 | local | |
| 340 | ||
| 341 | fun conv_mode x = | |
| 342 | ((Args.parens (Args.$$$ no_asmN) >> K simplify || | |
| 343 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || | |
| 344 | Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || | |
| 345 | Scan.succeed asm_full_simplify) |> Scan.lift) x; | |
| 346 | ||
| 347 | in | |
| 348 | ||
| 30528 | 349 | val simplified = conv_mode -- Attrib.thms >> | 
| 350 | (fn (f, ths) => Thm.rule_attribute (fn context => | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 351 | f ((if null ths then I else Raw_Simplifier.clear_ss) | 
| 32148 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
32091diff
changeset | 352 | (simpset_of (Context.proof_of context)) addsimps ths))); | 
| 16014 | 353 | |
| 354 | end; | |
| 355 | ||
| 356 | ||
| 357 | (* setup attributes *) | |
| 358 | ||
| 26463 | 359 | val _ = Context.>> (Context.map_theory | 
| 30528 | 360 | (Attrib.setup (Binding.name simpN) (Attrib.add_del simp_add simp_del) | 
| 361 | "declaration of Simplifier rewrite rule" #> | |
| 362 | Attrib.setup (Binding.name congN) (Attrib.add_del cong_add cong_del) | |
| 363 | "declaration of Simplifier congruence rule" #> | |
| 33671 | 364 | Attrib.setup (Binding.name "simproc") simproc_att | 
| 365 | "declaration of simplification procedures" #> | |
| 30528 | 366 | Attrib.setup (Binding.name "simplified") simplified "simplified rule")); | 
| 16014 | 367 | |
| 368 | ||
| 369 | ||
| 31300 | 370 | (** method syntax **) | 
| 16014 | 371 | |
| 372 | val cong_modifiers = | |
| 18728 | 373 | [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier), | 
| 374 | Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add), | |
| 375 | Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)]; | |
| 16014 | 376 | |
| 377 | val simp_modifiers = | |
| 18728 | 378 | [Args.$$$ simpN -- Args.colon >> K (I, simp_add), | 
| 379 | Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add), | |
| 380 | Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del), | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 381 | Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 382 | >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] | 
| 16014 | 383 | @ cong_modifiers; | 
| 384 | ||
| 385 | val simp_modifiers' = | |
| 18728 | 386 | [Args.add -- Args.colon >> K (I, simp_add), | 
| 387 | Args.del -- Args.colon >> K (I, simp_del), | |
| 18688 | 388 | Args.$$$ onlyN -- Args.colon | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 389 | >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] | 
| 16014 | 390 | @ cong_modifiers; | 
| 391 | ||
| 31300 | 392 | val simp_options = | 
| 393 | (Args.parens (Args.$$$ no_asmN) >> K simp_tac || | |
| 394 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || | |
| 395 | Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || | |
| 396 | Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || | |
| 397 | Scan.succeed asm_full_simp_tac); | |
| 16014 | 398 | |
| 31300 | 399 | fun simp_method more_mods meth = | 
| 35613 | 400 | Scan.lift simp_options --| | 
| 31300 | 401 | Method.sections (more_mods @ simp_modifiers') >> | 
| 35613 | 402 | (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts)); | 
| 16014 | 403 | |
| 404 | ||
| 405 | ||
| 18708 | 406 | (** setup **) | 
| 407 | ||
| 31300 | 408 | fun method_setup more_mods = | 
| 409 | Method.setup (Binding.name simpN) | |
| 410 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | |
| 411 | HEADGOAL (Method.insert_tac facts THEN' | |
| 32148 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
32091diff
changeset | 412 | (CHANGED_PROP oo tac) (simpset_of ctxt)))) | 
| 31300 | 413 | "simplification" #> | 
| 414 | Method.setup (Binding.name "simp_all") | |
| 415 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | |
| 416 | ALLGOALS (Method.insert_tac facts) THEN | |
| 42372 | 417 | (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) (simpset_of ctxt))) | 
| 31300 | 418 | "simplification (all goals)"; | 
| 16014 | 419 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 420 | fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ => | 
| 16014 | 421 | let | 
| 422 | val trivialities = Drule.reflexive_thm :: trivs; | |
| 423 | ||
| 43597 | 424 | fun unsafe_solver_tac ss = | 
| 425 | FIRST' [resolve_tac (trivialities @ Raw_Simplifier.prems_of ss), assume_tac]; | |
| 16014 | 426 | val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; | 
| 427 | ||
| 428 | (*no premature instantiation of variables during simplification*) | |
| 43597 | 429 | fun safe_solver_tac ss = | 
| 430 | FIRST' [match_tac (trivialities @ Raw_Simplifier.prems_of ss), eq_assume_tac]; | |
| 16014 | 431 | val safe_solver = mk_solver "easy safe" safe_solver_tac; | 
| 432 | ||
| 433 | fun mk_eq thm = | |
| 20872 | 434 | if can Logic.dest_equals (Thm.concl_of thm) then [thm] | 
| 16014 | 435 | else [thm RS reflect] handle THM _ => []; | 
| 436 | ||
| 44058 | 437 | fun mksimps thm = mk_eq (Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm); | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 438 | in | 
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 439 | empty_ss | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 440 | setSSolver safe_solver | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 441 | setSolver unsafe_solver | 
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 442 | |> set_subgoaler asm_simp_tac | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 443 | |> set_mksimps (K mksimps) | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 444 | end)); | 
| 16014 | 445 | |
| 446 | end; | |
| 447 | ||
| 32738 | 448 | structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier; | 
| 449 | open Basic_Simplifier; |