| author | wenzelm | 
| Sat, 13 Nov 2010 22:33:07 +0100 | |
| changeset 40533 | e38e80686ce5 | 
| parent 37720 | 50a9e2fa4f6b | 
| child 42373 | fb539d6d1ac2 | 
| 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 | Author: David von Oheimb, TU Muenchen | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 3 | |
| 5219 | 4 | Combination of classical reasoner and simplifier (depends on | 
| 16019 | 5 | 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 | 6 | *) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 7 | |
| 9860 | 8 | 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 | 9 | addSss addss addss' addIffs delIffs; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 10 | |
| 5219 | 11 | signature CLASIMP_DATA = | 
| 12 | sig | |
| 8469 | 13 | structure Splitter: SPLITTER | 
| 5219 | 14 | structure Classical: CLASSICAL | 
| 15 | structure Blast: BLAST | |
| 16 | sharing type Classical.claset = Blast.claset | |
| 9860 | 17 | val notE: thm | 
| 18 | val iffD1: thm | |
| 19 | val iffD2: thm | |
| 5219 | 20 | end | 
| 21 | ||
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 22 | signature CLASIMP = | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 23 | sig | 
| 5219 | 24 | include CLASIMP_DATA | 
| 25 | type claset | |
| 26 | type clasimpset | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 27 | val get_css: Context.generic -> clasimpset | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 28 | val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic | 
| 32148 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
30609diff
changeset | 29 | val clasimpset_of: Proof.context -> 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 | |
| 5483 | 43 | val clarsimp_tac: clasimpset -> int -> tactic | 
| 9860 | 44 | val mk_auto_tac: clasimpset -> int -> int -> tactic | 
| 45 | val auto_tac: clasimpset -> tactic | |
| 46 | val force_tac: clasimpset -> int -> tactic | |
| 47 | val fast_simp_tac: clasimpset -> int -> tactic | |
| 48 | val slow_simp_tac: clasimpset -> int -> tactic | |
| 49 | val best_simp_tac: clasimpset -> int -> tactic | |
| 18728 | 50 | val attrib: (clasimpset * thm list -> clasimpset) -> attribute | 
| 51 | val iff_add: attribute | |
| 52 | val iff_add': attribute | |
| 53 | val iff_del: attribute | |
| 30513 | 54 | val iff_modifiers: Method.modifier parser list | 
| 55 | val clasimp_modifiers: Method.modifier parser list | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 56 | val clasimp_setup: theory -> theory | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 57 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 58 | |
| 5219 | 59 | 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 | 60 | struct | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 61 | |
| 5219 | 62 | open Data; | 
| 63 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 64 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 65 | (* type clasimpset *) | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 66 | |
| 5219 | 67 | type claset = Classical.claset; | 
| 68 | type clasimpset = claset * simpset; | |
| 69 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 70 | fun get_css context = (Classical.get_cs context, Simplifier.get_ss context); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 71 | |
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 72 | fun map_css f context = | 
| 36601 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 73 | let | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 74 | val (cs, ss) = get_css context; | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 75 | val (cs', ss') = f (cs, Simplifier.context (Context.proof_of context) ss); | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 76 | in | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 77 | context | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 78 | |> Classical.map_cs (K cs') | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 79 | |> Simplifier.map_ss (K (Simplifier.inherit_context ss ss')) | 
| 
8a041e2d8122
map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
 wenzelm parents: 
35613diff
changeset | 80 | end; | 
| 17879 | 81 | |
| 32148 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 wenzelm parents: 
30609diff
changeset | 82 | fun clasimpset_of ctxt = (Classical.claset_of ctxt, Simplifier.simpset_of ctxt); | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 83 | |
| 5219 | 84 | |
| 85 | (* clasimpset operations *) | |
| 86 | ||
| 87 | (*this interface for updating a clasimpset is rudimentary and just for | |
| 88 | convenience for the most common cases*) | |
| 89 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 90 | fun pair_upd1 f ((a, b), x) = (f (a, x), b); | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 91 | fun pair_upd2 f ((a, b), x) = (a, f (b, x)); | 
| 5219 | 92 | |
| 93 | fun op addSIs2 arg = pair_upd1 Classical.addSIs arg; | |
| 94 | fun op addSEs2 arg = pair_upd1 Classical.addSEs arg; | |
| 95 | fun op addSDs2 arg = pair_upd1 Classical.addSDs arg; | |
| 96 | fun op addIs2 arg = pair_upd1 Classical.addIs arg; | |
| 97 | fun op addEs2 arg = pair_upd1 Classical.addEs arg; | |
| 98 | fun op addDs2 arg = pair_upd1 Classical.addDs arg; | |
| 99 | fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg; | |
| 100 | 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 | 101 | |
| 9402 | 102 | (*not totally safe: may instantiate unknowns that appear also in other subgoals*) | 
| 103 | val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true); | |
| 104 | ||
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 105 | (*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 | 106 | (*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 | 107 | fun cs addSss ss = Classical.addSafter (cs, ("safe_asm_full_simp_tac",
 | 
| 9772 | 108 | CHANGED o safe_asm_full_simp_tac ss)); | 
| 109 | fun cs addss  ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
 | |
| 110 | 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 | 111 | 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 | 112 | 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 | 113 | |
| 9860 | 114 | (* iffs: addition of rules to simpsets and clasets simultaneously *) | 
| 115 | ||
| 11344 | 116 | (*Takes (possibly conditional) theorems of the form A<->B to | 
| 9860 | 117 | the Safe Intr rule B==>A and | 
| 118 | the Safe Destruct rule A==>B. | |
| 119 | Also ~A goes to the Safe Elim rule A ==> ?R | |
| 11462 | 120 | Failing other cases, A is added as a Safe Intr rule*) | 
| 9860 | 121 | local | 
| 122 | ||
| 13026 | 123 | fun addIff decls1 decls2 simp ((cs, ss), th) = | 
| 124 | let | |
| 125 | val n = nprems_of th; | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 126 | val (elim, intro) = if n = 0 then decls1 else decls2; | 
| 13026 | 127 | val zero_rotate = zero_var_indexes o rotate_prems n; | 
| 128 | in | |
| 21709 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 paulson parents: 
21646diff
changeset | 129 | (elim (intro (cs, [zero_rotate (th RS iffD2)]), | 
| 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 paulson parents: 
21646diff
changeset | 130 | [Tactic.make_elim (zero_rotate (th RS iffD1))]) | 
| 
9cfd9eb9faec
removed use of put_name_hint, as the ATP linkup no longer needs this
 paulson parents: 
21646diff
changeset | 131 | handle THM _ => (elim (cs, [zero_rotate (th RS notE)]) | 
| 11902 | 132 | handle THM _ => intro (cs, [th])), | 
| 133 | simp (ss, [th])) | |
| 134 | end; | |
| 9860 | 135 | |
| 12375 | 136 | fun delIff delcs delss ((cs, ss), th) = | 
| 11902 | 137 | let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in | 
| 18688 | 138 | (delcs (cs, [zero_rotate (th RS iffD2), | 
| 139 | Tactic.make_elim (zero_rotate (th RS iffD1))]) | |
| 140 | handle THM _ => (delcs (cs, [zero_rotate (th RS notE)]) | |
| 12375 | 141 | handle THM _ => delcs (cs, [th])), | 
| 142 | delss (ss, [th])) | |
| 11902 | 143 | end; | 
| 9860 | 144 | |
| 18630 | 145 | fun modifier att (x, ths) = | 
| 30190 | 146 | fst (Library.foldl_map (Library.apply [att]) (x, rev ths)); | 
| 18630 | 147 | |
| 33369 | 148 | val addXIs = modifier (Context_Rules.intro_query NONE); | 
| 149 | val addXEs = modifier (Context_Rules.elim_query NONE); | |
| 150 | val delXs = modifier Context_Rules.rule_del; | |
| 18630 | 151 | |
| 9860 | 152 | in | 
| 153 | ||
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 154 | val op addIffs = | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 155 | Library.foldl | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 156 | (addIff (Classical.addSEs, Classical.addSIs) | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 157 | (Classical.addEs, Classical.addIs) | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 158 | Simplifier.addsimps); | 
| 15570 | 159 | 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 | 160 | |
| 18688 | 161 | fun addIffs_generic (x, ths) = | 
| 162 | Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1; | |
| 12375 | 163 | |
| 18688 | 164 | fun delIffs_generic (x, ths) = | 
| 165 | Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1; | |
| 12375 | 166 | |
| 9860 | 167 | end; | 
| 168 | ||
| 169 | ||
| 5219 | 170 | (* tacticals *) | 
| 171 | ||
| 12375 | 172 | fun clarsimp_tac (cs, ss) = | 
| 173 | safe_asm_full_simp_tac ss THEN_ALL_NEW | |
| 174 | Classical.clarify_tac (cs addSss ss); | |
| 175 | ||
| 5483 | 176 | |
| 5219 | 177 | (* auto_tac *) | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 178 | |
| 5219 | 179 | fun blast_depth_tac cs m i thm = | 
| 9772 | 180 | Blast.depth_tac cs m i thm | 
| 5554 | 181 |       handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
 | 
| 9772 | 182 | |
| 183 | (* a variant of depth_tac that avoids interference of the simplifier | |
| 5219 | 184 | with dup_step_tac when they are combined by auto_tac *) | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 185 | local | 
| 9772 | 186 | fun slow_step_tac' cs = Classical.appWrappers cs | 
| 187 | (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs); | |
| 188 | in fun nodup_depth_tac cs m i state = SELECT_GOAL | |
| 189 | (Classical.safe_steps_tac cs 1 THEN_ELSE | |
| 190 | (DEPTH_SOLVE (nodup_depth_tac cs m 1), | |
| 191 | Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac | |
| 192 | (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 | 193 | )) i state; | 
| 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 194 | end; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 195 | |
| 9402 | 196 | (*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 | 197 | in some of the subgoals*) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 198 | fun mk_auto_tac (cs, ss) m n = | 
| 5219 | 199 | let val cs' = cs addss ss | 
| 9772 | 200 | val maintac = | 
| 201 | 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 | 202 | ORELSE' | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 203 | (CHANGED o nodup_depth_tac cs' n); (* slower but more general *) | 
| 5219 | 204 | in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), | 
| 9772 | 205 | TRY (Classical.safe_tac cs), | 
| 37720 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 paulson parents: 
36960diff
changeset | 206 | REPEAT_DETERM (FIRSTGOAL maintac), | 
| 5219 | 207 | TRY (Classical.safe_tac (cs addSss ss)), | 
| 9772 | 208 | prune_params_tac] | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 209 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 210 | |
| 9772 | 211 | fun auto_tac css = mk_auto_tac css 4 2; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 212 | |
| 9772 | 213 | |
| 5219 | 214 | (* force_tac *) | 
| 215 | ||
| 4659 | 216 | (* 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 | 217 | fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [ | 
| 9772 | 218 | Classical.clarify_tac cs' 1, | 
| 219 | IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1), | |
| 220 | ALLGOALS (Classical.first_best_tac cs')]) end; | |
| 4659 | 221 | |
| 5219 | 222 | |
| 9805 | 223 | (* basic combinations *) | 
| 224 | ||
| 225 | fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end; | |
| 9591 | 226 | |
| 9805 | 227 | val fast_simp_tac = ADDSS Classical.fast_tac; | 
| 228 | val slow_simp_tac = ADDSS Classical.slow_tac; | |
| 229 | val best_simp_tac = ADDSS Classical.best_tac; | |
| 9591 | 230 | |
| 231 | ||
| 9860 | 232 | (* access clasimpset *) | 
| 8639 | 233 | |
| 15032 | 234 | |
| 9860 | 235 | |
| 236 | (* attributes *) | |
| 237 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 238 | fun attrib f = Thm.declaration_attribute (fn th => map_css (fn css => f (css, [th]))); | 
| 18688 | 239 | fun attrib' f (x, th) = (f (x, [th]), th); | 
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 240 | |
| 18688 | 241 | val iff_add = attrib (op addIffs); | 
| 242 | val iff_add' = attrib' addIffs_generic; | |
| 243 | val iff_del = attrib (op delIffs) o attrib' delIffs_generic; | |
| 244 | ||
| 30528 | 245 | fun iff_att x = (Scan.lift | 
| 18688 | 246 | (Args.del >> K iff_del || | 
| 247 | Scan.option Args.add -- Args.query >> K iff_add' || | |
| 30528 | 248 | Scan.option Args.add >> K iff_add)) x; | 
| 9860 | 249 | |
| 250 | ||
| 251 | (* method modifiers *) | |
| 252 | ||
| 253 | val iffN = "iff"; | |
| 254 | ||
| 255 | val iff_modifiers = | |
| 18728 | 256 | [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier), | 
| 257 | Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'), | |
| 258 | Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)]; | |
| 9860 | 259 | |
| 8469 | 260 | val clasimp_modifiers = | 
| 9860 | 261 | Simplifier.simp_modifiers @ Splitter.split_modifiers @ | 
| 262 | Classical.cla_modifiers @ iff_modifiers; | |
| 263 | ||
| 264 | ||
| 265 | (* methods *) | |
| 5926 | 266 | |
| 35613 | 267 | fun clasimp_meth tac ctxt = METHOD (fn facts => | 
| 268 | ALLGOALS (Method.insert_tac facts) THEN tac (clasimpset_of ctxt)); | |
| 7132 | 269 | |
| 35613 | 270 | fun clasimp_meth' tac ctxt = METHOD (fn facts => | 
| 271 | HEADGOAL (Method.insert_tac facts THEN' tac (clasimpset_of ctxt))); | |
| 5926 | 272 | |
| 273 | ||
| 30541 | 274 | fun clasimp_method' tac = | 
| 35613 | 275 | Method.sections clasimp_modifiers >> K (clasimp_meth' tac); | 
| 9772 | 276 | |
| 30541 | 277 | val auto_method = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36601diff
changeset | 278 | Scan.lift (Scan.option (Parse.nat -- Parse.nat)) --| | 
| 35613 | 279 | Method.sections clasimp_modifiers >> | 
| 280 | (fn NONE => clasimp_meth (CHANGED_PROP o auto_tac) | |
| 281 | | SOME (m, n) => clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n))); | |
| 9772 | 282 | |
| 283 | ||
| 284 | (* theory setup *) | |
| 285 | ||
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26425diff
changeset | 286 | val clasimp_setup = | 
| 30541 | 287 |   Attrib.setup @{binding iff} iff_att "declaration of Simplifier / Classical rules" #>
 | 
| 288 |   Method.setup @{binding fastsimp} (clasimp_method' fast_simp_tac) "combined fast and simp" #>
 | |
| 289 |   Method.setup @{binding slowsimp} (clasimp_method' slow_simp_tac) "combined slow and simp" #>
 | |
| 290 |   Method.setup @{binding bestsimp} (clasimp_method' best_simp_tac) "combined best and simp" #>
 | |
| 291 |   Method.setup @{binding force} (clasimp_method' force_tac) "force" #>
 | |
| 292 |   Method.setup @{binding auto} auto_method "auto" #>
 | |
| 293 |   Method.setup @{binding clarsimp} (clasimp_method' (CHANGED_PROP oo clarsimp_tac))
 | |
| 294 | "clarify simplified goal"; | |
| 5926 | 295 | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 296 | end; |