| author | nipkow | 
| Thu, 29 Sep 2005 15:31:34 +0200 | |
| changeset 17719 | 2e75155c5ed5 | 
| parent 17084 | fb0a80aef0be | 
| child 17879 | 88844eea4ce2 | 
| 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 | |
| 14 | structure Simplifier: SIMPLIFIER | |
| 8469 | 15 | structure Splitter: SPLITTER | 
| 5219 | 16 | structure Classical: CLASSICAL | 
| 17 | structure Blast: BLAST | |
| 18 | sharing type Classical.claset = Blast.claset | |
| 9860 | 19 | val notE: thm | 
| 20 | val iffD1: thm | |
| 21 | val iffD2: thm | |
| 22 | val cla_make_elim: thm -> thm | |
| 5219 | 23 | end | 
| 24 | ||
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 25 | signature CLASIMP = | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 26 | sig | 
| 5219 | 27 | include CLASIMP_DATA | 
| 28 | type claset | |
| 29 | type simpset | |
| 30 | type clasimpset | |
| 10317 | 31 | val clasimpset: unit -> clasimpset | 
| 9860 | 32 | val addSIs2: clasimpset * thm list -> clasimpset | 
| 33 | val addSEs2: clasimpset * thm list -> clasimpset | |
| 34 | val addSDs2: clasimpset * thm list -> clasimpset | |
| 35 | val addIs2: clasimpset * thm list -> clasimpset | |
| 36 | val addEs2: clasimpset * thm list -> clasimpset | |
| 37 | val addDs2: clasimpset * thm list -> clasimpset | |
| 38 | val addsimps2: clasimpset * thm list -> clasimpset | |
| 39 | val delsimps2: clasimpset * thm list -> clasimpset | |
| 40 | val addSss: claset * simpset -> claset | |
| 41 | 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 | 42 | val addss': claset * simpset -> claset | 
| 9860 | 43 | val addIffs: clasimpset * thm list -> clasimpset | 
| 44 | val delIffs: clasimpset * thm list -> clasimpset | |
| 45 | val AddIffs: thm list -> unit | |
| 46 | val DelIffs: thm list -> unit | |
| 47 | val CLASIMPSET: (clasimpset -> tactic) -> tactic | |
| 48 | val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic | |
| 5483 | 49 | val clarsimp_tac: clasimpset -> int -> tactic | 
| 9860 | 50 | val Clarsimp_tac: int -> tactic | 
| 51 | val mk_auto_tac: clasimpset -> int -> int -> tactic | |
| 52 | val auto_tac: clasimpset -> tactic | |
| 53 | val Auto_tac: tactic | |
| 54 | val auto: unit -> unit | |
| 55 | val force_tac: clasimpset -> int -> tactic | |
| 56 | val Force_tac: int -> tactic | |
| 57 | val force: int -> unit | |
| 58 | val fast_simp_tac: clasimpset -> int -> tactic | |
| 59 | val slow_simp_tac: clasimpset -> int -> tactic | |
| 60 | val best_simp_tac: clasimpset -> int -> tactic | |
| 8639 | 61 | val change_global_css: (clasimpset * thm list -> clasimpset) -> theory attribute | 
| 62 | val change_local_css: (clasimpset * thm list -> clasimpset) -> Proof.context attribute | |
| 9506 | 63 | val get_local_clasimpset: Proof.context -> clasimpset | 
| 15032 | 64 | val local_clasimpset_of: Proof.context -> clasimpset | 
| 9860 | 65 | val iff_add_global: theory attribute | 
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 66 | val iff_add_global': theory attribute | 
| 9860 | 67 | val iff_del_global: theory attribute | 
| 68 | val iff_add_local: Proof.context attribute | |
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 69 | val iff_add_local': Proof.context attribute | 
| 9860 | 70 | val iff_del_local: Proof.context attribute | 
| 71 | val iff_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | |
| 9506 | 72 | val clasimp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list | 
| 9860 | 73 | val setup: (theory -> theory) list | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 74 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 75 | |
| 5219 | 76 | 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 | 77 | struct | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 78 | |
| 5219 | 79 | open Data; | 
| 80 | ||
| 81 | type claset = Classical.claset; | |
| 82 | type simpset = Simplifier.simpset; | |
| 83 | type clasimpset = claset * simpset; | |
| 84 | ||
| 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 | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 137 | (elim (intro (cs, [suffix_thm name "_iff2" (zero_rotate (th RS Data.iffD2))]), | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 138 | [suffix_thm name "_iff1" (Data.cla_make_elim (zero_rotate (th RS Data.iffD1)))]) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 139 | handle THM _ => (elim (cs, [suffix_thm name "_iff" (zero_rotate (th RS Data.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 | 
| 12375 | 146 | (delcs (cs, [zero_rotate (th RS Data.iffD2), | 
| 11902 | 147 | Data.cla_make_elim (zero_rotate (th RS Data.iffD1))]) | 
| 12375 | 148 | handle THM _ => (delcs (cs, [zero_rotate (th RS Data.notE)]) | 
| 149 | handle THM _ => delcs (cs, [th])), | |
| 150 | delss (ss, [th])) | |
| 11902 | 151 | end; | 
| 9860 | 152 | |
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 153 | fun store_clasimp (cs, ss) = | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 154 | (Classical.claset_ref () := cs; Simplifier.simpset_ref () := ss); | 
| 9860 | 155 | |
| 156 | in | |
| 157 | ||
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 158 | val op addIffs = | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 159 | Library.foldl | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 160 | (addIff (Classical.addSEs, Classical.addSIs) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 161 | (Classical.addEs, Classical.addIs) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 162 | Simplifier.addsimps); | 
| 15570 | 163 | 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 | 164 | |
| 10317 | 165 | fun AddIffs thms = store_clasimp (clasimpset () addIffs thms); | 
| 166 | fun DelIffs thms = store_clasimp (clasimpset () delIffs thms); | |
| 9860 | 167 | |
| 12375 | 168 | fun addIffs_global (thy, ths) = | 
| 15570 | 169 | Library.foldl (addIff | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 170 | (ContextRules.addXEs_global, ContextRules.addXIs_global) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 171 | (ContextRules.addXEs_global, ContextRules.addXIs_global) #1) | 
| 13026 | 172 | ((thy, ()), ths) |> #1; | 
| 12375 | 173 | |
| 174 | fun addIffs_local (ctxt, ths) = | |
| 15570 | 175 | Library.foldl (addIff | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 176 | (ContextRules.addXEs_local, ContextRules.addXIs_local) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16019diff
changeset | 177 | (ContextRules.addXEs_local, ContextRules.addXIs_local) #1) | 
| 13026 | 178 | ((ctxt, ()), ths) |> #1; | 
| 12375 | 179 | |
| 180 | fun delIffs_global (thy, ths) = | |
| 15570 | 181 | Library.foldl (delIff ContextRules.delrules_global #1) ((thy, ()), ths) |> #1; | 
| 12375 | 182 | |
| 183 | fun delIffs_local (ctxt, ths) = | |
| 15570 | 184 | Library.foldl (delIff ContextRules.delrules_local #1) ((ctxt, ()), ths) |> #1; | 
| 12375 | 185 | |
| 9860 | 186 | end; | 
| 187 | ||
| 188 | ||
| 5219 | 189 | (* tacticals *) | 
| 190 | ||
| 191 | fun CLASIMPSET tacf state = | |
| 192 | Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state; | |
| 193 | ||
| 194 | fun CLASIMPSET' tacf i state = | |
| 195 | Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state; | |
| 4888 | 196 | |
| 197 | ||
| 12375 | 198 | fun clarsimp_tac (cs, ss) = | 
| 199 | safe_asm_full_simp_tac ss THEN_ALL_NEW | |
| 200 | Classical.clarify_tac (cs addSss ss); | |
| 201 | ||
| 10317 | 202 | fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i; | 
| 5483 | 203 | |
| 204 | ||
| 5219 | 205 | (* auto_tac *) | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 206 | |
| 5219 | 207 | fun blast_depth_tac cs m i thm = | 
| 9772 | 208 | Blast.depth_tac cs m i thm | 
| 5554 | 209 |       handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
 | 
| 9772 | 210 | |
| 211 | (* a variant of depth_tac that avoids interference of the simplifier | |
| 5219 | 212 | with dup_step_tac when they are combined by auto_tac *) | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 213 | local | 
| 9772 | 214 | fun slow_step_tac' cs = Classical.appWrappers cs | 
| 215 | (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs); | |
| 216 | in fun nodup_depth_tac cs m i state = SELECT_GOAL | |
| 217 | (Classical.safe_steps_tac cs 1 THEN_ELSE | |
| 218 | (DEPTH_SOLVE (nodup_depth_tac cs m 1), | |
| 219 | Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac | |
| 220 | (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 | 221 | )) i state; | 
| 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 222 | end; | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 223 | |
| 9402 | 224 | (*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 | 225 | in some of the subgoals*) | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 226 | fun mk_auto_tac (cs, ss) m n = | 
| 5219 | 227 | let val cs' = cs addss ss | 
| 9772 | 228 | val maintac = | 
| 229 | 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 | 230 | ORELSE' | 
| 5756 
8ef5288c24b0
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5567diff
changeset | 231 | (CHANGED o nodup_depth_tac cs' n); (* slower but more general *) | 
| 5219 | 232 | in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), | 
| 9772 | 233 | TRY (Classical.safe_tac cs), | 
| 234 | REPEAT (FIRSTGOAL maintac), | |
| 5219 | 235 | TRY (Classical.safe_tac (cs addSss ss)), | 
| 9772 | 236 | prune_params_tac] | 
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 237 | end; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 238 | |
| 9772 | 239 | fun auto_tac css = mk_auto_tac css 4 2; | 
| 10317 | 240 | 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 | 241 | fun auto () = by Auto_tac; | 
| 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 242 | |
| 9772 | 243 | |
| 5219 | 244 | (* force_tac *) | 
| 245 | ||
| 4659 | 246 | (* 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 | 247 | fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [ | 
| 9772 | 248 | Classical.clarify_tac cs' 1, | 
| 249 | IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1), | |
| 250 | ALLGOALS (Classical.first_best_tac cs')]) end; | |
| 10317 | 251 | 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 | 252 | fun force i = by (Force_tac i); | 
| 4659 | 253 | |
| 5219 | 254 | |
| 9805 | 255 | (* basic combinations *) | 
| 256 | ||
| 257 | fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end; | |
| 9591 | 258 | |
| 9805 | 259 | val fast_simp_tac = ADDSS Classical.fast_tac; | 
| 260 | val slow_simp_tac = ADDSS Classical.slow_tac; | |
| 261 | val best_simp_tac = ADDSS Classical.best_tac; | |
| 9591 | 262 | |
| 263 | ||
| 9860 | 264 | (* access clasimpset *) | 
| 8639 | 265 | |
| 266 | fun change_global_css f (thy, th) = | |
| 267 | let | |
| 268 | val cs_ref = Classical.claset_ref_of thy; | |
| 269 | val ss_ref = Simplifier.simpset_ref_of thy; | |
| 270 | val (cs', ss') = f ((! cs_ref, ! ss_ref), [th]); | |
| 271 | in cs_ref := cs'; ss_ref := ss'; (thy, th) end; | |
| 272 | ||
| 273 | fun change_local_css f (ctxt, th) = | |
| 274 | let | |
| 275 | val cs = Classical.get_local_claset ctxt; | |
| 276 | val ss = Simplifier.get_local_simpset ctxt; | |
| 277 | val (cs', ss') = f ((cs, ss), [th]); | |
| 278 | val ctxt' = | |
| 279 | ctxt | |
| 280 | |> Classical.put_local_claset cs' | |
| 281 | |> Simplifier.put_local_simpset ss'; | |
| 282 | in (ctxt', th) end; | |
| 283 | ||
| 7153 | 284 | fun get_local_clasimpset ctxt = | 
| 285 | (Classical.get_local_claset ctxt, Simplifier.get_local_simpset ctxt); | |
| 5926 | 286 | |
| 15032 | 287 | fun local_clasimpset_of ctxt = | 
| 288 | (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt); | |
| 289 | ||
| 9860 | 290 | |
| 291 | (* attributes *) | |
| 292 | ||
| 12375 | 293 | fun change_rules f (x, th) = (f (x, [th]), th); | 
| 294 | ||
| 9860 | 295 | val iff_add_global = change_global_css (op addIffs); | 
| 12375 | 296 | val iff_add_global' = change_rules addIffs_global; | 
| 297 | val iff_del_global = change_global_css (op delIffs) o change_rules delIffs_global; | |
| 9860 | 298 | val iff_add_local = change_local_css (op addIffs); | 
| 12375 | 299 | val iff_add_local' = change_rules addIffs_local; | 
| 300 | val iff_del_local = change_local_css (op delIffs) o change_rules delIffs_local; | |
| 9860 | 301 | |
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 302 | fun iff_att add add' del = Attrib.syntax (Scan.lift | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 303 | (Args.del >> K del || | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 304 | Scan.option Args.add -- Args.query >> K add' || | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 305 | Scan.option Args.add >> K add)); | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 306 | |
| 9860 | 307 | val iff_attr = | 
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 308 | (iff_att iff_add_global iff_add_global' iff_del_global, | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 309 | iff_att iff_add_local iff_add_local' iff_del_local); | 
| 9860 | 310 | |
| 311 | ||
| 312 | (* method modifiers *) | |
| 313 | ||
| 314 | val iffN = "iff"; | |
| 315 | ||
| 316 | val iff_modifiers = | |
| 10033 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 317 | [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add_local): Method.modifier), | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 318 | Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add_local'), | 
| 
fc4e7432b2b1
added iff_add_global', iff_add_local' (syntax "iff?");
 wenzelm parents: 
9952diff
changeset | 319 | Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del_local)]; | 
| 9860 | 320 | |
| 8469 | 321 | val clasimp_modifiers = | 
| 9860 | 322 | Simplifier.simp_modifiers @ Splitter.split_modifiers @ | 
| 323 | Classical.cla_modifiers @ iff_modifiers; | |
| 324 | ||
| 325 | ||
| 326 | (* methods *) | |
| 5926 | 327 | |
| 7559 | 328 | fun clasimp_meth tac prems ctxt = Method.METHOD (fn facts => | 
| 15032 | 329 | ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt)); | 
| 7132 | 330 | |
| 7559 | 331 | fun clasimp_meth' tac prems ctxt = Method.METHOD (fn facts => | 
| 15032 | 332 | HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt))); | 
| 5926 | 333 | |
| 7559 | 334 | val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth; | 
| 335 | val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth'; | |
| 5926 | 336 | |
| 337 | ||
| 9772 | 338 | fun auto_args m = | 
| 339 | Method.bang_sectioned_args' clasimp_modifiers (Scan.lift (Scan.option (Args.nat -- Args.nat))) m; | |
| 340 | ||
| 15531 | 341 | fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac) | 
| 342 | | auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n)); | |
| 9772 | 343 | |
| 344 | ||
| 345 | (* theory setup *) | |
| 346 | ||
| 5926 | 347 | val setup = | 
| 9860 | 348 | [Attrib.add_attributes | 
| 9893 | 349 |   [("iff", iff_attr, "declaration of Simplifier / Classical rules")],
 | 
| 9860 | 350 | Method.add_methods | 
| 9591 | 351 |    [("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"),
 | 
| 9805 | 352 |     ("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"),
 | 
| 353 |     ("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"),
 | |
| 9591 | 354 |     ("force", clasimp_method' force_tac, "force"),
 | 
| 9772 | 355 |     ("auto", auto_args auto_meth, "auto"),
 | 
| 10821 | 356 |     ("clarsimp", clasimp_method' (CHANGED_PROP oo clarsimp_tac), "clarify simplified goal")]];
 | 
| 5926 | 357 | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: diff
changeset | 358 | end; |