| author | wenzelm | 
| Wed, 16 Nov 2005 17:45:25 +0100 | |
| changeset 18179 | cf4b265007bf | 
| parent 18033 | ab8803126f84 | 
| child 18629 | bdf01da93ed4 | 
| permissions | -rw-r--r-- | 
| 16014 | 1 | (* Title: Pure/simplifier.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Generic simplifier, suitable for most logics (see also | |
| 6 | meta_simplifier.ML for the actual meta-level rewriting engine). | |
| 7 | *) | |
| 8 | ||
| 9 | signature BASIC_SIMPLIFIER = | |
| 10 | sig | |
| 11 | include BASIC_META_SIMPLIFIER | |
| 12 | val print_simpset: theory -> unit | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 13 | val change_simpset_of: theory -> (simpset -> simpset) -> unit | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 14 | val change_simpset: (simpset -> simpset) -> unit | 
| 16014 | 15 | val simpset_of: theory -> simpset | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 16 | val simpset: unit -> simpset | 
| 16014 | 17 | val SIMPSET: (simpset -> tactic) -> tactic | 
| 18 | val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic | |
| 19 | val Addsimps: thm list -> unit | |
| 20 | val Delsimps: thm list -> unit | |
| 21 | val Addsimprocs: simproc list -> unit | |
| 22 | val Delsimprocs: simproc list -> unit | |
| 23 | val Addcongs: thm list -> unit | |
| 24 | val Delcongs: thm list -> unit | |
| 25 | val local_simpset_of: Proof.context -> simpset | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 26 | val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic | 
| 16014 | 27 | val safe_asm_full_simp_tac: simpset -> int -> tactic | 
| 28 | val simp_tac: simpset -> int -> tactic | |
| 29 | val asm_simp_tac: simpset -> int -> tactic | |
| 30 | val full_simp_tac: simpset -> int -> tactic | |
| 31 | val asm_lr_simp_tac: simpset -> int -> tactic | |
| 32 | val asm_full_simp_tac: simpset -> int -> tactic | |
| 33 | val Simp_tac: int -> tactic | |
| 34 | val Asm_simp_tac: int -> tactic | |
| 35 | val Full_simp_tac: int -> tactic | |
| 36 | val Asm_lr_simp_tac: int -> tactic | |
| 37 | val Asm_full_simp_tac: int -> tactic | |
| 38 | val simplify: simpset -> thm -> thm | |
| 39 | val asm_simplify: simpset -> thm -> thm | |
| 40 | val full_simplify: simpset -> thm -> thm | |
| 41 | val asm_lr_simplify: simpset -> thm -> thm | |
| 42 | val asm_full_simplify: simpset -> thm -> thm | |
| 43 | end; | |
| 44 | ||
| 45 | signature SIMPLIFIER = | |
| 46 | sig | |
| 47 | include BASIC_SIMPLIFIER | |
| 17004 | 48 | val clear_ss: simpset -> simpset | 
| 17723 | 49 | val debug_bounds: bool ref | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 50 | val inherit_context: simpset -> simpset -> simpset | 
| 17898 | 51 | val the_context: simpset -> Proof.context | 
| 52 | val context: Proof.context -> simpset -> simpset | |
| 53 | val theory_context: theory -> simpset -> simpset | |
| 16458 | 54 | val simproc_i: theory -> string -> term list | 
| 55 | -> (theory -> simpset -> term -> thm option) -> simproc | |
| 56 | val simproc: theory -> string -> string list | |
| 57 | -> (theory -> simpset -> term -> thm option) -> simproc | |
| 16014 | 58 | val rewrite: simpset -> cterm -> thm | 
| 59 | val asm_rewrite: simpset -> cterm -> thm | |
| 60 | val full_rewrite: simpset -> cterm -> thm | |
| 61 | val asm_lr_rewrite: simpset -> cterm -> thm | |
| 62 | val asm_full_rewrite: simpset -> cterm -> thm | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 63 | val get_simpset: theory -> simpset | 
| 16014 | 64 | val print_local_simpset: Proof.context -> unit | 
| 65 | val get_local_simpset: Proof.context -> simpset | |
| 66 | val put_local_simpset: simpset -> Proof.context -> Proof.context | |
| 67 | val change_global_ss: (simpset * thm list -> simpset) -> theory attribute | |
| 68 | val change_local_ss: (simpset * thm list -> simpset) -> Proof.context attribute | |
| 69 | val simp_add_global: theory attribute | |
| 70 | val simp_del_global: theory attribute | |
| 71 | val simp_add_local: Proof.context attribute | |
| 72 | val simp_del_local: Proof.context attribute | |
| 73 | val cong_add_global: theory attribute | |
| 74 | val cong_del_global: theory attribute | |
| 75 | val cong_add_local: Proof.context attribute | |
| 76 | val cong_del_local: Proof.context attribute | |
| 18033 | 77 | val cong_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | 
| 78 | val simp_modifiers': (Args.T list -> (Method.modifier * Args.T list)) list | |
| 16014 | 79 | val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | 
| 80 | val method_setup: (Args.T list -> (Method.modifier * Args.T list)) list | |
| 81 | -> (theory -> theory) list | |
| 82 | val easy_setup: thm -> thm list -> (theory -> theory) list | |
| 83 | end; | |
| 84 | ||
| 85 | structure Simplifier: SIMPLIFIER = | |
| 86 | struct | |
| 87 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 88 | (** simpset data **) | 
| 16014 | 89 | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 90 | (* global simpsets *) | 
| 16014 | 91 | |
| 16458 | 92 | structure GlobalSimpset = TheoryDataFun | 
| 93 | (struct | |
| 16014 | 94 | val name = "Pure/simpset"; | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 95 | type T = simpset ref; | 
| 16014 | 96 | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 97 | val empty = ref empty_ss; | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 98 | fun copy (ref ss) = ref ss: T; (*create new reference!*) | 
| 17898 | 99 | fun extend (ref ss) = ref (MetaSimplifier.inherit_context empty_ss ss); | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 100 | fun merge _ (ref ss1, ref ss2) = ref (merge_ss (ss1, ss2)); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 101 | fun print _ (ref ss) = print_ss ss; | 
| 16458 | 102 | end); | 
| 16014 | 103 | |
| 104 | val _ = Context.add_setup [GlobalSimpset.init]; | |
| 105 | val print_simpset = GlobalSimpset.print; | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 106 | val get_simpset = ! o GlobalSimpset.get; | 
| 16014 | 107 | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 108 | val change_simpset_of = change o GlobalSimpset.get; | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 109 | fun change_simpset f = change_simpset_of (Context.the_context ()) f; | 
| 16014 | 110 | |
| 17898 | 111 | fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_simpset thy); | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 112 | val simpset = simpset_of o Context.the_context; | 
| 16014 | 113 | |
| 114 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 115 | fun SIMPSET tacf st = tacf (simpset_of (Thm.theory_of_thm st)) st; | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 116 | fun SIMPSET' tacf i st = tacf (simpset_of (Thm.theory_of_thm st)) i st; | 
| 16014 | 117 | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 118 | fun Addsimps args = change_simpset (fn ss => ss addsimps args); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 119 | fun Delsimps args = change_simpset (fn ss => ss delsimps args); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 120 | 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 | 121 | fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 122 | fun Addcongs args = change_simpset (fn ss => ss addcongs args); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 123 | fun Delcongs args = change_simpset (fn ss => ss delcongs args); | 
| 16014 | 124 | |
| 125 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 126 | (* local simpsets *) | 
| 16014 | 127 | |
| 16458 | 128 | structure LocalSimpset = ProofDataFun | 
| 129 | (struct | |
| 16014 | 130 | val name = "Pure/simpset"; | 
| 131 | type T = simpset; | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 132 | val init = get_simpset; | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 133 | fun print _ ss = print_ss ss; | 
| 16458 | 134 | end); | 
| 16014 | 135 | |
| 136 | val _ = Context.add_setup [LocalSimpset.init]; | |
| 137 | val print_local_simpset = LocalSimpset.print; | |
| 138 | val get_local_simpset = LocalSimpset.get; | |
| 139 | val put_local_simpset = LocalSimpset.put; | |
| 140 | ||
| 17898 | 141 | fun local_simpset_of ctxt = MetaSimplifier.context ctxt (get_local_simpset ctxt); | 
| 16014 | 142 | |
| 143 | ||
| 144 | (* attributes *) | |
| 145 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 146 | fun change_global_ss f (thy, th) = (change_simpset_of thy (fn ss => f (ss, [th])); (thy, th)); | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 147 | fun change_local_ss f (ctxt, th) = (LocalSimpset.map (fn ss => f (ss, [th])) ctxt, th); | 
| 16014 | 148 | |
| 149 | val simp_add_global = change_global_ss (op addsimps); | |
| 150 | val simp_del_global = change_global_ss (op delsimps); | |
| 16806 | 151 | val simp_add_local = change_local_ss (op addsimps); | 
| 16014 | 152 | val simp_del_local = change_local_ss (op delsimps); | 
| 153 | ||
| 154 | val cong_add_global = change_global_ss (op addcongs); | |
| 155 | val cong_del_global = change_global_ss (op delcongs); | |
| 156 | val cong_add_local = change_local_ss (op addcongs); | |
| 157 | val cong_del_local = change_local_ss (op delcongs); | |
| 158 | ||
| 159 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 160 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 161 | (** simplification tactics and rules **) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 162 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 163 | fun solve_all_tac solvers ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 164 | let | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 165 |     val (_, {subgoal_tac, ...}) = MetaSimplifier.rep_ss ss;
 | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 166 | val solve_tac = subgoal_tac (MetaSimplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 167 | in DEPTH_SOLVE (solve_tac 1) end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 168 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 169 | (*NOTE: may instantiate unknowns that appear also in other subgoals*) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 170 | fun generic_simp_tac safe mode ss = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 171 | let | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 172 |     val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = MetaSimplifier.rep_ss ss;
 | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 173 | val loop_tac = FIRST' (map (fn (_, tac) => tac ss) loop_tacs); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 174 | val solve_tac = FIRST' (map (MetaSimplifier.solver ss) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 175 | (if safe then solvers else unsafe_solvers)); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 176 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 177 | fun simp_loop_tac i = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 178 | asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 179 | (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 | 180 | in simp_loop_tac end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 181 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 182 | local | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 183 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 184 | fun simp rew mode ss thm = | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 185 | let | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 186 |     val (_, {solvers = (unsafe_solvers, _), ...}) = MetaSimplifier.rep_ss ss;
 | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 187 | val tacf = solve_all_tac unsafe_solvers; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 188 | 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 | 189 | in rew mode prover ss thm end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 190 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 191 | in | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 192 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 193 | val simp_thm = simp MetaSimplifier.rewrite_thm; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 194 | val simp_cterm = simp MetaSimplifier.rewrite_cterm; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 195 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 196 | end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 197 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 198 | |
| 16806 | 199 | (* tactics *) | 
| 200 | ||
| 16014 | 201 | val simp_tac = generic_simp_tac false (false, false, false); | 
| 202 | val asm_simp_tac = generic_simp_tac false (false, true, false); | |
| 203 | val full_simp_tac = generic_simp_tac false (true, false, false); | |
| 204 | val asm_lr_simp_tac = generic_simp_tac false (true, true, false); | |
| 205 | val asm_full_simp_tac = generic_simp_tac false (true, true, true); | |
| 206 | val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); | |
| 207 | ||
| 208 | (*the abstraction over the proof state delays the dereferencing*) | |
| 209 | fun Simp_tac i st = simp_tac (simpset ()) i st; | |
| 210 | fun Asm_simp_tac i st = asm_simp_tac (simpset ()) i st; | |
| 211 | fun Full_simp_tac i st = full_simp_tac (simpset ()) i st; | |
| 212 | fun Asm_lr_simp_tac i st = asm_lr_simp_tac (simpset ()) i st; | |
| 213 | fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st; | |
| 214 | ||
| 16806 | 215 | |
| 216 | (* conversions *) | |
| 217 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 218 | val simplify = simp_thm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 219 | val asm_simplify = simp_thm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 220 | val full_simplify = simp_thm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 221 | val asm_lr_simplify = simp_thm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 222 | val asm_full_simplify = simp_thm (true, true, true); | 
| 16014 | 223 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 224 | val rewrite = simp_cterm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 225 | val asm_rewrite = simp_cterm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 226 | val full_rewrite = simp_cterm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 227 | val asm_lr_rewrite = simp_cterm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 228 | val asm_full_rewrite = simp_cterm (true, true, true); | 
| 16014 | 229 | |
| 230 | ||
| 231 | ||
| 232 | (** concrete syntax of attributes **) | |
| 233 | ||
| 234 | (* add / del *) | |
| 235 | ||
| 236 | val simpN = "simp"; | |
| 237 | val congN = "cong"; | |
| 238 | val addN = "add"; | |
| 239 | val delN = "del"; | |
| 240 | val onlyN = "only"; | |
| 241 | val no_asmN = "no_asm"; | |
| 242 | val no_asm_useN = "no_asm_use"; | |
| 243 | val no_asm_simpN = "no_asm_simp"; | |
| 244 | val asm_lrN = "asm_lr"; | |
| 245 | ||
| 246 | val simp_attr = | |
| 247 | (Attrib.add_del_args simp_add_global simp_del_global, | |
| 248 | Attrib.add_del_args simp_add_local simp_del_local); | |
| 249 | ||
| 250 | val cong_attr = | |
| 251 | (Attrib.add_del_args cong_add_global cong_del_global, | |
| 252 | Attrib.add_del_args cong_add_local cong_del_local); | |
| 253 | ||
| 254 | ||
| 255 | (* conversions *) | |
| 256 | ||
| 257 | local | |
| 258 | ||
| 259 | fun conv_mode x = | |
| 260 | ((Args.parens (Args.$$$ no_asmN) >> K simplify || | |
| 261 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || | |
| 262 | Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || | |
| 263 | Scan.succeed asm_full_simplify) |> Scan.lift) x; | |
| 264 | ||
| 265 | fun simplified_att get args = | |
| 266 | Attrib.syntax (conv_mode -- args >> (fn (f, ths) => | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 267 | Drule.rule_attribute (fn x => | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 268 | f ((if null ths then I else MetaSimplifier.clear_ss) (get x) addsimps ths)))); | 
| 16014 | 269 | |
| 270 | in | |
| 271 | ||
| 272 | val simplified_attr = | |
| 273 | (simplified_att simpset_of Attrib.global_thmss, | |
| 274 | simplified_att local_simpset_of Attrib.local_thmss); | |
| 275 | ||
| 276 | end; | |
| 277 | ||
| 278 | ||
| 279 | (* setup attributes *) | |
| 280 | ||
| 281 | val _ = Context.add_setup | |
| 282 | [Attrib.add_attributes | |
| 283 | [(simpN, simp_attr, "declaration of simplification rule"), | |
| 284 | (congN, cong_attr, "declaration of Simplifier congruence rule"), | |
| 285 |     ("simplified", simplified_attr, "simplified rule")]];
 | |
| 286 | ||
| 287 | ||
| 288 | ||
| 289 | (** proof methods **) | |
| 290 | ||
| 291 | (* simplification *) | |
| 292 | ||
| 293 | val simp_options = | |
| 294 | (Args.parens (Args.$$$ no_asmN) >> K simp_tac || | |
| 295 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || | |
| 296 | Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || | |
| 297 | Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || | |
| 298 | Scan.succeed asm_full_simp_tac); | |
| 299 | ||
| 16685 | 300 | fun simp_flags x = (Scan.repeat | 
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 301 | (Args.parens (Args.$$$ "depth_limit" -- Args.colon |-- Args.nat) | 
| 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 302 | >> setmp MetaSimplifier.simp_depth_limit) | 
| 16709 | 303 | >> (curry (Library.foldl op o) I o rev)) x; | 
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 304 | |
| 16014 | 305 | val cong_modifiers = | 
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 306 | [Args.$$$ congN -- Args.colon >> K ((I, cong_add_local): Method.modifier), | 
| 16014 | 307 | Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add_local), | 
| 308 | Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del_local)]; | |
| 309 | ||
| 310 | val simp_modifiers = | |
| 311 | [Args.$$$ simpN -- Args.colon >> K (I, simp_add_local), | |
| 312 | Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add_local), | |
| 313 | Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del_local), | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 314 | Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 315 | >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add_local)] | 
| 16014 | 316 | @ cong_modifiers; | 
| 317 | ||
| 318 | val simp_modifiers' = | |
| 319 | [Args.add -- Args.colon >> K (I, simp_add_local), | |
| 320 | Args.del -- Args.colon >> K (I, simp_del_local), | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 321 | Args.$$$ onlyN -- Args.colon >> K (LocalSimpset.map MetaSimplifier.clear_ss, simp_add_local)] | 
| 16014 | 322 | @ cong_modifiers; | 
| 323 | ||
| 324 | fun simp_args more_mods = | |
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 325 | Method.sectioned_args (Args.bang_facts -- Scan.lift simp_options -- Scan.lift simp_flags) | 
| 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 326 | (more_mods @ simp_modifiers'); | 
| 16014 | 327 | |
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 328 | fun simp_method ((prems, tac), FLAGS) ctxt = Method.METHOD (fn facts => | 
| 16014 | 329 | ALLGOALS (Method.insert_tac (prems @ facts)) THEN | 
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 330 | (FLAGS o CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt)); | 
| 16014 | 331 | |
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 332 | fun simp_method' ((prems, tac), FLAGS) ctxt = Method.METHOD (fn facts => | 
| 16014 | 333 | HEADGOAL (Method.insert_tac (prems @ facts) THEN' | 
| 16684 
7b58002668c0
methods: added simp_flags argument, added "depth_limit" flag;
 wenzelm parents: 
16458diff
changeset | 334 | ((FLAGS o CHANGED_PROP) oo tac) (local_simpset_of ctxt))); | 
| 16014 | 335 | |
| 336 | ||
| 337 | (* setup methods *) | |
| 338 | ||
| 339 | fun setup_methods more_mods = Method.add_methods | |
| 340 | [(simpN, simp_args more_mods simp_method', "simplification"), | |
| 341 |   ("simp_all", simp_args more_mods simp_method, "simplification (all goals)")];
 | |
| 342 | ||
| 343 | fun method_setup mods = [setup_methods mods]; | |
| 344 | ||
| 345 | ||
| 346 | (** easy_setup **) | |
| 347 | ||
| 348 | fun easy_setup reflect trivs = | |
| 349 | let | |
| 350 | val trivialities = Drule.reflexive_thm :: trivs; | |
| 351 | ||
| 352 | fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac]; | |
| 353 | val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; | |
| 354 | ||
| 355 | (*no premature instantiation of variables during simplification*) | |
| 356 | fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac]; | |
| 357 | val safe_solver = mk_solver "easy safe" safe_solver_tac; | |
| 358 | ||
| 359 | fun mk_eq thm = | |
| 360 | if Logic.is_equals (Thm.concl_of thm) then [thm] | |
| 361 | else [thm RS reflect] handle THM _ => []; | |
| 362 | ||
| 363 | fun mksimps thm = mk_eq (Drule.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm); | |
| 364 | ||
| 365 | fun init_ss thy = | |
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 366 | (GlobalSimpset.get thy := | 
| 16014 | 367 | empty_ss setsubgoaler asm_simp_tac | 
| 368 | setSSolver safe_solver | |
| 369 | setSolver unsafe_solver | |
| 370 | setmksimps mksimps; thy); | |
| 371 | in method_setup [] @ [init_ss] end; | |
| 372 | ||
| 17883 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 373 | |
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 374 | open MetaSimplifier; | 
| 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 wenzelm parents: 
17723diff
changeset | 375 | |
| 16014 | 376 | end; | 
| 377 | ||
| 378 | structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier; | |
| 379 | open BasicSimplifier; |