| author | paulson | 
| Wed, 22 Nov 2006 20:08:07 +0100 | |
| changeset 21470 | 7c1b59ddcd56 | 
| parent 21313 | 26fc3a45547c | 
| child 21551 | d276e7d25017 | 
| permissions | -rw-r--r-- | 
| 21163 | 1 | (* Title: HOL/simpdata.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1991 University of Cambridge | |
| 5 | ||
| 6 | Instantiation of the generic simplifier for HOL. | |
| 7 | *) | |
| 8 | ||
| 9 | (** tools setup **) | |
| 10 | ||
| 11 | structure Quantifier1 = Quantifier1Fun | |
| 12 | (struct | |
| 13 | (*abstract syntax*) | |
| 14 |   fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t)
 | |
| 15 | | dest_eq _ = NONE; | |
| 16 |   fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t)
 | |
| 17 | | dest_conj _ = NONE; | |
| 18 |   fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t)
 | |
| 19 | | dest_imp _ = NONE; | |
| 20 | val conj = HOLogic.conj | |
| 21 | val imp = HOLogic.imp | |
| 22 | (*rules*) | |
| 23 | val iff_reflection = HOL.eq_reflection | |
| 24 | val iffI = HOL.iffI | |
| 25 | val iff_trans = HOL.trans | |
| 26 | val conjI= HOL.conjI | |
| 27 | val conjE= HOL.conjE | |
| 28 | val impI = HOL.impI | |
| 29 | val mp = HOL.mp | |
| 30 | val uncurry = thm "uncurry" | |
| 31 | val exI = HOL.exI | |
| 32 | val exE = HOL.exE | |
| 33 | val iff_allI = thm "iff_allI" | |
| 34 | val iff_exI = thm "iff_exI" | |
| 35 | val all_comm = thm "all_comm" | |
| 36 | val ex_comm = thm "ex_comm" | |
| 37 | end); | |
| 38 | ||
| 39 | structure HOL = | |
| 40 | struct | |
| 41 | ||
| 42 | open HOL; | |
| 43 | ||
| 44 | val Eq_FalseI = thm "Eq_FalseI"; | |
| 45 | val Eq_TrueI = thm "Eq_TrueI"; | |
| 46 | val simp_implies_def = thm "simp_implies_def"; | |
| 47 | val simp_impliesI = thm "simp_impliesI"; | |
| 48 | ||
| 49 | fun mk_meta_eq r = r RS eq_reflection; | |
| 50 | fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r; | |
| 51 | ||
| 52 | fun mk_eq thm = case concl_of thm | |
| 53 | (*expects Trueprop if not == *) | |
| 54 |   of Const ("==",_) $ _ $ _ => thm
 | |
| 55 |    | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq thm
 | |
| 56 |    | _ $ (Const ("Not", _) $ _) => thm RS Eq_FalseI
 | |
| 57 | | _ => thm RS Eq_TrueI; | |
| 58 | ||
| 59 | fun mk_eq_True r = | |
| 60 | SOME (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE; | |
| 61 | ||
| 62 | (* Produce theorems of the form | |
| 63 | (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y) | |
| 64 | *) | |
| 65 | fun lift_meta_eq_to_obj_eq i st = | |
| 66 | let | |
| 67 |     fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
 | |
| 68 | | count_imp _ = 0; | |
| 69 | val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) | |
| 70 | in if j = 0 then meta_eq_to_obj_eq | |
| 71 | else | |
| 72 | let | |
| 73 |         val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
 | |
| 74 | fun mk_simp_implies Q = foldr (fn (R, S) => | |
| 75 |           Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps
 | |
| 76 |         val aT = TFree ("'a", HOLogic.typeS);
 | |
| 77 |         val x = Free ("x", aT);
 | |
| 78 |         val y = Free ("y", aT)
 | |
| 79 | in Goal.prove_global (Thm.theory_of_thm st) [] | |
| 80 | [mk_simp_implies (Logic.mk_equals (x, y))] | |
| 81 | (mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y)))) | |
| 82 | (fn prems => EVERY | |
| 83 | [rewrite_goals_tac [simp_implies_def], | |
| 84 | REPEAT (ares_tac (meta_eq_to_obj_eq :: map (rewrite_rule [simp_implies_def]) prems) 1)]) | |
| 85 | end | |
| 86 | end; | |
| 87 | ||
| 88 | (*Congruence rules for = (instead of ==)*) | |
| 89 | fun mk_meta_cong rl = zero_var_indexes | |
| 90 | (let val rl' = Seq.hd (TRYALL (fn i => fn st => | |
| 91 | rtac (lift_meta_eq_to_obj_eq i st) i st) rl) | |
| 92 | in mk_meta_eq rl' handle THM _ => | |
| 93 | if can Logic.dest_equals (concl_of rl') then rl' | |
| 94 | else error "Conclusion of congruence rules must be =-equality" | |
| 95 | end); | |
| 96 | ||
| 97 | fun mk_atomize pairs = | |
| 98 | let | |
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 99 | fun atoms thm = | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 100 | let | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 101 | fun res th = map (fn rl => th RS rl); (*exception THM*) | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 102 | fun res_fixed rls = | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 103 | if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 104 | else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.thm_context thm) [thm]; | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 105 | in | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 106 | case concl_of thm | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 107 |           of Const ("Trueprop", _) $ p => (case head_of p
 | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 108 | of Const (a, _) => (case AList.lookup (op =) pairs a | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 109 | of SOME rls => (maps atoms (res_fixed rls) handle THM _ => [thm]) | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 110 | | NONE => [thm]) | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 111 | | _ => [thm]) | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 112 | | _ => [thm] | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 113 | end; | 
| 21163 | 114 | in atoms end; | 
| 115 | ||
| 116 | fun mksimps pairs = | |
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 117 | map_filter (try mk_eq) o mk_atomize pairs o gen_all; | 
| 21163 | 118 | |
| 119 | fun unsafe_solver_tac prems = | |
| 120 | (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN' | |
| 121 | FIRST'[resolve_tac(reflexive_thm :: TrueI :: refl :: prems), atac, etac FalseE]; | |
| 122 | val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; | |
| 123 | ||
| 124 | (*No premature instantiation of variables during simplification*) | |
| 125 | fun safe_solver_tac prems = | |
| 126 | (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN' | |
| 127 | FIRST'[match_tac(reflexive_thm :: TrueI :: refl :: prems), | |
| 128 | eq_assume_tac, ematch_tac [FalseE]]; | |
| 129 | val safe_solver = mk_solver "HOL safe" safe_solver_tac; | |
| 130 | ||
| 131 | end; | |
| 132 | ||
| 133 | structure SplitterData = | |
| 134 | struct | |
| 135 | structure Simplifier = Simplifier | |
| 136 | val mk_eq = HOL.mk_eq | |
| 137 | val meta_eq_to_iff = HOL.meta_eq_to_obj_eq | |
| 138 | val iffD = HOL.iffD2 | |
| 139 | val disjE = HOL.disjE | |
| 140 | val conjE = HOL.conjE | |
| 141 | val exE = HOL.exE | |
| 142 | val contrapos = HOL.contrapos_nn | |
| 143 | val contrapos2 = HOL.contrapos_pp | |
| 144 | val notnotD = HOL.notnotD | |
| 145 | end; | |
| 146 | ||
| 147 | structure Splitter = SplitterFun(SplitterData); | |
| 148 | ||
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 149 | |
| 21163 | 150 | (* integration of simplifier with classical reasoner *) | 
| 151 | ||
| 152 | structure Clasimp = ClasimpFun | |
| 153 | (structure Simplifier = Simplifier and Splitter = Splitter | |
| 154 | and Classical = Classical and Blast = Blast | |
| 155 | val iffD1 = HOL.iffD1 val iffD2 = HOL.iffD2 val notE = HOL.notE); | |
| 156 | ||
| 157 | structure HOL = | |
| 158 | struct | |
| 159 | ||
| 160 | open HOL; | |
| 161 | ||
| 162 | val mksimps_pairs = | |
| 163 |   [("op -->", [mp]), ("op &", [thm "conjunct1", thm "conjunct2"]),
 | |
| 164 |    ("All", [spec]), ("True", []), ("False", []),
 | |
| 165 |    ("HOL.If", [thm "if_bool_eq_conj" RS iffD1])];
 | |
| 166 | ||
| 167 | val simpset_basic = | |
| 168 | Simplifier.theory_context (the_context ()) empty_ss | |
| 169 | setsubgoaler asm_simp_tac | |
| 170 | setSSolver safe_solver | |
| 171 | setSolver unsafe_solver | |
| 172 | setmksimps (mksimps mksimps_pairs) | |
| 173 | setmkeqTrue mk_eq_True | |
| 174 | setmkcong mk_meta_cong; | |
| 175 | ||
| 176 | fun simplify rews = Simplifier.full_simplify (simpset_basic addsimps rews); | |
| 177 | ||
| 178 | fun unfold_tac ths = | |
| 179 | let val ss0 = Simplifier.clear_ss simpset_basic addsimps ths | |
| 180 | in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; | |
| 181 | ||
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 182 | |
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 183 | |
| 21163 | 184 | (** simprocs **) | 
| 185 | ||
| 186 | (* simproc for proving "(y = x) == False" from premise "~(x = y)" *) | |
| 187 | ||
| 188 | val use_neq_simproc = ref true; | |
| 189 | ||
| 190 | local | |
| 191 | val thy = the_context (); | |
| 192 | val neq_to_EQ_False = thm "not_sym" RS HOL.Eq_FalseI; | |
| 193 | fun neq_prover sg ss (eq $ lhs $ rhs) = | |
| 194 | let | |
| 195 | fun test thm = (case #prop (rep_thm thm) of | |
| 196 | _ $ (Not $ (eq' $ l' $ r')) => | |
| 197 | Not = HOLogic.Not andalso eq' = eq andalso | |
| 198 | r' aconv lhs andalso l' aconv rhs | |
| 199 | | _ => false) | |
| 200 | in if !use_neq_simproc then case find_first test (prems_of_ss ss) | |
| 201 | of NONE => NONE | |
| 202 | | SOME thm => SOME (thm RS neq_to_EQ_False) | |
| 203 | else NONE | |
| 204 | end | |
| 205 | in | |
| 206 | ||
| 207 | val neq_simproc = Simplifier.simproc thy "neq_simproc" ["x = y"] neq_prover; | |
| 208 | ||
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 209 | end; | 
| 21163 | 210 | |
| 211 | ||
| 212 | (* simproc for Let *) | |
| 213 | ||
| 214 | val use_let_simproc = ref true; | |
| 215 | ||
| 216 | local | |
| 217 | val thy = the_context (); | |
| 218 | val Let_folded = thm "Let_folded"; | |
| 219 | val Let_unfold = thm "Let_unfold"; | |
| 220 | val (f_Let_unfold, x_Let_unfold) = | |
| 221 | let val [(_$(f$x)$_)] = prems_of Let_unfold | |
| 222 | in (cterm_of thy f, cterm_of thy x) end | |
| 223 | val (f_Let_folded, x_Let_folded) = | |
| 224 | let val [(_$(f$x)$_)] = prems_of Let_folded | |
| 225 | in (cterm_of thy f, cterm_of thy x) end; | |
| 226 | val g_Let_folded = | |
| 227 | let val [(_$_$(g$_))] = prems_of Let_folded in cterm_of thy g end; | |
| 228 | in | |
| 229 | ||
| 230 | val let_simproc = | |
| 231 | Simplifier.simproc thy "let_simp" ["Let x f"] | |
| 232 | (fn sg => fn ss => fn t => | |
| 233 | let val ctxt = Simplifier.the_context ss; | |
| 234 | val ([t'], ctxt') = Variable.import_terms false [t] ctxt; | |
| 235 | in Option.map (hd o Variable.export ctxt' ctxt o single) | |
| 236 |       (case t' of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
 | |
| 237 | if not (!use_let_simproc) then NONE | |
| 238 | else if is_Free x orelse is_Bound x orelse is_Const x | |
| 239 | then SOME (thm "Let_def") | |
| 240 | else | |
| 241 | let | |
| 242 | val n = case f of (Abs (x,_,_)) => x | _ => "x"; | |
| 243 | val cx = cterm_of sg x; | |
| 244 |              val {T=xT,...} = rep_cterm cx;
 | |
| 245 | val cf = cterm_of sg f; | |
| 246 | val fx_g = Simplifier.rewrite ss (Thm.capply cf cx); | |
| 247 | val (_$_$g) = prop_of fx_g; | |
| 248 | val g' = abstract_over (x,g); | |
| 249 | in (if (g aconv g') | |
| 250 | then | |
| 251 | let | |
| 252 | val rl = cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] Let_unfold; | |
| 253 | in SOME (rl OF [fx_g]) end | |
| 254 | else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*) | |
| 255 | else let | |
| 256 | val abs_g'= Abs (n,xT,g'); | |
| 257 | val g'x = abs_g'$x; | |
| 258 | val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x)); | |
| 259 | val rl = cterm_instantiate | |
| 260 | [(f_Let_folded,cterm_of sg f),(x_Let_folded,cx), | |
| 261 | (g_Let_folded,cterm_of sg abs_g')] | |
| 262 | Let_folded; | |
| 263 | in SOME (rl OF [transitive fx_g g_g'x]) | |
| 264 | end) | |
| 265 | end | |
| 266 | | _ => NONE) | |
| 267 | end) | |
| 268 | ||
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 269 | end; | 
| 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 270 | |
| 21163 | 271 | |
| 272 | (* generic refutation procedure *) | |
| 273 | ||
| 274 | (* parameters: | |
| 275 | ||
| 276 | test: term -> bool | |
| 277 | tests if a term is at all relevant to the refutation proof; | |
| 278 | if not, then it can be discarded. Can improve performance, | |
| 279 | esp. if disjunctions can be discarded (no case distinction needed!). | |
| 280 | ||
| 281 | prep_tac: int -> tactic | |
| 282 | A preparation tactic to be applied to the goal once all relevant premises | |
| 283 | have been moved to the conclusion. | |
| 284 | ||
| 285 | ref_tac: int -> tactic | |
| 286 | the actual refutation tactic. Should be able to deal with goals | |
| 287 | [| A1; ...; An |] ==> False | |
| 288 | where the Ai are atomic, i.e. no top-level &, | or EX | |
| 289 | *) | |
| 290 | ||
| 291 | local | |
| 292 | val nnf_simpset = | |
| 293 | empty_ss setmkeqTrue mk_eq_True | |
| 294 | setmksimps (mksimps mksimps_pairs) | |
| 295 | addsimps [thm "imp_conv_disj", thm "iff_conv_conj_imp", thm "de_Morgan_disj", thm "de_Morgan_conj", | |
| 296 | thm "not_all", thm "not_ex", thm "not_not"]; | |
| 297 | fun prem_nnf_tac i st = | |
| 298 | full_simp_tac (Simplifier.theory_context (Thm.theory_of_thm st) nnf_simpset) i st; | |
| 299 | in | |
| 300 | fun refute_tac test prep_tac ref_tac = | |
| 301 | let val refute_prems_tac = | |
| 302 | REPEAT_DETERM | |
| 303 | (eresolve_tac [conjE, exE] 1 ORELSE | |
| 304 | filter_prems_tac test 1 ORELSE | |
| 305 | etac disjE 1) THEN | |
| 306 | ((etac notE 1 THEN eq_assume_tac 1) ORELSE | |
| 307 | ref_tac 1); | |
| 308 | in EVERY'[TRY o filter_prems_tac test, | |
| 309 | REPEAT_DETERM o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac, | |
| 310 | SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] | |
| 311 | end; | |
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 312 | end; | 
| 21163 | 313 | |
| 314 | val defALL_regroup = | |
| 315 | Simplifier.simproc (the_context ()) | |
| 316 | "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; | |
| 317 | ||
| 318 | val defEX_regroup = | |
| 319 | Simplifier.simproc (the_context ()) | |
| 320 | "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; | |
| 321 | ||
| 322 | ||
| 323 | val simpset_simprocs = simpset_basic | |
| 324 | addsimprocs [defALL_regroup, defEX_regroup, neq_simproc, let_simproc] | |
| 325 | ||
| 21313 
26fc3a45547c
mk_atomize: careful matching against rules admits overloading;
 wenzelm parents: 
21163diff
changeset | 326 | end; |