| author | wenzelm | 
| Wed, 24 May 2006 22:44:56 +0200 | |
| changeset 19716 | 52c22fccdaaf | 
| parent 18728 | 6790126ab5f6 | 
| child 21646 | c07b5b0e8492 | 
| permissions | -rw-r--r-- | 
| 9772 | 1 | (* Title: Provers/clasimp.ML | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 2 | ID: $Id$ | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 3 | Author: David von Oheimb, TU Muenchen | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 4 | |
| 5219 | 5 | Combination of classical reasoner and simplifier (depends on | 
| 16019 | 6 | splitter.ML, classical.ML, blast.ML). | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 7 | *) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 8 | |
| 9860 | 9 | infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 addsimps2 delsimps2 | 
| 13603 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 berghofe parents: 
13026diff
changeset | 10 | addSss addss addss' addIffs delIffs; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 11 | |
| 5219 | 12 | signature CLASIMP_DATA = | 
| 13 | sig | |
| 8469 | 14 | structure Splitter: SPLITTER | 
| 5219 | 15 | structure Classical: CLASSICAL | 
| 16 | structure Blast: BLAST | |
| 17 | sharing type Classical.claset = Blast.claset | |
| 9860 | 18 | val notE: thm | 
| 19 | val iffD1: thm | |
| 20 | val iffD2: thm | |
| 5219 | 21 | end | 
| 22 | ||
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 23 | signature CLASIMP = | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 24 | sig | 
| 5219 | 25 | include CLASIMP_DATA | 
| 26 | type claset | |
| 27 | type clasimpset | |
| 17879 | 28 | val change_clasimpset: (clasimpset -> clasimpset) -> unit | 
| 10317 | 29 | val clasimpset: unit -> clasimpset | 
| 9860 | 30 | val addSIs2: clasimpset * thm list -> clasimpset | 
| 31 | val addSEs2: clasimpset * thm list -> clasimpset | |
| 32 | val addSDs2: clasimpset * thm list -> clasimpset | |
| 33 | val addIs2: clasimpset * thm list -> clasimpset | |
| 34 | val addEs2: clasimpset * thm list -> clasimpset | |
| 35 | val addDs2: clasimpset * thm list -> clasimpset | |
| 36 | val addsimps2: clasimpset * thm list -> clasimpset | |
| 37 | val delsimps2: clasimpset * thm list -> clasimpset | |
| 38 | val addSss: claset * simpset -> claset | |
| 39 | val addss: claset * simpset -> claset | |
| 13603 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 berghofe parents: 
13026diff
changeset | 40 | val addss': claset * simpset -> claset | 
| 9860 | 41 | val addIffs: clasimpset * thm list -> clasimpset | 
| 42 | val delIffs: clasimpset * thm list -> clasimpset | |
| 43 | val AddIffs: thm list -> unit | |
| 44 | val DelIffs: thm list -> unit | |
| 45 | val CLASIMPSET: (clasimpset -> tactic) -> tactic | |
| 46 | val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic | |
| 5483 | 47 | val clarsimp_tac: clasimpset -> int -> tactic | 
| 9860 | 48 | val Clarsimp_tac: int -> tactic | 
| 49 | val mk_auto_tac: clasimpset -> int -> int -> tactic | |
| 50 | val auto_tac: clasimpset -> tactic | |
| 51 | val Auto_tac: tactic | |
| 52 | val auto: unit -> unit | |
| 53 | val force_tac: clasimpset -> int -> tactic | |
| 54 | val Force_tac: int -> tactic | |
| 55 | val force: int -> unit | |
| 56 | val fast_simp_tac: clasimpset -> int -> tactic | |
| 57 | val slow_simp_tac: clasimpset -> int -> tactic | |
| 58 | val best_simp_tac: clasimpset -> int -> tactic | |
| 18728 | 59 | val attrib: (clasimpset * thm list -> clasimpset) -> attribute | 
| 9506 | 60 | val get_local_clasimpset: Proof.context -> clasimpset | 
| 15032 | 61 | val local_clasimpset_of: Proof.context -> clasimpset | 
| 18728 | 62 | val iff_add: attribute | 
| 63 | val iff_add': attribute | |
| 64 | val iff_del: attribute | |
| 9860 | 65 | val iff_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | 
| 9506 | 66 | val clasimp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | 
| 18708 | 67 | val setup: theory -> theory | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 68 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 69 | |
| 5219 | 70 | functor ClasimpFun(Data: CLASIMP_DATA): CLASIMP = | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 71 | struct | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 72 | |
| 5219 | 73 | open Data; | 
| 74 | ||
| 75 | type claset = Classical.claset; | |
| 76 | type clasimpset = claset * simpset; | |
| 77 | ||
| 17879 | 78 | fun change_clasimpset_of thy f = | 
| 79 | let val (cs', ss') = f (Classical.get_claset thy, Simplifier.get_simpset thy) in | |
| 80 | Classical.change_claset_of thy (fn _ => cs'); | |
| 81 | Simplifier.change_simpset_of thy (fn _ => ss') | |
| 82 | end; | |
| 83 | ||
| 84 | fun change_clasimpset f = change_clasimpset_of (Context.the_context ()) f; | |
| 10317 | 85 | fun clasimpset () = (Classical.claset (), Simplifier.simpset ()); | 
| 86 | ||
| 5219 | 87 | |
| 88 | (* clasimpset operations *) | |
| 89 | ||
| 90 | (*this interface for updating a clasimpset is rudimentary and just for | |
| 91 | convenience for the most common cases*) | |
| 92 | ||
| 93 | fun pair_upd1 f ((a,b),x) = (f(a,x), b); | |
| 94 | fun pair_upd2 f ((a,b),x) = (a, f(b,x)); | |
| 95 | ||
| 96 | fun op addSIs2 arg = pair_upd1 Classical.addSIs arg; | |
| 97 | fun op addSEs2 arg = pair_upd1 Classical.addSEs arg; | |
| 98 | fun op addSDs2 arg = pair_upd1 Classical.addSDs arg; | |
| 99 | fun op addIs2 arg = pair_upd1 Classical.addIs arg; | |
| 100 | fun op addEs2 arg = pair_upd1 Classical.addEs arg; | |
| 101 | fun op addDs2 arg = pair_upd1 Classical.addDs arg; | |
| 102 | fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg; | |
| 103 | fun op delsimps2 arg = pair_upd2 Simplifier.delsimps arg; | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 104 | |
| 9402 | 105 | (*not totally safe: may instantiate unknowns that appear also in other subgoals*) | 
| 106 | val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true); | |
| 107 | ||
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 108 | (*Add a simpset to a classical set!*) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 109 | (*Caution: only one simpset added can be added by each of addSss and addss*) | 
| 11181 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 oheimb parents: 
10821diff
changeset | 110 | fun cs addSss ss = Classical.addSafter (cs, ("safe_asm_full_simp_tac",
 | 
| 9772 | 111 | CHANGED o safe_asm_full_simp_tac ss)); | 
| 112 | fun cs addss  ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
 | |
| 113 | CHANGED o Simplifier.asm_full_simp_tac ss)); | |
| 13603 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 berghofe parents: 
13026diff
changeset | 114 | fun cs addss' ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
 | 
| 
57f364d1d3b2
Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
 berghofe parents: 
13026diff
changeset | 115 | CHANGED o Simplifier.asm_lr_simp_tac ss)); | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 116 | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 117 | (*Attach a suffix, provided we have a name to start with.*) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 118 | fun suffix_thm "" _ th = th | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 119 | | suffix_thm a b th = Thm.name_thm (a^b, th); | 
| 9860 | 120 | |
| 121 | (* iffs: addition of rules to simpsets and clasets simultaneously *) | |
| 122 | ||
| 11344 | 123 | (*Takes (possibly conditional) theorems of the form A<->B to | 
| 9860 | 124 | the Safe Intr rule B==>A and | 
| 125 | the Safe Destruct rule A==>B. | |
| 126 | Also ~A goes to the Safe Elim rule A ==> ?R | |
| 11462 | 127 | Failing other cases, A is added as a Safe Intr rule*) | 
| 9860 | 128 | local | 
| 129 | ||
| 13026 | 130 | fun addIff decls1 decls2 simp ((cs, ss), th) = | 
| 131 | let | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 132 | val name = Thm.name_of_thm th; | 
| 13026 | 133 | val n = nprems_of th; | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 134 | val (elim, intro) = if n = 0 then decls1 else decls2; | 
| 13026 | 135 | val zero_rotate = zero_var_indexes o rotate_prems n; | 
| 136 | in | |
| 18688 | 137 | (elim (intro (cs, [suffix_thm name "_iff2" (zero_rotate (th RS iffD2))]), | 
| 138 | [suffix_thm name "_iff1" (Tactic.make_elim (zero_rotate (th RS iffD1)))]) | |
| 139 | handle THM _ => (elim (cs, [suffix_thm name "_iff1" (zero_rotate (th RS notE))]) | |
| 11902 | 140 | handle THM _ => intro (cs, [th])), | 
| 141 | simp (ss, [th])) | |
| 142 | end; | |
| 9860 | 143 | |
| 12375 | 144 | fun delIff delcs delss ((cs, ss), th) = | 
| 11902 | 145 | let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in | 
| 18688 | 146 | (delcs (cs, [zero_rotate (th RS iffD2), | 
| 147 | Tactic.make_elim (zero_rotate (th RS iffD1))]) | |
| 148 | handle THM _ => (delcs (cs, [zero_rotate (th RS notE)]) | |
| 12375 | 149 | handle THM _ => delcs (cs, [th])), | 
| 150 | delss (ss, [th])) | |
| 11902 | 151 | end; | 
| 9860 | 152 | |
| 18630 | 153 | fun modifier att (x, ths) = | 
| 18728 | 154 | fst (foldl_map (Library.apply [att]) (x, rev ths)); | 
| 18630 | 155 | |
| 18688 | 156 | val addXIs = modifier (ContextRules.intro_query NONE); | 
| 157 | val addXEs = modifier (ContextRules.elim_query NONE); | |
| 158 | val addXDs = modifier (ContextRules.dest_query NONE); | |
| 159 | val delXs = modifier ContextRules.rule_del; | |
| 18630 | 160 | |
| 9860 | 161 | in | 
| 162 | ||
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 163 | val op addIffs = | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 164 | Library.foldl | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 165 | (addIff (Classical.addSEs, Classical.addSIs) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 166 | (Classical.addEs, Classical.addIs) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 167 | Simplifier.addsimps); | 
| 15570 | 168 | val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps); | 
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 169 | |
| 17879 | 170 | fun AddIffs thms = change_clasimpset (fn css => css addIffs thms); | 
| 171 | fun DelIffs thms = change_clasimpset (fn css => css delIffs thms); | |
| 9860 | 172 | |
| 18688 | 173 | fun addIffs_generic (x, ths) = | 
| 174 | Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1; | |
| 12375 | 175 | |
| 18688 | 176 | fun delIffs_generic (x, ths) = | 
| 177 | Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1; | |
| 12375 | 178 | |
| 9860 | 179 | end; | 
| 180 | ||
| 181 | ||
| 5219 | 182 | (* tacticals *) | 
| 183 | ||
| 184 | fun CLASIMPSET tacf state = | |
| 185 | Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state; | |
| 186 | ||
| 187 | fun CLASIMPSET' tacf i state = | |
| 188 | Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state; | |
| 4888 | 189 | |
| 190 | ||
| 12375 | 191 | fun clarsimp_tac (cs, ss) = | 
| 192 | safe_asm_full_simp_tac ss THEN_ALL_NEW | |
| 193 | Classical.clarify_tac (cs addSss ss); | |
| 194 | ||
| 10317 | 195 | fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i; | 
| 5483 | 196 | |
| 197 | ||
| 5219 | 198 | (* auto_tac *) | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 199 | |
| 5219 | 200 | fun blast_depth_tac cs m i thm = | 
| 9772 | 201 | Blast.depth_tac cs m i thm | 
| 5554 | 202 |       handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
 | 
| 9772 | 203 | |
| 204 | (* a variant of depth_tac that avoids interference of the simplifier | |
| 5219 | 205 | with dup_step_tac when they are combined by auto_tac *) | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 206 | local | 
| 9772 | 207 | fun slow_step_tac' cs = Classical.appWrappers cs | 
| 208 | (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs); | |
| 209 | in fun nodup_depth_tac cs m i state = SELECT_GOAL | |
| 210 | (Classical.safe_steps_tac cs 1 THEN_ELSE | |
| 211 | (DEPTH_SOLVE (nodup_depth_tac cs m 1), | |
| 212 | Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac | |
| 213 | (slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m-1) 1)) | |
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 214 | )) i state; | 
| 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 215 | end; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 216 | |
| 9402 | 217 | (*Designed to be idempotent, except if blast_depth_tac instantiates variables | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 218 | in some of the subgoals*) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 219 | fun mk_auto_tac (cs, ss) m n = | 
| 5219 | 220 | let val cs' = cs addss ss | 
| 9772 | 221 | val maintac = | 
| 222 | blast_depth_tac cs m (* fast but can't use wrappers *) | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 223 | ORELSE' | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 224 | (CHANGED o nodup_depth_tac cs' n); (* slower but more general *) | 
| 5219 | 225 | in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), | 
| 9772 | 226 | TRY (Classical.safe_tac cs), | 
| 227 | REPEAT (FIRSTGOAL maintac), | |
| 5219 | 228 | TRY (Classical.safe_tac (cs addSss ss)), | 
| 9772 | 229 | prune_params_tac] | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 230 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 231 | |
| 9772 | 232 | fun auto_tac css = mk_auto_tac css 4 2; | 
| 10317 | 233 | fun Auto_tac st = auto_tac (clasimpset ()) st; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 234 | fun auto () = by Auto_tac; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 235 | |
| 9772 | 236 | |
| 5219 | 237 | (* force_tac *) | 
| 238 | ||
| 4659 | 239 | (* aimed to solve the given subgoal totally, using whatever tools possible *) | 
| 4717 
1370ad043564
renamed smart_tac to force_tac, slight improvement of force_tac
 oheimb parents: 
4659diff
changeset | 240 | fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [ | 
| 9772 | 241 | Classical.clarify_tac cs' 1, | 
| 242 | IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1), | |
| 243 | ALLGOALS (Classical.first_best_tac cs')]) end; | |
| 10317 | 244 | fun Force_tac i = force_tac (clasimpset ()) i; | 
| 4717 
1370ad043564
renamed smart_tac to force_tac, slight improvement of force_tac
 oheimb parents: 
4659diff
changeset | 245 | fun force i = by (Force_tac i); | 
| 4659 | 246 | |
| 5219 | 247 | |
| 9805 | 248 | (* basic combinations *) | 
| 249 | ||
| 250 | fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end; | |
| 9591 | 251 | |
| 9805 | 252 | val fast_simp_tac = ADDSS Classical.fast_tac; | 
| 253 | val slow_simp_tac = ADDSS Classical.slow_tac; | |
| 254 | val best_simp_tac = ADDSS Classical.best_tac; | |
| 9591 | 255 | |
| 256 | ||
| 9860 | 257 | (* access clasimpset *) | 
| 8639 | 258 | |
| 7153 | 259 | fun get_local_clasimpset ctxt = | 
| 260 | (Classical.get_local_claset ctxt, Simplifier.get_local_simpset ctxt); | |
| 5926 | 261 | |
| 15032 | 262 | fun local_clasimpset_of ctxt = | 
| 263 | (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt); | |
| 264 | ||
| 9860 | 265 | |
| 266 | (* attributes *) | |
| 267 | ||
| 18728 | 268 | fun attrib f = Thm.declaration_attribute (fn th => | 
| 18688 | 269 | fn Context.Theory thy => (change_clasimpset_of thy (fn css => f (css, [th])); Context.Theory thy) | 
| 270 | | Context.Proof ctxt => | |
| 271 | let | |
| 272 | val cs = Classical.get_local_claset ctxt; | |
| 273 | val ss = Simplifier.get_local_simpset ctxt; | |
| 274 | val (cs', ss') = f ((cs, ss), [th]); | |
| 275 | val ctxt' = | |
| 276 | ctxt | |
| 277 | |> Classical.put_local_claset cs' | |
| 278 | |> Simplifier.put_local_simpset ss'; | |
| 279 | in Context.Proof ctxt' end); | |
| 9860 | 280 | |
| 18688 | 281 | fun attrib' f (x, th) = (f (x, [th]), th); | 
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 282 | |
| 18688 | 283 | val iff_add = attrib (op addIffs); | 
| 284 | val iff_add' = attrib' addIffs_generic; | |
| 285 | val iff_del = attrib (op delIffs) o attrib' delIffs_generic; | |
| 286 | ||
| 287 | val iff_att = Attrib.syntax (Scan.lift | |
| 288 | (Args.del >> K iff_del || | |
| 289 | Scan.option Args.add -- Args.query >> K iff_add' || | |
| 290 | Scan.option Args.add >> K iff_add)); | |
| 9860 | 291 | |
| 292 | ||
| 293 | (* method modifiers *) | |
| 294 | ||
| 295 | val iffN = "iff"; | |
| 296 | ||
| 297 | val iff_modifiers = | |
| 18728 | 298 | [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier), | 
| 299 | Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'), | |
| 300 | Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)]; | |
| 9860 | 301 | |
| 8469 | 302 | val clasimp_modifiers = | 
| 9860 | 303 | Simplifier.simp_modifiers @ Splitter.split_modifiers @ | 
| 304 | Classical.cla_modifiers @ iff_modifiers; | |
| 305 | ||
| 306 | ||
| 307 | (* methods *) | |
| 5926 | 308 | |
| 7559 | 309 | fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts => | 
| 15032 | 310 | ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt)); | 
| 7132 | 311 | |
| 7559 | 312 | fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts => | 
| 15032 | 313 | HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt))); | 
| 5926 | 314 | |
| 7559 | 315 | val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth; | 
| 316 | val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth'; | |
| 5926 | 317 | |
| 318 | ||
| 9772 | 319 | fun auto_args m = | 
| 320 | Method.bang_sectioned_args' clasimp_modifiers (Scan.lift (Scan.option (Args.nat -- Args.nat))) m; | |
| 321 | ||
| 15531 | 322 | fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac) | 
| 323 | | auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n)); | |
| 9772 | 324 | |
| 325 | ||
| 326 | (* theory setup *) | |
| 327 | ||
| 5926 | 328 | val setup = | 
| 18708 | 329 | (Attrib.add_attributes | 
| 18728 | 330 |    [("iff", iff_att, "declaration of Simplifier / Classical rules")] #>
 | 
| 9860 | 331 | Method.add_methods | 
| 9591 | 332 |    [("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"),
 | 
| 9805 | 333 |     ("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"),
 | 
| 334 |     ("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"),
 | |
| 9591 | 335 |     ("force", clasimp_method' force_tac, "force"),
 | 
| 9772 | 336 |     ("auto", auto_args auto_meth, "auto"),
 | 
| 18708 | 337 |     ("clarsimp", clasimp_method' (CHANGED_PROP oo clarsimp_tac), "clarify simplified goal")]);
 | 
| 5926 | 338 | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 339 | end; |