| author | wenzelm | 
| Fri, 23 Dec 2005 18:36:26 +0100 | |
| changeset 18511 | beed2bc052a3 | 
| parent 18456 | 8cc35e95450a | 
| child 18522 | 9bdfb6eaf8ab | 
| 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 | 
| 16417 | 10 | uses ("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML")
 | 
| 15481 | 11 |       ("eqrule_FOL_data.ML")
 | 
| 12 |       ("~~/src/Provers/eqsubst.ML")
 | |
| 18456 | 13 | begin | 
| 9487 | 14 | |
| 15 | ||
| 16 | subsection {* The classical axiom *}
 | |
| 4093 | 17 | |
| 7355 
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
 wenzelm parents: 
5887diff
changeset | 18 | axioms | 
| 
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
 wenzelm parents: 
5887diff
changeset | 19 | classical: "(~P ==> P) ==> P" | 
| 4093 | 20 | |
| 9487 | 21 | |
| 11678 | 22 | subsection {* Lemmas and proof tools *}
 | 
| 9487 | 23 | |
| 7355 
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
 wenzelm parents: 
5887diff
changeset | 24 | use "FOL_lemmas1.ML" | 
| 12127 
219e543496a3
theorems case_split = case_split_thm [case_names True False, cases type: o];
 wenzelm parents: 
11988diff
changeset | 25 | theorems case_split = case_split_thm [case_names True False, cases type: o] | 
| 9525 | 26 | |
| 10383 | 27 | use "cladata.ML" | 
| 28 | setup Cla.setup | |
| 14156 | 29 | setup cla_setup | 
| 30 | setup case_setup | |
| 10383 | 31 | |
| 9487 | 32 | use "blastdata.ML" | 
| 33 | setup Blast.setup | |
| 13550 | 34 | |
| 35 | ||
| 36 | lemma ex1_functional: "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" | |
| 37 | by blast | |
| 38 | ||
| 39 | ML {*
 | |
| 40 | val ex1_functional = thm "ex1_functional"; | |
| 41 | *} | |
| 9487 | 42 | |
| 43 | use "simpdata.ML" | |
| 44 | setup simpsetup | |
| 45 | setup "Simplifier.method_setup Splitter.split_modifiers" | |
| 46 | setup Splitter.setup | |
| 47 | setup Clasimp.setup | |
| 48 | ||
| 15481 | 49 | |
| 50 | subsection {* Lucas Dixon's eqstep tactic *}
 | |
| 51 | ||
| 52 | use "~~/src/Provers/eqsubst.ML"; | |
| 53 | use "eqrule_FOL_data.ML"; | |
| 54 | ||
| 55 | setup EQSubstTac.setup | |
| 56 | ||
| 57 | ||
| 14085 | 58 | subsection {* Other simple lemmas *}
 | 
| 59 | ||
| 60 | lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)" | |
| 61 | by blast | |
| 62 | ||
| 63 | lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" | |
| 64 | by blast | |
| 65 | ||
| 66 | lemma not_disj_iff_imp: "~P | Q <-> (P-->Q)" | |
| 67 | by blast | |
| 68 | ||
| 69 | (** Monotonicity of implications **) | |
| 70 | ||
| 71 | lemma conj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)" | |
| 72 | by fast (*or (IntPr.fast_tac 1)*) | |
| 73 | ||
| 74 | lemma disj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)" | |
| 75 | by fast (*or (IntPr.fast_tac 1)*) | |
| 76 | ||
| 77 | lemma imp_mono: "[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)" | |
| 78 | by fast (*or (IntPr.fast_tac 1)*) | |
| 79 | ||
| 80 | lemma imp_refl: "P-->P" | |
| 81 | by (rule impI, assumption) | |
| 82 | ||
| 83 | (*The quantifier monotonicity rules are also intuitionistically valid*) | |
| 84 | lemma ex_mono: "(!!x. P(x) --> Q(x)) ==> (EX x. P(x)) --> (EX x. Q(x))" | |
| 85 | by blast | |
| 86 | ||
| 87 | lemma all_mono: "(!!x. P(x) --> Q(x)) ==> (ALL x. P(x)) --> (ALL x. Q(x))" | |
| 88 | by blast | |
| 89 | ||
| 11678 | 90 | |
| 91 | subsection {* Proof by cases and induction *}
 | |
| 92 | ||
| 93 | text {* Proper handling of non-atomic rule statements. *}
 | |
| 94 | ||
| 95 | constdefs | |
| 18456 | 96 | induct_forall where "induct_forall(P) == \<forall>x. P(x)" | 
| 97 | induct_implies where "induct_implies(A, B) == A \<longrightarrow> B" | |
| 98 | induct_equal where "induct_equal(x, y) == x = y" | |
| 99 | induct_conj where "induct_conj(A, B) == A \<and> B" | |
| 11678 | 100 | |
| 101 | lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(\<lambda>x. P(x)))" | |
| 18456 | 102 | by (unfold atomize_all induct_forall_def) | 
| 11678 | 103 | |
| 104 | lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))" | |
| 18456 | 105 | by (unfold atomize_imp induct_implies_def) | 
| 11678 | 106 | |
| 107 | lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))" | |
| 18456 | 108 | by (unfold atomize_eq induct_equal_def) | 
| 11678 | 109 | |
| 18456 | 110 | lemma induct_conj_eq: | 
| 111 | includes meta_conjunction_syntax | |
| 112 | shows "(A && B) == Trueprop(induct_conj(A, B))" | |
| 113 | by (unfold atomize_conj induct_conj_def) | |
| 11988 | 114 | |
| 18456 | 115 | lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq | 
| 116 | lemmas induct_rulify [symmetric, standard] = induct_atomize | |
| 117 | lemmas induct_rulify_fallback = | |
| 118 | induct_forall_def induct_implies_def induct_equal_def induct_conj_def | |
| 11678 | 119 | |
| 18456 | 120 | hide const induct_forall induct_implies induct_equal induct_conj | 
| 11678 | 121 | |
| 122 | ||
| 123 | text {* Method setup. *}
 | |
| 124 | ||
| 125 | ML {*
 | |
| 126 | structure InductMethod = InductMethodFun | |
| 127 | (struct | |
| 128 | val cases_default = thm "case_split"; | |
| 129 | val atomize = thms "induct_atomize"; | |
| 18456 | 130 | val rulify = thms "induct_rulify"; | 
| 131 | val rulify_fallback = thms "induct_rulify_fallback"; | |
| 11678 | 132 | end); | 
| 133 | *} | |
| 134 | ||
| 135 | setup InductMethod.setup | |
| 136 | ||
| 4854 | 137 | end |