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