| author | bulwahn | 
| Mon, 18 Apr 2011 09:10:23 +0200 | |
| changeset 42391 | d7b58dc35cc5 | 
| parent 42375 | 774df7c59508 | 
| child 42464 | ae16b8abf1a8 | 
| 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 | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 11 | val change_simpset: (simpset -> simpset) -> unit | 
| 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 | 12 | val global_simpset_of: theory -> simpset | 
| 16014 | 13 | val Addsimprocs: simproc list -> unit | 
| 14 | val Delsimprocs: simproc list -> unit | |
| 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 | 15 | val simpset_of: Proof.context -> simpset | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 16 | val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic | 
| 16014 | 17 | val safe_asm_full_simp_tac: simpset -> int -> tactic | 
| 35613 | 18 | val simp_tac: simpset -> int -> tactic | 
| 19 | val asm_simp_tac: simpset -> int -> tactic | |
| 20 | val full_simp_tac: simpset -> int -> tactic | |
| 21 | val asm_lr_simp_tac: simpset -> int -> tactic | |
| 22 | val asm_full_simp_tac: simpset -> int -> tactic | |
| 23 | val simplify: simpset -> thm -> thm | |
| 24 | val asm_simplify: simpset -> thm -> thm | |
| 25 | val full_simplify: simpset -> thm -> thm | |
| 26 | val asm_lr_simplify: simpset -> thm -> thm | |
| 16014 | 27 | val asm_full_simplify: simpset -> thm -> thm | 
| 28 | end; | |
| 29 | ||
| 30 | signature SIMPLIFIER = | |
| 31 | sig | |
| 32 | include BASIC_SIMPLIFIER | |
| 30356 | 33 | val pretty_ss: Proof.context -> simpset -> Pretty.T | 
| 17004 | 34 | val clear_ss: simpset -> simpset | 
| 32738 | 35 | val debug_bounds: bool Unsynchronized.ref | 
| 41226 | 36 | val add_simp: thm -> simpset -> simpset | 
| 37 | val del_simp: thm -> simpset -> simpset | |
| 38 | val add_prems: thm list -> simpset -> simpset | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 39 | val inherit_context: simpset -> simpset -> simpset | 
| 17898 | 40 | val the_context: simpset -> Proof.context | 
| 41 | val context: Proof.context -> simpset -> simpset | |
| 37441 | 42 | val global_context: theory -> simpset -> simpset | 
| 36600 | 43 | val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37441diff
changeset | 44 | val simproc_global_i: theory -> string -> term list | 
| 16458 | 45 | -> (theory -> simpset -> term -> thm option) -> simproc | 
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37441diff
changeset | 46 | val simproc_global: theory -> string -> string list | 
| 16458 | 47 | -> (theory -> simpset -> term -> thm option) -> simproc | 
| 35613 | 48 | val rewrite: simpset -> conv | 
| 49 | val asm_rewrite: simpset -> conv | |
| 50 | val full_rewrite: simpset -> conv | |
| 51 | val asm_lr_rewrite: simpset -> conv | |
| 23598 | 52 | val asm_full_rewrite: simpset -> conv | 
| 22379 | 53 | val get_ss: Context.generic -> simpset | 
| 54 | val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic | |
| 18728 | 55 | val attrib: (simpset * thm list -> simpset) -> attribute | 
| 56 | val simp_add: attribute | |
| 57 | val simp_del: attribute | |
| 58 | val cong_add: attribute | |
| 59 | val cong_del: attribute | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 60 | val map_simpset: (simpset -> simpset) -> theory -> theory | 
| 24024 | 61 | val get_simproc: Context.generic -> xstring -> simproc | 
| 22236 | 62 |   val def_simproc: {name: string, lhss: string list,
 | 
| 63 | proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> | |
| 22201 | 64 | local_theory -> local_theory | 
| 22236 | 65 |   val def_simproc_i: {name: string, lhss: term list,
 | 
| 66 | proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} -> | |
| 22201 | 67 | local_theory -> local_theory | 
| 30513 | 68 | val cong_modifiers: Method.modifier parser list | 
| 69 | val simp_modifiers': Method.modifier parser list | |
| 70 | val simp_modifiers: Method.modifier parser list | |
| 71 | val method_setup: Method.modifier parser list -> theory -> theory | |
| 18708 | 72 | val easy_setup: thm -> thm list -> theory -> theory | 
| 16014 | 73 | end; | 
| 74 | ||
| 75 | structure Simplifier: SIMPLIFIER = | |
| 76 | struct | |
| 77 | ||
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 78 | open Raw_Simplifier; | 
| 21708 | 79 | |
| 80 | ||
| 30356 | 81 | (** pretty printing **) | 
| 82 | ||
| 83 | fun pretty_ss ctxt ss = | |
| 84 | let | |
| 85 | val pretty_cterm = Syntax.pretty_term ctxt o Thm.term_of; | |
| 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 | 86 | val pretty_thm = Display.pretty_thm ctxt; | 
| 30356 | 87 | fun pretty_proc (name, lhss) = Pretty.big_list (name ^ ":") (map pretty_cterm lhss); | 
| 88 | fun pretty_cong (name, thm) = | |
| 89 | Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, pretty_thm thm]; | |
| 90 | ||
| 91 |     val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss;
 | |
| 92 | in | |
| 93 | [Pretty.big_list "simplification rules:" (map (pretty_thm o #2) simps), | |
| 94 | Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)), | |
| 95 | Pretty.big_list "congruences:" (map pretty_cong congs), | |
| 96 |       Pretty.strs ("loopers:" :: map quote loopers),
 | |
| 97 |       Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
 | |
| 98 |       Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
 | |
| 99 | |> Pretty.chunks | |
| 100 | end; | |
| 101 | ||
| 102 | ||
| 103 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 104 | (** simpset data **) | 
| 16014 | 105 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36610diff
changeset | 106 | structure Simpset = Generic_Data | 
| 22846 | 107 | ( | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 108 | type T = simpset; | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 109 | val empty = empty_ss; | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 110 | fun extend ss = Raw_Simplifier.inherit_context empty_ss ss; | 
| 33519 | 111 | val merge = merge_ss; | 
| 22846 | 112 | ); | 
| 16014 | 113 | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36610diff
changeset | 114 | val get_ss = Simpset.get; | 
| 41253 
42f24340ae53
more explicit references to structure Raw_Simplifier;
 wenzelm parents: 
41228diff
changeset | 115 | |
| 
42f24340ae53
more explicit references to structure Raw_Simplifier;
 wenzelm parents: 
41228diff
changeset | 116 | fun map_ss f context = | 
| 
42f24340ae53
more explicit references to structure Raw_Simplifier;
 wenzelm parents: 
41228diff
changeset | 117 | Simpset.map (Raw_Simplifier.with_context (Context.proof_of context) f) context; | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 118 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 119 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 120 | (* attributes *) | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 121 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 122 | fun attrib f = Thm.declaration_attribute (fn th => map_ss (fn ss => f (ss, [th]))); | 
| 16014 | 123 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 124 | val simp_add = attrib (op addsimps); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 125 | val simp_del = attrib (op delsimps); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 126 | val cong_add = attrib (op addcongs); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 127 | val cong_del = attrib (op delcongs); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 128 | |
| 16014 | 129 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 130 | (* global simpset *) | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 131 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 132 | fun map_simpset f = Context.theory_map (map_ss f); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 133 | fun change_simpset f = Context.>> (Context.map_theory (map_simpset f)); | 
| 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 | 134 | fun global_simpset_of thy = | 
| 42360 | 135 | Raw_Simplifier.context (Proof_Context.init_global thy) (get_ss (Context.Theory thy)); | 
| 16014 | 136 | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 137 | fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 138 | fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args); | 
| 16014 | 139 | |
| 140 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 141 | (* local simpset *) | 
| 16014 | 142 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 143 | fun simpset_of ctxt = Raw_Simplifier.context ctxt (get_ss (Context.Proof ctxt)); | 
| 16014 | 144 | |
| 27338 | 145 | val _ = ML_Antiquote.value "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 | 146 | (Scan.succeed "Simplifier.simpset_of (ML_Context.the_local_context ())"); | 
| 22132 | 147 | |
| 16014 | 148 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 149 | |
| 22201 | 150 | (** named simprocs **) | 
| 151 | ||
| 22204 | 152 | (* data *) | 
| 153 | ||
| 33519 | 154 | structure Simprocs = Generic_Data | 
| 22201 | 155 | ( | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 156 | type T = simproc Name_Space.table; | 
| 33159 | 157 | val empty : T = Name_Space.empty_table "simproc"; | 
| 22201 | 158 | val extend = I; | 
| 33519 | 159 | fun merge simprocs = Name_Space.merge_tables simprocs; | 
| 22201 | 160 | ); | 
| 161 | ||
| 22204 | 162 | |
| 163 | (* get simprocs *) | |
| 164 | ||
| 24024 | 165 | fun get_simproc context xname = | 
| 22204 | 166 | let | 
| 24024 | 167 | val (space, tab) = Simprocs.get context; | 
| 33095 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 wenzelm parents: 
33092diff
changeset | 168 | val name = Name_Space.intern space xname; | 
| 22204 | 169 | in | 
| 170 | (case Symtab.lookup tab name of | |
| 22236 | 171 | SOME proc => proc | 
| 22204 | 172 |     | NONE => error ("Undefined simplification procedure: " ^ quote name))
 | 
| 173 | end; | |
| 174 | ||
| 27338 | 175 | val _ = ML_Antiquote.value "simproc" (Scan.lift Args.name >> (fn name => | 
| 176 | "Simplifier.get_simproc (ML_Context.the_generic_context ()) " ^ ML_Syntax.print_string name)); | |
| 22204 | 177 | |
| 178 | ||
| 179 | (* define simprocs *) | |
| 22201 | 180 | |
| 181 | local | |
| 182 | ||
| 22236 | 183 | fun gen_simproc prep {name, lhss, proc, identifier} lthy =
 | 
| 22201 | 184 | let | 
| 28991 | 185 | val b = Binding.name name; | 
| 33671 | 186 | val naming = Local_Theory.naming_of lthy; | 
| 22236 | 187 | val simproc = make_simproc | 
| 33169 | 188 |       {name = Name_Space.full_name naming b,
 | 
| 22236 | 189 | lhss = | 
| 190 | let | |
| 191 | val lhss' = prep lthy lhss; | |
| 192 | val ctxt' = lthy | |
| 193 | |> fold Variable.declare_term lhss' | |
| 194 | |> fold Variable.auto_fixes lhss'; | |
| 195 | in Variable.export_terms ctxt' lthy lhss' end | |
| 42360 | 196 | |> map (Thm.cterm_of (Proof_Context.theory_of lthy)), | 
| 22236 | 197 | proc = proc, | 
| 33551 
c40ced05b10a
define simprocs: do not apply target_morphism prematurely, this is already done in LocalTheory.declaration;
 wenzelm parents: 
33519diff
changeset | 198 | identifier = identifier}; | 
| 22201 | 199 | in | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 200 | lthy |> Local_Theory.declaration false (fn phi => fn context => | 
| 22201 | 201 | let | 
| 28991 | 202 | val b' = Morphism.binding phi b; | 
| 22236 | 203 | val simproc' = morph_simproc phi simproc; | 
| 22201 | 204 | in | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 205 | context | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 206 | |> Simprocs.map | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 207 | (#2 o Name_Space.define (Context.proof_of context) true naming (b', simproc')) | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 208 | |> map_ss (fn ss => ss addsimprocs [simproc']) | 
| 22201 | 209 | end) | 
| 210 | end; | |
| 211 | ||
| 212 | in | |
| 213 | ||
| 24509 
23ee6b7788c2
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24124diff
changeset | 214 | val def_simproc = gen_simproc Syntax.read_terms; | 
| 
23ee6b7788c2
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 wenzelm parents: 
24124diff
changeset | 215 | val def_simproc_i = gen_simproc Syntax.check_terms; | 
| 22201 | 216 | |
| 217 | end; | |
| 218 | ||
| 219 | ||
| 220 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 221 | (** simplification tactics and rules **) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 222 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 223 | fun solve_all_tac solvers ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 224 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 225 |     val (_, {subgoal_tac, ...}) = Raw_Simplifier.internal_ss ss;
 | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 226 | 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 | 227 | in DEPTH_SOLVE (solve_tac 1) end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 228 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 229 | (*NOTE: may instantiate unknowns that appear also in other subgoals*) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 230 | fun generic_simp_tac safe mode ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 231 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 232 |     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 | 233 | 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 | 234 | val solve_tac = FIRST' (map (Raw_Simplifier.solver ss) | 
| 22717 | 235 | (rev (if safe then solvers else unsafe_solvers))); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 236 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 237 | fun simp_loop_tac i = | 
| 41253 
42f24340ae53
more explicit references to structure Raw_Simplifier;
 wenzelm parents: 
41228diff
changeset | 238 | Raw_Simplifier.asm_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 | 239 | (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 240 | in simp_loop_tac end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 241 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 242 | local | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 243 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 244 | fun simp rew mode ss thm = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 245 | let | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 246 |     val (_, {solvers = (unsafe_solvers, _), ...}) = Raw_Simplifier.internal_ss ss;
 | 
| 22717 | 247 | val tacf = solve_all_tac (rev unsafe_solvers); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 248 | 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 | 249 | in rew mode prover ss thm end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 250 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 251 | in | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 252 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 253 | val simp_thm = simp Raw_Simplifier.rewrite_thm; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 254 | val simp_cterm = simp Raw_Simplifier.rewrite_cterm; | 
| 17967 
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 | end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 257 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 258 | |
| 16806 | 259 | (* tactics *) | 
| 260 | ||
| 16014 | 261 | val simp_tac = generic_simp_tac false (false, false, false); | 
| 262 | val asm_simp_tac = generic_simp_tac false (false, true, false); | |
| 263 | val full_simp_tac = generic_simp_tac false (true, false, false); | |
| 264 | val asm_lr_simp_tac = generic_simp_tac false (true, true, false); | |
| 265 | val asm_full_simp_tac = generic_simp_tac false (true, true, true); | |
| 266 | val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); | |
| 267 | ||
| 16806 | 268 | |
| 269 | (* conversions *) | |
| 270 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 271 | val simplify = simp_thm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 272 | val asm_simplify = simp_thm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 273 | val full_simplify = simp_thm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 274 | val asm_lr_simplify = simp_thm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 275 | val asm_full_simplify = simp_thm (true, true, true); | 
| 16014 | 276 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 277 | val rewrite = simp_cterm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 278 | val asm_rewrite = simp_cterm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 279 | val full_rewrite = simp_cterm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 280 | val asm_lr_rewrite = simp_cterm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 281 | val asm_full_rewrite = simp_cterm (true, true, true); | 
| 16014 | 282 | |
| 283 | ||
| 284 | ||
| 285 | (** concrete syntax of attributes **) | |
| 286 | ||
| 287 | (* add / del *) | |
| 288 | ||
| 289 | val simpN = "simp"; | |
| 290 | val congN = "cong"; | |
| 291 | val onlyN = "only"; | |
| 292 | val no_asmN = "no_asm"; | |
| 293 | val no_asm_useN = "no_asm_use"; | |
| 294 | val no_asm_simpN = "no_asm_simp"; | |
| 295 | val asm_lrN = "asm_lr"; | |
| 296 | ||
| 297 | ||
| 24024 | 298 | (* simprocs *) | 
| 299 | ||
| 300 | local | |
| 301 | ||
| 302 | val add_del = | |
| 303 | (Args.del -- Args.colon >> K (op delsimprocs) || | |
| 304 | Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) | |
| 305 | >> (fn f => fn simproc => fn phi => Thm.declaration_attribute | |
| 306 | (K (map_ss (fn ss => f (ss, [morph_simproc phi simproc]))))); | |
| 307 | ||
| 308 | in | |
| 309 | ||
| 30528 | 310 | val simproc_att = | 
| 311 | Scan.peek (fn context => | |
| 24024 | 312 | add_del :|-- (fn decl => | 
| 313 | Scan.repeat1 (Args.named_attribute (decl o get_simproc context)) | |
| 30528 | 314 | >> (Library.apply o map Morphism.form))); | 
| 24024 | 315 | |
| 316 | end; | |
| 24124 
4399175e3014
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24024diff
changeset | 317 | |
| 24024 | 318 | |
| 16014 | 319 | (* conversions *) | 
| 320 | ||
| 321 | local | |
| 322 | ||
| 323 | fun conv_mode x = | |
| 324 | ((Args.parens (Args.$$$ no_asmN) >> K simplify || | |
| 325 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || | |
| 326 | Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || | |
| 327 | Scan.succeed asm_full_simplify) |> Scan.lift) x; | |
| 328 | ||
| 329 | in | |
| 330 | ||
| 30528 | 331 | val simplified = conv_mode -- Attrib.thms >> | 
| 332 | (fn (f, ths) => Thm.rule_attribute (fn context => | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 333 | 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 | 334 | (simpset_of (Context.proof_of context)) addsimps ths))); | 
| 16014 | 335 | |
| 336 | end; | |
| 337 | ||
| 338 | ||
| 339 | (* setup attributes *) | |
| 340 | ||
| 26463 | 341 | val _ = Context.>> (Context.map_theory | 
| 30528 | 342 | (Attrib.setup (Binding.name simpN) (Attrib.add_del simp_add simp_del) | 
| 343 | "declaration of Simplifier rewrite rule" #> | |
| 344 | Attrib.setup (Binding.name congN) (Attrib.add_del cong_add cong_del) | |
| 345 | "declaration of Simplifier congruence rule" #> | |
| 33671 | 346 | Attrib.setup (Binding.name "simproc") simproc_att | 
| 347 | "declaration of simplification procedures" #> | |
| 30528 | 348 | Attrib.setup (Binding.name "simplified") simplified "simplified rule")); | 
| 16014 | 349 | |
| 350 | ||
| 351 | ||
| 31300 | 352 | (** method syntax **) | 
| 16014 | 353 | |
| 354 | val cong_modifiers = | |
| 18728 | 355 | [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier), | 
| 356 | Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add), | |
| 357 | Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)]; | |
| 16014 | 358 | |
| 359 | val simp_modifiers = | |
| 18728 | 360 | [Args.$$$ simpN -- Args.colon >> K (I, simp_add), | 
| 361 | Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add), | |
| 362 | 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 | 363 | Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 364 | >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] | 
| 16014 | 365 | @ cong_modifiers; | 
| 366 | ||
| 367 | val simp_modifiers' = | |
| 18728 | 368 | [Args.add -- Args.colon >> K (I, simp_add), | 
| 369 | Args.del -- Args.colon >> K (I, simp_del), | |
| 18688 | 370 | Args.$$$ onlyN -- Args.colon | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 371 | >> K (Context.proof_map (map_ss Raw_Simplifier.clear_ss), simp_add)] | 
| 16014 | 372 | @ cong_modifiers; | 
| 373 | ||
| 31300 | 374 | val simp_options = | 
| 375 | (Args.parens (Args.$$$ no_asmN) >> K simp_tac || | |
| 376 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || | |
| 377 | Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || | |
| 378 | Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || | |
| 379 | Scan.succeed asm_full_simp_tac); | |
| 16014 | 380 | |
| 31300 | 381 | fun simp_method more_mods meth = | 
| 35613 | 382 | Scan.lift simp_options --| | 
| 31300 | 383 | Method.sections (more_mods @ simp_modifiers') >> | 
| 35613 | 384 | (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts)); | 
| 16014 | 385 | |
| 386 | ||
| 387 | ||
| 18708 | 388 | (** setup **) | 
| 389 | ||
| 31300 | 390 | fun method_setup more_mods = | 
| 391 | Method.setup (Binding.name simpN) | |
| 392 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | |
| 393 | 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 | 394 | (CHANGED_PROP oo tac) (simpset_of ctxt)))) | 
| 31300 | 395 | "simplification" #> | 
| 396 | Method.setup (Binding.name "simp_all") | |
| 397 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | |
| 398 | ALLGOALS (Method.insert_tac facts) THEN | |
| 42372 | 399 | (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) (simpset_of ctxt))) | 
| 31300 | 400 | "simplification (all goals)"; | 
| 16014 | 401 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 402 | fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ => | 
| 16014 | 403 | let | 
| 404 | val trivialities = Drule.reflexive_thm :: trivs; | |
| 405 | ||
| 406 | fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac]; | |
| 407 | val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; | |
| 408 | ||
| 409 | (*no premature instantiation of variables during simplification*) | |
| 410 | fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac]; | |
| 411 | val safe_solver = mk_solver "easy safe" safe_solver_tac; | |
| 412 | ||
| 413 | fun mk_eq thm = | |
| 20872 | 414 | if can Logic.dest_equals (Thm.concl_of thm) then [thm] | 
| 16014 | 415 | else [thm RS reflect] handle THM _ => []; | 
| 416 | ||
| 26653 | 417 | fun mksimps thm = mk_eq (Thm.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm); | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 418 | in | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 419 | empty_ss setsubgoaler asm_simp_tac | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 420 | setSSolver safe_solver | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 421 | setSolver unsafe_solver | 
| 36543 
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
 wenzelm parents: 
35613diff
changeset | 422 | setmksimps (K mksimps) | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 423 | end)); | 
| 16014 | 424 | |
| 425 | end; | |
| 426 | ||
| 32738 | 427 | structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier; | 
| 428 | open Basic_Simplifier; |