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