| author | berghofe | 
| Wed, 21 May 2008 14:04:41 +0200 | |
| changeset 26964 | df1f238a05f7 | 
| parent 26496 | 49ae9456eba9 | 
| child 27115 | 0dcafa5c9e3f | 
| permissions | -rw-r--r-- | 
| 9487 | 1 | (* Title: FOL/FOL.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson and Markus Wenzel | |
| 11678 | 4 | *) | 
| 9487 | 5 | |
| 11678 | 6 | header {* Classical first-order logic *}
 | 
| 4093 | 7 | |
| 18456 | 8 | theory FOL | 
| 15481 | 9 | imports IFOL | 
| 23154 | 10 | uses | 
| 24097 | 11 | "~~/src/Provers/classical.ML" | 
| 12 | "~~/src/Provers/blast.ML" | |
| 13 | "~~/src/Provers/clasimp.ML" | |
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 14 | "~~/src/Tools/induct.ML" | 
| 23154 | 15 |   ("cladata.ML")
 | 
| 16 |   ("blastdata.ML")
 | |
| 17 |   ("simpdata.ML")
 | |
| 18456 | 18 | begin | 
| 9487 | 19 | |
| 20 | ||
| 21 | subsection {* The classical axiom *}
 | |
| 4093 | 22 | |
| 7355 
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
 wenzelm parents: 
5887diff
changeset | 23 | axioms | 
| 
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
 wenzelm parents: 
5887diff
changeset | 24 | classical: "(~P ==> P) ==> P" | 
| 4093 | 25 | |
| 9487 | 26 | |
| 11678 | 27 | subsection {* Lemmas and proof tools *}
 | 
| 9487 | 28 | |
| 21539 | 29 | lemma ccontr: "(\<not> P \<Longrightarrow> False) \<Longrightarrow> P" | 
| 30 | by (erule FalseE [THEN classical]) | |
| 31 | ||
| 32 | (*** Classical introduction rules for | and EX ***) | |
| 33 | ||
| 34 | lemma disjCI: "(~Q ==> P) ==> P|Q" | |
| 35 | apply (rule classical) | |
| 36 | apply (assumption | erule meta_mp | rule disjI1 notI)+ | |
| 37 | apply (erule notE disjI2)+ | |
| 38 | done | |
| 39 | ||
| 40 | (*introduction rule involving only EX*) | |
| 41 | lemma ex_classical: | |
| 42 | assumes r: "~(EX x. P(x)) ==> P(a)" | |
| 43 | shows "EX x. P(x)" | |
| 44 | apply (rule classical) | |
| 45 | apply (rule exI, erule r) | |
| 46 | done | |
| 47 | ||
| 48 | (*version of above, simplifying ~EX to ALL~ *) | |
| 49 | lemma exCI: | |
| 50 | assumes r: "ALL x. ~P(x) ==> P(a)" | |
| 51 | shows "EX x. P(x)" | |
| 52 | apply (rule ex_classical) | |
| 53 | apply (rule notI [THEN allI, THEN r]) | |
| 54 | apply (erule notE) | |
| 55 | apply (erule exI) | |
| 56 | done | |
| 57 | ||
| 58 | lemma excluded_middle: "~P | P" | |
| 59 | apply (rule disjCI) | |
| 60 | apply assumption | |
| 61 | done | |
| 62 | ||
| 63 | (*For disjunctive case analysis*) | |
| 64 | ML {*
 | |
| 22139 | 65 | fun excluded_middle_tac sP = | 
| 66 |     res_inst_tac [("Q",sP)] (@{thm excluded_middle} RS @{thm disjE})
 | |
| 21539 | 67 | *} | 
| 68 | ||
| 69 | lemma case_split_thm: | |
| 70 | assumes r1: "P ==> Q" | |
| 71 | and r2: "~P ==> Q" | |
| 72 | shows Q | |
| 73 | apply (rule excluded_middle [THEN disjE]) | |
| 74 | apply (erule r2) | |
| 75 | apply (erule r1) | |
| 76 | done | |
| 77 | ||
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 78 | lemmas case_split = case_split_thm [case_names True False] | 
| 21539 | 79 | |
| 80 | (*HOL's more natural case analysis tactic*) | |
| 81 | ML {*
 | |
| 22139 | 82 |   fun case_tac a = res_inst_tac [("P",a)] @{thm case_split_thm}
 | 
| 21539 | 83 | *} | 
| 84 | ||
| 85 | ||
| 86 | (*** Special elimination rules *) | |
| 87 | ||
| 88 | ||
| 89 | (*Classical implies (-->) elimination. *) | |
| 90 | lemma impCE: | |
| 91 | assumes major: "P-->Q" | |
| 92 | and r1: "~P ==> R" | |
| 93 | and r2: "Q ==> R" | |
| 94 | shows R | |
| 95 | apply (rule excluded_middle [THEN disjE]) | |
| 96 | apply (erule r1) | |
| 97 | apply (rule r2) | |
| 98 | apply (erule major [THEN mp]) | |
| 99 | done | |
| 100 | ||
| 101 | (*This version of --> elimination works on Q before P. It works best for | |
| 102 | those cases in which P holds "almost everywhere". Can't install as | |
| 103 | default: would break old proofs.*) | |
| 104 | lemma impCE': | |
| 105 | assumes major: "P-->Q" | |
| 106 | and r1: "Q ==> R" | |
| 107 | and r2: "~P ==> R" | |
| 108 | shows R | |
| 109 | apply (rule excluded_middle [THEN disjE]) | |
| 110 | apply (erule r2) | |
| 111 | apply (rule r1) | |
| 112 | apply (erule major [THEN mp]) | |
| 113 | done | |
| 114 | ||
| 115 | (*Double negation law*) | |
| 116 | lemma notnotD: "~~P ==> P" | |
| 117 | apply (rule classical) | |
| 118 | apply (erule notE) | |
| 119 | apply assumption | |
| 120 | done | |
| 121 | ||
| 122 | lemma contrapos2: "[| Q; ~ P ==> ~ Q |] ==> P" | |
| 123 | apply (rule classical) | |
| 124 | apply (drule (1) meta_mp) | |
| 125 | apply (erule (1) notE) | |
| 126 | done | |
| 127 | ||
| 128 | (*** Tactics for implication and contradiction ***) | |
| 129 | ||
| 130 | (*Classical <-> elimination. Proof substitutes P=Q in | |
| 131 | ~P ==> ~Q and P ==> Q *) | |
| 132 | lemma iffCE: | |
| 133 | assumes major: "P<->Q" | |
| 134 | and r1: "[| P; Q |] ==> R" | |
| 135 | and r2: "[| ~P; ~Q |] ==> R" | |
| 136 | shows R | |
| 137 | apply (rule major [unfolded iff_def, THEN conjE]) | |
| 138 | apply (elim impCE) | |
| 139 | apply (erule (1) r2) | |
| 140 | apply (erule (1) notE)+ | |
| 141 | apply (erule (1) r1) | |
| 142 | done | |
| 143 | ||
| 144 | ||
| 145 | (*Better for fast_tac: needs no quantifier duplication!*) | |
| 146 | lemma alt_ex1E: | |
| 147 | assumes major: "EX! x. P(x)" | |
| 148 | and r: "!!x. [| P(x); ALL y y'. P(y) & P(y') --> y=y' |] ==> R" | |
| 149 | shows R | |
| 150 | using major | |
| 151 | proof (rule ex1E) | |
| 152 | fix x | |
| 153 | assume * : "\<forall>y. P(y) \<longrightarrow> y = x" | |
| 154 | assume "P(x)" | |
| 155 | then show R | |
| 156 | proof (rule r) | |
| 157 |     { fix y y'
 | |
| 158 | assume "P(y)" and "P(y')" | |
| 159 | with * have "x = y" and "x = y'" by - (tactic "IntPr.fast_tac 1")+ | |
| 160 | then have "y = y'" by (rule subst) | |
| 161 | } note r' = this | |
| 162 | show "\<forall>y y'. P(y) \<and> P(y') \<longrightarrow> y = y'" by (intro strip, elim conjE) (rule r') | |
| 163 | qed | |
| 164 | qed | |
| 9525 | 165 | |
| 26411 | 166 | lemma imp_elim: "P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R" | 
| 167 | by (rule classical) iprover | |
| 168 | ||
| 169 | lemma swap: "~ P ==> (~ R ==> P) ==> R" | |
| 170 | by (rule classical) iprover | |
| 171 | ||
| 10383 | 172 | use "cladata.ML" | 
| 173 | setup Cla.setup | |
| 14156 | 174 | setup cla_setup | 
| 175 | setup case_setup | |
| 10383 | 176 | |
| 9487 | 177 | use "blastdata.ML" | 
| 178 | setup Blast.setup | |
| 13550 | 179 | |
| 180 | ||
| 181 | lemma ex1_functional: "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" | |
| 21539 | 182 | by blast | 
| 20223 | 183 | |
| 184 | (* Elimination of True from asumptions: *) | |
| 185 | lemma True_implies_equals: "(True ==> PROP P) == PROP P" | |
| 186 | proof | |
| 187 | assume "True \<Longrightarrow> PROP P" | |
| 188 | from this and TrueI show "PROP P" . | |
| 189 | next | |
| 190 | assume "PROP P" | |
| 191 | then show "PROP P" . | |
| 192 | qed | |
| 9487 | 193 | |
| 21539 | 194 | lemma uncurry: "P --> Q --> R ==> P & Q --> R" | 
| 195 | by blast | |
| 196 | ||
| 197 | lemma iff_allI: "(!!x. P(x) <-> Q(x)) ==> (ALL x. P(x)) <-> (ALL x. Q(x))" | |
| 198 | by blast | |
| 199 | ||
| 200 | lemma iff_exI: "(!!x. P(x) <-> Q(x)) ==> (EX x. P(x)) <-> (EX x. Q(x))" | |
| 201 | by blast | |
| 202 | ||
| 203 | lemma all_comm: "(ALL x y. P(x,y)) <-> (ALL y x. P(x,y))" by blast | |
| 204 | ||
| 205 | lemma ex_comm: "(EX x y. P(x,y)) <-> (EX y x. P(x,y))" by blast | |
| 206 | ||
| 26286 | 207 | |
| 208 | ||
| 209 | (*** Classical simplification rules ***) | |
| 210 | ||
| 211 | (*Avoids duplication of subgoals after expand_if, when the true and false | |
| 212 | cases boil down to the same thing.*) | |
| 213 | lemma cases_simp: "(P --> Q) & (~P --> Q) <-> Q" by blast | |
| 214 | ||
| 215 | ||
| 216 | (*** Miniscoping: pushing quantifiers in | |
| 217 | We do NOT distribute of ALL over &, or dually that of EX over | | |
| 218 | Baaz and Leitsch, On Skolemization and Proof Complexity (1994) | |
| 219 | show that this step can increase proof length! | |
| 220 | ***) | |
| 221 | ||
| 222 | (*existential miniscoping*) | |
| 223 | lemma int_ex_simps: | |
| 224 | "!!P Q. (EX x. P(x) & Q) <-> (EX x. P(x)) & Q" | |
| 225 | "!!P Q. (EX x. P & Q(x)) <-> P & (EX x. Q(x))" | |
| 226 | "!!P Q. (EX x. P(x) | Q) <-> (EX x. P(x)) | Q" | |
| 227 | "!!P Q. (EX x. P | Q(x)) <-> P | (EX x. Q(x))" | |
| 228 | by iprover+ | |
| 229 | ||
| 230 | (*classical rules*) | |
| 231 | lemma cla_ex_simps: | |
| 232 | "!!P Q. (EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q" | |
| 233 | "!!P Q. (EX x. P --> Q(x)) <-> P --> (EX x. Q(x))" | |
| 234 | by blast+ | |
| 235 | ||
| 236 | lemmas ex_simps = int_ex_simps cla_ex_simps | |
| 237 | ||
| 238 | (*universal miniscoping*) | |
| 239 | lemma int_all_simps: | |
| 240 | "!!P Q. (ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q" | |
| 241 | "!!P Q. (ALL x. P & Q(x)) <-> P & (ALL x. Q(x))" | |
| 242 | "!!P Q. (ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q" | |
| 243 | "!!P Q. (ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))" | |
| 244 | by iprover+ | |
| 245 | ||
| 246 | (*classical rules*) | |
| 247 | lemma cla_all_simps: | |
| 248 | "!!P Q. (ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q" | |
| 249 | "!!P Q. (ALL x. P | Q(x)) <-> P | (ALL x. Q(x))" | |
| 250 | by blast+ | |
| 251 | ||
| 252 | lemmas all_simps = int_all_simps cla_all_simps | |
| 253 | ||
| 254 | ||
| 255 | (*** Named rewrite rules proved for IFOL ***) | |
| 256 | ||
| 257 | lemma imp_disj1: "(P-->Q) | R <-> (P-->Q | R)" by blast | |
| 258 | lemma imp_disj2: "Q | (P-->R) <-> (P-->Q | R)" by blast | |
| 259 | ||
| 260 | lemma de_Morgan_conj: "(~(P & Q)) <-> (~P | ~Q)" by blast | |
| 261 | ||
| 262 | lemma not_imp: "~(P --> Q) <-> (P & ~Q)" by blast | |
| 263 | lemma not_iff: "~(P <-> Q) <-> (P <-> ~Q)" by blast | |
| 264 | ||
| 265 | lemma not_all: "(~ (ALL x. P(x))) <-> (EX x.~P(x))" by blast | |
| 266 | lemma imp_all: "((ALL x. P(x)) --> Q) <-> (EX x. P(x) --> Q)" by blast | |
| 267 | ||
| 268 | ||
| 269 | lemmas meta_simps = | |
| 270 | triv_forall_equality (* prunes params *) | |
| 271 | True_implies_equals (* prune asms `True' *) | |
| 272 | ||
| 273 | lemmas IFOL_simps = | |
| 274 | refl [THEN P_iff_T] conj_simps disj_simps not_simps | |
| 275 | imp_simps iff_simps quant_simps | |
| 276 | ||
| 277 | lemma notFalseI: "~False" by iprover | |
| 278 | ||
| 279 | lemma cla_simps_misc: | |
| 280 | "~(P&Q) <-> ~P | ~Q" | |
| 281 | "P | ~P" | |
| 282 | "~P | P" | |
| 283 | "~ ~ P <-> P" | |
| 284 | "(~P --> P) <-> P" | |
| 285 | "(~P <-> ~Q) <-> (P<->Q)" by blast+ | |
| 286 | ||
| 287 | lemmas cla_simps = | |
| 288 | de_Morgan_conj de_Morgan_disj imp_disj1 imp_disj2 | |
| 289 | not_imp not_all not_ex cases_simp cla_simps_misc | |
| 290 | ||
| 291 | ||
| 9487 | 292 | use "simpdata.ML" | 
| 293 | setup simpsetup | |
| 294 | setup "Simplifier.method_setup Splitter.split_modifiers" | |
| 295 | setup Splitter.setup | |
| 26496 
49ae9456eba9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26411diff
changeset | 296 | setup clasimp_setup | 
| 18591 | 297 | setup EqSubst.setup | 
| 15481 | 298 | |
| 299 | ||
| 14085 | 300 | subsection {* Other simple lemmas *}
 | 
| 301 | ||
| 302 | lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)" | |
| 303 | by blast | |
| 304 | ||
| 305 | lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" | |
| 306 | by blast | |
| 307 | ||
| 308 | lemma not_disj_iff_imp: "~P | Q <-> (P-->Q)" | |
| 309 | by blast | |
| 310 | ||
| 311 | (** Monotonicity of implications **) | |
| 312 | ||
| 313 | lemma conj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)" | |
| 314 | by fast (*or (IntPr.fast_tac 1)*) | |
| 315 | ||
| 316 | lemma disj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)" | |
| 317 | by fast (*or (IntPr.fast_tac 1)*) | |
| 318 | ||
| 319 | lemma imp_mono: "[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)" | |
| 320 | by fast (*or (IntPr.fast_tac 1)*) | |
| 321 | ||
| 322 | lemma imp_refl: "P-->P" | |
| 323 | by (rule impI, assumption) | |
| 324 | ||
| 325 | (*The quantifier monotonicity rules are also intuitionistically valid*) | |
| 326 | lemma ex_mono: "(!!x. P(x) --> Q(x)) ==> (EX x. P(x)) --> (EX x. Q(x))" | |
| 327 | by blast | |
| 328 | ||
| 329 | lemma all_mono: "(!!x. P(x) --> Q(x)) ==> (ALL x. P(x)) --> (ALL x. Q(x))" | |
| 330 | by blast | |
| 331 | ||
| 11678 | 332 | |
| 333 | subsection {* Proof by cases and induction *}
 | |
| 334 | ||
| 335 | text {* Proper handling of non-atomic rule statements. *}
 | |
| 336 | ||
| 337 | constdefs | |
| 18456 | 338 | induct_forall where "induct_forall(P) == \<forall>x. P(x)" | 
| 339 | induct_implies where "induct_implies(A, B) == A \<longrightarrow> B" | |
| 340 | induct_equal where "induct_equal(x, y) == x = y" | |
| 341 | induct_conj where "induct_conj(A, B) == A \<and> B" | |
| 11678 | 342 | |
| 343 | lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(\<lambda>x. P(x)))" | |
| 18816 | 344 | unfolding atomize_all induct_forall_def . | 
| 11678 | 345 | |
| 346 | lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))" | |
| 18816 | 347 | unfolding atomize_imp induct_implies_def . | 
| 11678 | 348 | |
| 349 | lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))" | |
| 18816 | 350 | unfolding atomize_eq induct_equal_def . | 
| 11678 | 351 | |
| 18456 | 352 | lemma induct_conj_eq: | 
| 353 | includes meta_conjunction_syntax | |
| 354 | shows "(A && B) == Trueprop(induct_conj(A, B))" | |
| 18816 | 355 | unfolding atomize_conj induct_conj_def . | 
| 11988 | 356 | |
| 18456 | 357 | lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq | 
| 358 | lemmas induct_rulify [symmetric, standard] = induct_atomize | |
| 359 | lemmas induct_rulify_fallback = | |
| 360 | induct_forall_def induct_implies_def induct_equal_def induct_conj_def | |
| 11678 | 361 | |
| 18456 | 362 | hide const induct_forall induct_implies induct_equal induct_conj | 
| 11678 | 363 | |
| 364 | ||
| 365 | text {* Method setup. *}
 | |
| 366 | ||
| 367 | ML {*
 | |
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 368 | structure Induct = InductFun | 
| 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 369 | ( | 
| 22139 | 370 |     val cases_default = @{thm case_split}
 | 
| 371 |     val atomize = @{thms induct_atomize}
 | |
| 372 |     val rulify = @{thms induct_rulify}
 | |
| 373 |     val rulify_fallback = @{thms induct_rulify_fallback}
 | |
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 374 | ); | 
| 11678 | 375 | *} | 
| 376 | ||
| 24830 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 377 | setup Induct.setup | 
| 
a7b3ab44d993
moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
 wenzelm parents: 
24097diff
changeset | 378 | declare case_split [cases type: o] | 
| 11678 | 379 | |
| 4854 | 380 | end |