| author | wenzelm | 
| Mon, 23 Aug 2010 20:50:00 +0200 | |
| changeset 38657 | 2e0ebdaac59b | 
| parent 38522 | de7984a7172b | 
| child 38800 | 34c84817e39c | 
| permissions | -rw-r--r-- | 
| 1477 | 1 | (* Title: FOLP/IFOLP.thy | 
| 2 | Author: Martin D Coen, Cambridge University Computer Laboratory | |
| 1142 | 3 | Copyright 1992 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 17480 | 6 | header {* Intuitionistic First-Order Logic with Proofs *}
 | 
| 7 | ||
| 8 | theory IFOLP | |
| 9 | imports Pure | |
| 26322 | 10 | uses ("hypsubst.ML") ("intprover.ML")
 | 
| 17480 | 11 | begin | 
| 0 | 12 | |
| 26956 
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
 wenzelm parents: 
26480diff
changeset | 13 | setup PureThy.old_appl_syntax_setup | 
| 
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
 wenzelm parents: 
26480diff
changeset | 14 | |
| 17480 | 15 | classes "term" | 
| 36452 | 16 | default_sort "term" | 
| 0 | 17 | |
| 17480 | 18 | typedecl p | 
| 19 | typedecl o | |
| 0 | 20 | |
| 17480 | 21 | consts | 
| 0 | 22 | (*** Judgements ***) | 
| 1477 | 23 | Proof :: "[o,p]=>prop" | 
| 0 | 24 |  EqProof        ::   "[p,p,o]=>prop"    ("(3_ /= _ :/ _)" [10,10,10] 5)
 | 
| 17480 | 25 | |
| 0 | 26 | (*** Logical Connectives -- Type Formers ***) | 
| 35128 | 27 | "op =" :: "['a,'a] => o" (infixl "=" 50) | 
| 17480 | 28 | True :: "o" | 
| 29 | False :: "o" | |
| 2714 | 30 |  Not            ::      "o => o"        ("~ _" [40] 40)
 | 
| 35128 | 31 | "op &" :: "[o,o] => o" (infixr "&" 35) | 
| 32 | "op |" :: "[o,o] => o" (infixr "|" 30) | |
| 33 | "op -->" :: "[o,o] => o" (infixr "-->" 25) | |
| 34 | "op <->" :: "[o,o] => o" (infixr "<->" 25) | |
| 0 | 35 | (*Quantifiers*) | 
| 1477 | 36 |  All            ::      "('a => o) => o"        (binder "ALL " 10)
 | 
| 37 |  Ex             ::      "('a => o) => o"        (binder "EX " 10)
 | |
| 38 |  Ex1            ::      "('a => o) => o"        (binder "EX! " 10)
 | |
| 0 | 39 | (*Rewriting gadgets*) | 
| 1477 | 40 | NORM :: "o => o" | 
| 41 | norm :: "'a => 'a" | |
| 0 | 42 | |
| 648 
e27c9ec2b48b
FOLP/IFOLP.thy: tightening precedences to eliminate syntactic ambiguities.
 lcp parents: 
283diff
changeset | 43 | (*** Proof Term Formers: precedence must exceed 50 ***) | 
| 1477 | 44 | tt :: "p" | 
| 45 | contr :: "p=>p" | |
| 17480 | 46 | fst :: "p=>p" | 
| 47 | snd :: "p=>p" | |
| 1477 | 48 |  pair           :: "[p,p]=>p"           ("(1<_,/_>)")
 | 
| 49 | split :: "[p, [p,p]=>p] =>p" | |
| 17480 | 50 | inl :: "p=>p" | 
| 51 | inr :: "p=>p" | |
| 1477 | 52 | when :: "[p, p=>p, p=>p]=>p" | 
| 53 | lambda :: "(p => p) => p" (binder "lam " 55) | |
| 35128 | 54 | "op `" :: "[p,p]=>p" (infixl "`" 60) | 
| 648 
e27c9ec2b48b
FOLP/IFOLP.thy: tightening precedences to eliminate syntactic ambiguities.
 lcp parents: 
283diff
changeset | 55 | alll :: "['a=>p]=>p" (binder "all " 55) | 
| 35128 | 56 | "op ^" :: "[p,'a]=>p" (infixl "^" 55) | 
| 1477 | 57 |  exists         :: "['a,p]=>p"          ("(1[_,/_])")
 | 
| 0 | 58 | xsplit :: "[p,['a,p]=>p]=>p" | 
| 59 | ideq :: "'a=>p" | |
| 60 | idpeel :: "[p,'a=>p]=>p" | |
| 17480 | 61 | nrm :: p | 
| 62 | NRM :: p | |
| 0 | 63 | |
| 35113 | 64 | syntax "_Proof" :: "[p,o]=>prop"    ("(_ /: _)" [51, 10] 5)
 | 
| 65 | ||
| 17480 | 66 | ML {*
 | 
| 67 | ||
| 68 | (*show_proofs:=true displays the proof terms -- they are ENORMOUS*) | |
| 32740 | 69 | val show_proofs = Unsynchronized.ref false; | 
| 17480 | 70 | |
| 26322 | 71 | fun proof_tr [p,P] = Const (@{const_name Proof}, dummyT) $ P $ p;
 | 
| 17480 | 72 | |
| 73 | fun proof_tr' [P,p] = | |
| 35113 | 74 |   if ! show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
 | 
| 75 | else P (*this case discards the proof term*); | |
| 17480 | 76 | *} | 
| 77 | ||
| 35113 | 78 | parse_translation {* [(@{syntax_const "_Proof"}, proof_tr)] *}
 | 
| 79 | print_translation {* [(@{const_syntax Proof}, proof_tr')] *}
 | |
| 17480 | 80 | |
| 81 | axioms | |
| 0 | 82 | |
| 83 | (**** Propositional logic ****) | |
| 84 | ||
| 85 | (*Equality*) | |
| 86 | (* Like Intensional Equality in MLTT - but proofs distinct from terms *) | |
| 87 | ||
| 17480 | 88 | ieqI: "ideq(a) : a=a" | 
| 89 | ieqE: "[| p : a=b; !!x. f(x) : P(x,x) |] ==> idpeel(p,f) : P(a,b)" | |
| 0 | 90 | |
| 91 | (* Truth and Falsity *) | |
| 92 | ||
| 17480 | 93 | TrueI: "tt : True" | 
| 94 | FalseE: "a:False ==> contr(a):P" | |
| 0 | 95 | |
| 96 | (* Conjunction *) | |
| 97 | ||
| 17480 | 98 | conjI: "[| a:P; b:Q |] ==> <a,b> : P&Q" | 
| 99 | conjunct1: "p:P&Q ==> fst(p):P" | |
| 100 | conjunct2: "p:P&Q ==> snd(p):Q" | |
| 0 | 101 | |
| 102 | (* Disjunction *) | |
| 103 | ||
| 17480 | 104 | disjI1: "a:P ==> inl(a):P|Q" | 
| 105 | disjI2: "b:Q ==> inr(b):P|Q" | |
| 106 | disjE: "[| a:P|Q; !!x. x:P ==> f(x):R; !!x. x:Q ==> g(x):R | |
| 107 | |] ==> when(a,f,g):R" | |
| 0 | 108 | |
| 109 | (* Implication *) | |
| 110 | ||
| 17480 | 111 | impI: "(!!x. x:P ==> f(x):Q) ==> lam x. f(x):P-->Q" | 
| 112 | mp: "[| f:P-->Q; a:P |] ==> f`a:Q" | |
| 0 | 113 | |
| 114 | (*Quantifiers*) | |
| 115 | ||
| 17480 | 116 | allI: "(!!x. f(x) : P(x)) ==> all x. f(x) : ALL x. P(x)" | 
| 117 | spec: "(f:ALL x. P(x)) ==> f^x : P(x)" | |
| 0 | 118 | |
| 17480 | 119 | exI: "p : P(x) ==> [x,p] : EX x. P(x)" | 
| 120 | exE: "[| p: EX x. P(x); !!x u. u:P(x) ==> f(x,u) : R |] ==> xsplit(p,f):R" | |
| 0 | 121 | |
| 122 | (**** Equality between proofs ****) | |
| 123 | ||
| 17480 | 124 | prefl: "a : P ==> a = a : P" | 
| 125 | psym: "a = b : P ==> b = a : P" | |
| 126 | ptrans: "[| a = b : P; b = c : P |] ==> a = c : P" | |
| 0 | 127 | |
| 17480 | 128 | idpeelB: "[| !!x. f(x) : P(x,x) |] ==> idpeel(ideq(a),f) = f(a) : P(a,a)" | 
| 0 | 129 | |
| 17480 | 130 | fstB: "a:P ==> fst(<a,b>) = a : P" | 
| 131 | sndB: "b:Q ==> snd(<a,b>) = b : Q" | |
| 132 | pairEC: "p:P&Q ==> p = <fst(p),snd(p)> : P&Q" | |
| 0 | 133 | |
| 17480 | 134 | whenBinl: "[| a:P; !!x. x:P ==> f(x) : Q |] ==> when(inl(a),f,g) = f(a) : Q" | 
| 135 | whenBinr: "[| b:P; !!x. x:P ==> g(x) : Q |] ==> when(inr(b),f,g) = g(b) : Q" | |
| 136 | plusEC: "a:P|Q ==> when(a,%x. inl(x),%y. inr(y)) = a : P|Q" | |
| 0 | 137 | |
| 17480 | 138 | applyB: "[| a:P; !!x. x:P ==> b(x) : Q |] ==> (lam x. b(x)) ` a = b(a) : Q" | 
| 139 | funEC: "f:P ==> f = lam x. f`x : P" | |
| 0 | 140 | |
| 17480 | 141 | specB: "[| !!x. f(x) : P(x) |] ==> (all x. f(x)) ^ a = f(a) : P(a)" | 
| 0 | 142 | |
| 143 | ||
| 144 | (**** Definitions ****) | |
| 145 | ||
| 17480 | 146 | not_def: "~P == P-->False" | 
| 147 | iff_def: "P<->Q == (P-->Q) & (Q-->P)" | |
| 0 | 148 | |
| 149 | (*Unique existence*) | |
| 17480 | 150 | ex1_def: "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" | 
| 0 | 151 | |
| 152 | (*Rewriting -- special constants to flag normalized terms and formulae*) | |
| 17480 | 153 | norm_eq: "nrm : norm(x) = x" | 
| 154 | NORM_iff: "NRM : NORM(P) <-> P" | |
| 155 | ||
| 26322 | 156 | (*** Sequent-style elimination rules for & --> and ALL ***) | 
| 157 | ||
| 36319 | 158 | schematic_lemma conjE: | 
| 26322 | 159 | assumes "p:P&Q" | 
| 160 | and "!!x y.[| x:P; y:Q |] ==> f(x,y):R" | |
| 161 | shows "?a:R" | |
| 162 | apply (rule assms(2)) | |
| 163 | apply (rule conjunct1 [OF assms(1)]) | |
| 164 | apply (rule conjunct2 [OF assms(1)]) | |
| 165 | done | |
| 166 | ||
| 36319 | 167 | schematic_lemma impE: | 
| 26322 | 168 | assumes "p:P-->Q" | 
| 169 | and "q:P" | |
| 170 | and "!!x. x:Q ==> r(x):R" | |
| 171 | shows "?p:R" | |
| 172 | apply (rule assms mp)+ | |
| 173 | done | |
| 174 | ||
| 36319 | 175 | schematic_lemma allE: | 
| 26322 | 176 | assumes "p:ALL x. P(x)" | 
| 177 | and "!!y. y:P(x) ==> q(y):R" | |
| 178 | shows "?p:R" | |
| 179 | apply (rule assms spec)+ | |
| 180 | done | |
| 181 | ||
| 182 | (*Duplicates the quantifier; for use with eresolve_tac*) | |
| 36319 | 183 | schematic_lemma all_dupE: | 
| 26322 | 184 | assumes "p:ALL x. P(x)" | 
| 185 | and "!!y z.[| y:P(x); z:ALL x. P(x) |] ==> q(y,z):R" | |
| 186 | shows "?p:R" | |
| 187 | apply (rule assms spec)+ | |
| 188 | done | |
| 189 | ||
| 190 | ||
| 191 | (*** Negation rules, which translate between ~P and P-->False ***) | |
| 192 | ||
| 36319 | 193 | schematic_lemma notI: | 
| 26322 | 194 | assumes "!!x. x:P ==> q(x):False" | 
| 195 | shows "?p:~P" | |
| 196 | unfolding not_def | |
| 197 | apply (assumption | rule assms impI)+ | |
| 198 | done | |
| 199 | ||
| 36319 | 200 | schematic_lemma notE: "p:~P \<Longrightarrow> q:P \<Longrightarrow> ?p:R" | 
| 26322 | 201 | unfolding not_def | 
| 202 | apply (drule (1) mp) | |
| 203 | apply (erule FalseE) | |
| 204 | done | |
| 205 | ||
| 206 | (*This is useful with the special implication rules for each kind of P. *) | |
| 36319 | 207 | schematic_lemma not_to_imp: | 
| 26322 | 208 | assumes "p:~P" | 
| 209 | and "!!x. x:(P-->False) ==> q(x):Q" | |
| 210 | shows "?p:Q" | |
| 211 | apply (assumption | rule assms impI notE)+ | |
| 212 | done | |
| 213 | ||
| 214 | (* For substitution int an assumption P, reduce Q to P-->Q, substitute into | |
| 27150 | 215 | this implication, then apply impI to move P back into the assumptions.*) | 
| 36319 | 216 | schematic_lemma rev_mp: "[| p:P; q:P --> Q |] ==> ?p:Q" | 
| 26322 | 217 | apply (assumption | rule mp)+ | 
| 218 | done | |
| 219 | ||
| 220 | ||
| 221 | (*Contrapositive of an inference rule*) | |
| 36319 | 222 | schematic_lemma contrapos: | 
| 26322 | 223 | assumes major: "p:~Q" | 
| 224 | and minor: "!!y. y:P==>q(y):Q" | |
| 225 | shows "?a:~P" | |
| 226 | apply (rule major [THEN notE, THEN notI]) | |
| 227 | apply (erule minor) | |
| 228 | done | |
| 229 | ||
| 230 | (** Unique assumption tactic. | |
| 231 | Ignores proof objects. | |
| 232 | Fails unless one assumption is equal and exactly one is unifiable | |
| 233 | **) | |
| 234 | ||
| 235 | ML {*
 | |
| 236 | local | |
| 237 |   fun discard_proof (Const (@{const_name Proof}, _) $ P $ _) = P;
 | |
| 238 | in | |
| 239 | val uniq_assume_tac = | |
| 240 | SUBGOAL | |
| 241 | (fn (prem,i) => | |
| 242 | let val hyps = map discard_proof (Logic.strip_assums_hyp prem) | |
| 243 | and concl = discard_proof (Logic.strip_assums_concl prem) | |
| 244 | in | |
| 245 | if exists (fn hyp => hyp aconv concl) hyps | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
27152diff
changeset | 246 | then case distinct (op =) (filter (fn hyp => Term.could_unify (hyp, concl)) hyps) of | 
| 26322 | 247 | [_] => assume_tac i | 
| 248 | | _ => no_tac | |
| 249 | else no_tac | |
| 250 | end); | |
| 251 | end; | |
| 252 | *} | |
| 253 | ||
| 254 | ||
| 255 | (*** Modus Ponens Tactics ***) | |
| 256 | ||
| 257 | (*Finds P-->Q and P in the assumptions, replaces implication by Q *) | |
| 258 | ML {*
 | |
| 259 |   fun mp_tac i = eresolve_tac [@{thm notE}, make_elim @{thm mp}] i  THEN  assume_tac i
 | |
| 260 | *} | |
| 261 | ||
| 262 | (*Like mp_tac but instantiates no variables*) | |
| 263 | ML {*
 | |
| 264 |   fun int_uniq_mp_tac i = eresolve_tac [@{thm notE}, @{thm impE}] i  THEN  uniq_assume_tac i
 | |
| 265 | *} | |
| 266 | ||
| 267 | ||
| 268 | (*** If-and-only-if ***) | |
| 269 | ||
| 36319 | 270 | schematic_lemma iffI: | 
| 26322 | 271 | assumes "!!x. x:P ==> q(x):Q" | 
| 272 | and "!!x. x:Q ==> r(x):P" | |
| 273 | shows "?p:P<->Q" | |
| 274 | unfolding iff_def | |
| 275 | apply (assumption | rule assms conjI impI)+ | |
| 276 | done | |
| 277 | ||
| 278 | ||
| 279 | (*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *) | |
| 280 | ||
| 36319 | 281 | schematic_lemma iffE: | 
| 26322 | 282 | assumes "p:P <-> Q" | 
| 283 | and "!!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R" | |
| 284 | shows "?p:R" | |
| 285 | apply (rule conjE) | |
| 286 | apply (rule assms(1) [unfolded iff_def]) | |
| 287 | apply (rule assms(2)) | |
| 288 | apply assumption+ | |
| 289 | done | |
| 290 | ||
| 291 | (* Destruct rules for <-> similar to Modus Ponens *) | |
| 292 | ||
| 36319 | 293 | schematic_lemma iffD1: "[| p:P <-> Q; q:P |] ==> ?p:Q" | 
| 26322 | 294 | unfolding iff_def | 
| 295 | apply (rule conjunct1 [THEN mp], assumption+) | |
| 296 | done | |
| 297 | ||
| 36319 | 298 | schematic_lemma iffD2: "[| p:P <-> Q; q:Q |] ==> ?p:P" | 
| 26322 | 299 | unfolding iff_def | 
| 300 | apply (rule conjunct2 [THEN mp], assumption+) | |
| 301 | done | |
| 302 | ||
| 36319 | 303 | schematic_lemma iff_refl: "?p:P <-> P" | 
| 26322 | 304 | apply (rule iffI) | 
| 305 | apply assumption+ | |
| 306 | done | |
| 307 | ||
| 36319 | 308 | schematic_lemma iff_sym: "p:Q <-> P ==> ?p:P <-> Q" | 
| 26322 | 309 | apply (erule iffE) | 
| 310 | apply (rule iffI) | |
| 311 | apply (erule (1) mp)+ | |
| 312 | done | |
| 313 | ||
| 36319 | 314 | schematic_lemma iff_trans: "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R" | 
| 26322 | 315 | apply (rule iffI) | 
| 316 | apply (assumption | erule iffE | erule (1) impE)+ | |
| 317 | done | |
| 318 | ||
| 319 | (*** Unique existence. NOTE THAT the following 2 quantifications | |
| 320 | EX!x such that [EX!y such that P(x,y)] (sequential) | |
| 321 | EX!x,y such that P(x,y) (simultaneous) | |
| 322 | do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential. | |
| 323 | ***) | |
| 324 | ||
| 36319 | 325 | schematic_lemma ex1I: | 
| 26322 | 326 | assumes "p:P(a)" | 
| 327 | and "!!x u. u:P(x) ==> f(u) : x=a" | |
| 328 | shows "?p:EX! x. P(x)" | |
| 329 | unfolding ex1_def | |
| 330 | apply (assumption | rule assms exI conjI allI impI)+ | |
| 331 | done | |
| 332 | ||
| 36319 | 333 | schematic_lemma ex1E: | 
| 26322 | 334 | assumes "p:EX! x. P(x)" | 
| 335 | and "!!x u v. [| u:P(x); v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R" | |
| 336 | shows "?a : R" | |
| 337 | apply (insert assms(1) [unfolded ex1_def]) | |
| 338 | apply (erule exE conjE | assumption | rule assms(1))+ | |
| 29305 | 339 | apply (erule assms(2), assumption) | 
| 26322 | 340 | done | 
| 341 | ||
| 342 | ||
| 343 | (*** <-> congruence rules for simplification ***) | |
| 344 | ||
| 345 | (*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*) | |
| 346 | ML {*
 | |
| 347 | fun iff_tac prems i = | |
| 348 |     resolve_tac (prems RL [@{thm iffE}]) i THEN
 | |
| 349 |     REPEAT1 (eresolve_tac [asm_rl, @{thm mp}] i)
 | |
| 350 | *} | |
| 351 | ||
| 36319 | 352 | schematic_lemma conj_cong: | 
| 26322 | 353 | assumes "p:P <-> P'" | 
| 354 | and "!!x. x:P' ==> q(x):Q <-> Q'" | |
| 355 | shows "?p:(P&Q) <-> (P'&Q')" | |
| 356 | apply (insert assms(1)) | |
| 357 | apply (assumption | rule iffI conjI | | |
| 358 |     erule iffE conjE mp | tactic {* iff_tac @{thms assms} 1 *})+
 | |
| 359 | done | |
| 360 | ||
| 36319 | 361 | schematic_lemma disj_cong: | 
| 26322 | 362 | "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')" | 
| 363 |   apply (erule iffE disjE disjI1 disjI2 | assumption | rule iffI | tactic {* mp_tac 1 *})+
 | |
| 364 | done | |
| 365 | ||
| 36319 | 366 | schematic_lemma imp_cong: | 
| 26322 | 367 | assumes "p:P <-> P'" | 
| 368 | and "!!x. x:P' ==> q(x):Q <-> Q'" | |
| 369 | shows "?p:(P-->Q) <-> (P'-->Q')" | |
| 370 | apply (insert assms(1)) | |
| 371 |   apply (assumption | rule iffI impI | erule iffE | tactic {* mp_tac 1 *} |
 | |
| 372 |     tactic {* iff_tac @{thms assms} 1 *})+
 | |
| 373 | done | |
| 374 | ||
| 36319 | 375 | schematic_lemma iff_cong: | 
| 26322 | 376 | "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')" | 
| 377 |   apply (erule iffE | assumption | rule iffI | tactic {* mp_tac 1 *})+
 | |
| 378 | done | |
| 379 | ||
| 36319 | 380 | schematic_lemma not_cong: | 
| 26322 | 381 | "p:P <-> P' ==> ?p:~P <-> ~P'" | 
| 382 |   apply (assumption | rule iffI notI | tactic {* mp_tac 1 *} | erule iffE notE)+
 | |
| 383 | done | |
| 384 | ||
| 36319 | 385 | schematic_lemma all_cong: | 
| 26322 | 386 | assumes "!!x. f(x):P(x) <-> Q(x)" | 
| 387 | shows "?p:(ALL x. P(x)) <-> (ALL x. Q(x))" | |
| 388 |   apply (assumption | rule iffI allI | tactic {* mp_tac 1 *} | erule allE |
 | |
| 389 |     tactic {* iff_tac @{thms assms} 1 *})+
 | |
| 390 | done | |
| 391 | ||
| 36319 | 392 | schematic_lemma ex_cong: | 
| 26322 | 393 | assumes "!!x. f(x):P(x) <-> Q(x)" | 
| 394 | shows "?p:(EX x. P(x)) <-> (EX x. Q(x))" | |
| 395 |   apply (erule exE | assumption | rule iffI exI | tactic {* mp_tac 1 *} |
 | |
| 396 |     tactic {* iff_tac @{thms assms} 1 *})+
 | |
| 397 | done | |
| 398 | ||
| 399 | (*NOT PROVED | |
| 400 | bind_thm ("ex1_cong", prove_goal (the_context ())
 | |
| 401 | "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX! x.P(x)) <-> (EX! x.Q(x))" | |
| 402 | (fn prems => | |
| 403 | [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1 | |
| 404 | ORELSE mp_tac 1 | |
| 405 | ORELSE iff_tac prems 1)) ])) | |
| 406 | *) | |
| 407 | ||
| 408 | (*** Equality rules ***) | |
| 409 | ||
| 410 | lemmas refl = ieqI | |
| 411 | ||
| 36319 | 412 | schematic_lemma subst: | 
| 26322 | 413 | assumes prem1: "p:a=b" | 
| 414 | and prem2: "q:P(a)" | |
| 415 | shows "?p : P(b)" | |
| 416 | apply (rule prem2 [THEN rev_mp]) | |
| 417 | apply (rule prem1 [THEN ieqE]) | |
| 418 | apply (rule impI) | |
| 419 | apply assumption | |
| 420 | done | |
| 421 | ||
| 36319 | 422 | schematic_lemma sym: "q:a=b ==> ?c:b=a" | 
| 26322 | 423 | apply (erule subst) | 
| 424 | apply (rule refl) | |
| 425 | done | |
| 426 | ||
| 36319 | 427 | schematic_lemma trans: "[| p:a=b; q:b=c |] ==> ?d:a=c" | 
| 26322 | 428 | apply (erule (1) subst) | 
| 429 | done | |
| 430 | ||
| 431 | (** ~ b=a ==> ~ a=b **) | |
| 36319 | 432 | schematic_lemma not_sym: "p:~ b=a ==> ?q:~ a=b" | 
| 26322 | 433 | apply (erule contrapos) | 
| 434 | apply (erule sym) | |
| 435 | done | |
| 436 | ||
| 437 | (*calling "standard" reduces maxidx to 0*) | |
| 438 | lemmas ssubst = sym [THEN subst, standard] | |
| 439 | ||
| 440 | (*A special case of ex1E that would otherwise need quantifier expansion*) | |
| 36319 | 441 | schematic_lemma ex1_equalsE: "[| p:EX! x. P(x); q:P(a); r:P(b) |] ==> ?d:a=b" | 
| 26322 | 442 | apply (erule ex1E) | 
| 443 | apply (rule trans) | |
| 444 | apply (rule_tac [2] sym) | |
| 445 | apply (assumption | erule spec [THEN mp])+ | |
| 446 | done | |
| 447 | ||
| 448 | (** Polymorphic congruence rules **) | |
| 449 | ||
| 36319 | 450 | schematic_lemma subst_context: "[| p:a=b |] ==> ?d:t(a)=t(b)" | 
| 26322 | 451 | apply (erule ssubst) | 
| 452 | apply (rule refl) | |
| 453 | done | |
| 454 | ||
| 36319 | 455 | schematic_lemma subst_context2: "[| p:a=b; q:c=d |] ==> ?p:t(a,c)=t(b,d)" | 
| 26322 | 456 | apply (erule ssubst)+ | 
| 457 | apply (rule refl) | |
| 458 | done | |
| 459 | ||
| 36319 | 460 | schematic_lemma subst_context3: "[| p:a=b; q:c=d; r:e=f |] ==> ?p:t(a,c,e)=t(b,d,f)" | 
| 26322 | 461 | apply (erule ssubst)+ | 
| 462 | apply (rule refl) | |
| 463 | done | |
| 464 | ||
| 465 | (*Useful with eresolve_tac for proving equalties from known equalities. | |
| 466 | a = b | |
| 467 | | | | |
| 468 | c = d *) | |
| 36319 | 469 | schematic_lemma box_equals: "[| p:a=b; q:a=c; r:b=d |] ==> ?p:c=d" | 
| 26322 | 470 | apply (rule trans) | 
| 471 | apply (rule trans) | |
| 472 | apply (rule sym) | |
| 473 | apply assumption+ | |
| 474 | done | |
| 475 | ||
| 476 | (*Dual of box_equals: for proving equalities backwards*) | |
| 36319 | 477 | schematic_lemma simp_equals: "[| p:a=c; q:b=d; r:c=d |] ==> ?p:a=b" | 
| 26322 | 478 | apply (rule trans) | 
| 479 | apply (rule trans) | |
| 480 | apply (assumption | rule sym)+ | |
| 481 | done | |
| 482 | ||
| 483 | (** Congruence rules for predicate letters **) | |
| 484 | ||
| 36319 | 485 | schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')" | 
| 26322 | 486 | apply (rule iffI) | 
| 487 |    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
 | |
| 488 | done | |
| 489 | ||
| 36319 | 490 | schematic_lemma pred2_cong: "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')" | 
| 26322 | 491 | apply (rule iffI) | 
| 492 |    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
 | |
| 493 | done | |
| 494 | ||
| 36319 | 495 | schematic_lemma pred3_cong: "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')" | 
| 26322 | 496 | apply (rule iffI) | 
| 497 |    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
 | |
| 498 | done | |
| 499 | ||
| 27152 
192954a9a549
changed pred_congs: merely cover pred1_cong pred2_cong pred3_cong;
 wenzelm parents: 
27150diff
changeset | 500 | lemmas pred_congs = pred1_cong pred2_cong pred3_cong | 
| 26322 | 501 | |
| 502 | (*special case for the equality predicate!*) | |
| 503 | lemmas eq_cong = pred2_cong [where P = "op =", standard] | |
| 504 | ||
| 505 | ||
| 506 | (*** Simplifications of assumed implications. | |
| 507 | Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE | |
| 508 | used with mp_tac (restricted to atomic formulae) is COMPLETE for | |
| 509 | intuitionistic propositional logic. See | |
| 510 | R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic | |
| 511 | (preprint, University of St Andrews, 1991) ***) | |
| 512 | ||
| 36319 | 513 | schematic_lemma conj_impE: | 
| 26322 | 514 | assumes major: "p:(P&Q)-->S" | 
| 515 | and minor: "!!x. x:P-->(Q-->S) ==> q(x):R" | |
| 516 | shows "?p:R" | |
| 517 | apply (assumption | rule conjI impI major [THEN mp] minor)+ | |
| 518 | done | |
| 519 | ||
| 36319 | 520 | schematic_lemma disj_impE: | 
| 26322 | 521 | assumes major: "p:(P|Q)-->S" | 
| 522 | and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R" | |
| 523 | shows "?p:R" | |
| 524 |   apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
 | |
| 525 |       resolve_tac [@{thm disjI1}, @{thm disjI2}, @{thm impI},
 | |
| 526 |         @{thm major} RS @{thm mp}, @{thm minor}] 1) *})
 | |
| 527 | done | |
| 528 | ||
| 529 | (*Simplifies the implication. Classical version is stronger. | |
| 530 | Still UNSAFE since Q must be provable -- backtracking needed. *) | |
| 36319 | 531 | schematic_lemma imp_impE: | 
| 26322 | 532 | assumes major: "p:(P-->Q)-->S" | 
| 533 | and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q" | |
| 534 | and r2: "!!x. x:S ==> r(x):R" | |
| 535 | shows "?p:R" | |
| 536 | apply (assumption | rule impI major [THEN mp] r1 r2)+ | |
| 537 | done | |
| 538 | ||
| 539 | (*Simplifies the implication. Classical version is stronger. | |
| 540 | Still UNSAFE since ~P must be provable -- backtracking needed. *) | |
| 36319 | 541 | schematic_lemma not_impE: | 
| 26322 | 542 | assumes major: "p:~P --> S" | 
| 543 | and r1: "!!y. y:P ==> q(y):False" | |
| 544 | and r2: "!!y. y:S ==> r(y):R" | |
| 545 | shows "?p:R" | |
| 546 | apply (assumption | rule notI impI major [THEN mp] r1 r2)+ | |
| 547 | done | |
| 548 | ||
| 549 | (*Simplifies the implication. UNSAFE. *) | |
| 36319 | 550 | schematic_lemma iff_impE: | 
| 26322 | 551 | assumes major: "p:(P<->Q)-->S" | 
| 552 | and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q" | |
| 553 | and r2: "!!x y.[| x:Q; y:P-->S |] ==> r(x,y):P" | |
| 554 | and r3: "!!x. x:S ==> s(x):R" | |
| 555 | shows "?p:R" | |
| 556 | apply (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+ | |
| 557 | done | |
| 558 | ||
| 559 | (*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*) | |
| 36319 | 560 | schematic_lemma all_impE: | 
| 26322 | 561 | assumes major: "p:(ALL x. P(x))-->S" | 
| 562 | and r1: "!!x. q:P(x)" | |
| 563 | and r2: "!!y. y:S ==> r(y):R" | |
| 564 | shows "?p:R" | |
| 565 | apply (assumption | rule allI impI major [THEN mp] r1 r2)+ | |
| 566 | done | |
| 567 | ||
| 568 | (*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *) | |
| 36319 | 569 | schematic_lemma ex_impE: | 
| 26322 | 570 | assumes major: "p:(EX x. P(x))-->S" | 
| 571 | and r: "!!y. y:P(a)-->S ==> q(y):R" | |
| 572 | shows "?p:R" | |
| 573 | apply (assumption | rule exI impI major [THEN mp] r)+ | |
| 574 | done | |
| 575 | ||
| 576 | ||
| 36319 | 577 | schematic_lemma rev_cut_eq: | 
| 26322 | 578 | assumes "p:a=b" | 
| 579 | and "!!x. x:a=b ==> f(x):R" | |
| 580 | shows "?p:R" | |
| 581 | apply (rule assms)+ | |
| 582 | done | |
| 583 | ||
| 584 | lemma thin_refl: "!!X. [|p:x=x; PROP W|] ==> PROP W" . | |
| 585 | ||
| 586 | use "hypsubst.ML" | |
| 587 | ||
| 588 | ML {*
 | |
| 589 | ||
| 590 | (*** Applying HypsubstFun to generate hyp_subst_tac ***) | |
| 591 | ||
| 592 | structure Hypsubst_Data = | |
| 593 | struct | |
| 594 | (*Take apart an equality judgement; otherwise raise Match!*) | |
| 595 |   fun dest_eq (Const (@{const_name Proof}, _) $
 | |
| 596 |     (Const (@{const_name "op ="}, _)  $ t $ u) $ _) = (t, u);
 | |
| 597 | ||
| 598 |   val imp_intr = @{thm impI}
 | |
| 599 | ||
| 600 | (*etac rev_cut_eq moves an equality to be the last premise. *) | |
| 601 |   val rev_cut_eq = @{thm rev_cut_eq}
 | |
| 602 | ||
| 603 |   val rev_mp = @{thm rev_mp}
 | |
| 604 |   val subst = @{thm subst}
 | |
| 605 |   val sym = @{thm sym}
 | |
| 606 |   val thin_refl = @{thm thin_refl}
 | |
| 607 | end; | |
| 608 | ||
| 609 | structure Hypsubst = HypsubstFun(Hypsubst_Data); | |
| 610 | open Hypsubst; | |
| 611 | *} | |
| 612 | ||
| 613 | use "intprover.ML" | |
| 614 | ||
| 615 | ||
| 616 | (*** Rewrite rules ***) | |
| 617 | ||
| 36319 | 618 | schematic_lemma conj_rews: | 
| 26322 | 619 | "?p1 : P & True <-> P" | 
| 620 | "?p2 : True & P <-> P" | |
| 621 | "?p3 : P & False <-> False" | |
| 622 | "?p4 : False & P <-> False" | |
| 623 | "?p5 : P & P <-> P" | |
| 624 | "?p6 : P & ~P <-> False" | |
| 625 | "?p7 : ~P & P <-> False" | |
| 626 | "?p8 : (P & Q) & R <-> P & (Q & R)" | |
| 627 |   apply (tactic {* fn st => IntPr.fast_tac 1 st *})+
 | |
| 628 | done | |
| 629 | ||
| 36319 | 630 | schematic_lemma disj_rews: | 
| 26322 | 631 | "?p1 : P | True <-> True" | 
| 632 | "?p2 : True | P <-> True" | |
| 633 | "?p3 : P | False <-> P" | |
| 634 | "?p4 : False | P <-> P" | |
| 635 | "?p5 : P | P <-> P" | |
| 636 | "?p6 : (P | Q) | R <-> P | (Q | R)" | |
| 637 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 638 | done | |
| 639 | ||
| 36319 | 640 | schematic_lemma not_rews: | 
| 26322 | 641 | "?p1 : ~ False <-> True" | 
| 642 | "?p2 : ~ True <-> False" | |
| 643 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 644 | done | |
| 645 | ||
| 36319 | 646 | schematic_lemma imp_rews: | 
| 26322 | 647 | "?p1 : (P --> False) <-> ~P" | 
| 648 | "?p2 : (P --> True) <-> True" | |
| 649 | "?p3 : (False --> P) <-> True" | |
| 650 | "?p4 : (True --> P) <-> P" | |
| 651 | "?p5 : (P --> P) <-> True" | |
| 652 | "?p6 : (P --> ~P) <-> ~P" | |
| 653 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 654 | done | |
| 655 | ||
| 36319 | 656 | schematic_lemma iff_rews: | 
| 26322 | 657 | "?p1 : (True <-> P) <-> P" | 
| 658 | "?p2 : (P <-> True) <-> P" | |
| 659 | "?p3 : (P <-> P) <-> True" | |
| 660 | "?p4 : (False <-> P) <-> ~P" | |
| 661 | "?p5 : (P <-> False) <-> ~P" | |
| 662 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 663 | done | |
| 664 | ||
| 36319 | 665 | schematic_lemma quant_rews: | 
| 26322 | 666 | "?p1 : (ALL x. P) <-> P" | 
| 667 | "?p2 : (EX x. P) <-> P" | |
| 668 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 669 | done | |
| 670 | ||
| 671 | (*These are NOT supplied by default!*) | |
| 36319 | 672 | schematic_lemma distrib_rews1: | 
| 26322 | 673 | "?p1 : ~(P|Q) <-> ~P & ~Q" | 
| 674 | "?p2 : P & (Q | R) <-> P&Q | P&R" | |
| 675 | "?p3 : (Q | R) & P <-> Q&P | R&P" | |
| 676 | "?p4 : (P | Q --> R) <-> (P --> R) & (Q --> R)" | |
| 677 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 678 | done | |
| 679 | ||
| 36319 | 680 | schematic_lemma distrib_rews2: | 
| 26322 | 681 | "?p1 : ~(EX x. NORM(P(x))) <-> (ALL x. ~NORM(P(x)))" | 
| 682 | "?p2 : ((EX x. NORM(P(x))) --> Q) <-> (ALL x. NORM(P(x)) --> Q)" | |
| 683 | "?p3 : (EX x. NORM(P(x))) & NORM(Q) <-> (EX x. NORM(P(x)) & NORM(Q))" | |
| 684 | "?p4 : NORM(Q) & (EX x. NORM(P(x))) <-> (EX x. NORM(Q) & NORM(P(x)))" | |
| 685 |   apply (tactic {* IntPr.fast_tac 1 *})+
 | |
| 686 | done | |
| 687 | ||
| 688 | lemmas distrib_rews = distrib_rews1 distrib_rews2 | |
| 689 | ||
| 36319 | 690 | schematic_lemma P_Imp_P_iff_T: "p:P ==> ?p:(P <-> True)" | 
| 26322 | 691 |   apply (tactic {* IntPr.fast_tac 1 *})
 | 
| 692 | done | |
| 693 | ||
| 36319 | 694 | schematic_lemma not_P_imp_P_iff_F: "p:~P ==> ?p:(P <-> False)" | 
| 26322 | 695 |   apply (tactic {* IntPr.fast_tac 1 *})
 | 
| 696 | done | |
| 0 | 697 | |
| 698 | end |