author | haftmann |
Tue, 08 Aug 2006 08:19:15 +0200 | |
changeset 20351 | c7658e811ffb |
parent 20223 | 89d2758ecddf |
child 21539 | c5cf9243ad62 |
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 |
||
20223 | 37 |
ML {* val ex1_functional = thm "ex1_functional" *} |
38 |
||
39 |
(* Elimination of True from asumptions: *) |
|
40 |
lemma True_implies_equals: "(True ==> PROP P) == PROP P" |
|
41 |
proof |
|
42 |
assume "True \<Longrightarrow> PROP P" |
|
43 |
from this and TrueI show "PROP P" . |
|
44 |
next |
|
45 |
assume "PROP P" |
|
46 |
then show "PROP P" . |
|
47 |
qed |
|
9487 | 48 |
|
49 |
use "simpdata.ML" |
|
50 |
setup simpsetup |
|
51 |
setup "Simplifier.method_setup Splitter.split_modifiers" |
|
52 |
setup Splitter.setup |
|
53 |
setup Clasimp.setup |
|
18591 | 54 |
setup EqSubst.setup |
15481 | 55 |
|
56 |
||
14085 | 57 |
subsection {* Other simple lemmas *} |
58 |
||
59 |
lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)" |
|
60 |
by blast |
|
61 |
||
62 |
lemma [simp]: "((P-->Q) <-> (P-->R)) <-> (P --> (Q<->R))" |
|
63 |
by blast |
|
64 |
||
65 |
lemma not_disj_iff_imp: "~P | Q <-> (P-->Q)" |
|
66 |
by blast |
|
67 |
||
68 |
(** Monotonicity of implications **) |
|
69 |
||
70 |
lemma conj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)" |
|
71 |
by fast (*or (IntPr.fast_tac 1)*) |
|
72 |
||
73 |
lemma disj_mono: "[| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)" |
|
74 |
by fast (*or (IntPr.fast_tac 1)*) |
|
75 |
||
76 |
lemma imp_mono: "[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)" |
|
77 |
by fast (*or (IntPr.fast_tac 1)*) |
|
78 |
||
79 |
lemma imp_refl: "P-->P" |
|
80 |
by (rule impI, assumption) |
|
81 |
||
82 |
(*The quantifier monotonicity rules are also intuitionistically valid*) |
|
83 |
lemma ex_mono: "(!!x. P(x) --> Q(x)) ==> (EX x. P(x)) --> (EX x. Q(x))" |
|
84 |
by blast |
|
85 |
||
86 |
lemma all_mono: "(!!x. P(x) --> Q(x)) ==> (ALL x. P(x)) --> (ALL x. Q(x))" |
|
87 |
by blast |
|
88 |
||
11678 | 89 |
|
90 |
subsection {* Proof by cases and induction *} |
|
91 |
||
92 |
text {* Proper handling of non-atomic rule statements. *} |
|
93 |
||
94 |
constdefs |
|
18456 | 95 |
induct_forall where "induct_forall(P) == \<forall>x. P(x)" |
96 |
induct_implies where "induct_implies(A, B) == A \<longrightarrow> B" |
|
97 |
induct_equal where "induct_equal(x, y) == x = y" |
|
98 |
induct_conj where "induct_conj(A, B) == A \<and> B" |
|
11678 | 99 |
|
100 |
lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(\<lambda>x. P(x)))" |
|
18816 | 101 |
unfolding atomize_all induct_forall_def . |
11678 | 102 |
|
103 |
lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))" |
|
18816 | 104 |
unfolding atomize_imp induct_implies_def . |
11678 | 105 |
|
106 |
lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))" |
|
18816 | 107 |
unfolding atomize_eq induct_equal_def . |
11678 | 108 |
|
18456 | 109 |
lemma induct_conj_eq: |
110 |
includes meta_conjunction_syntax |
|
111 |
shows "(A && B) == Trueprop(induct_conj(A, B))" |
|
18816 | 112 |
unfolding atomize_conj induct_conj_def . |
11988 | 113 |
|
18456 | 114 |
lemmas induct_atomize = induct_forall_eq induct_implies_eq induct_equal_eq induct_conj_eq |
115 |
lemmas induct_rulify [symmetric, standard] = induct_atomize |
|
116 |
lemmas induct_rulify_fallback = |
|
117 |
induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
11678 | 118 |
|
18456 | 119 |
hide const induct_forall induct_implies induct_equal induct_conj |
11678 | 120 |
|
121 |
||
122 |
text {* Method setup. *} |
|
123 |
||
124 |
ML {* |
|
125 |
structure InductMethod = InductMethodFun |
|
126 |
(struct |
|
127 |
val cases_default = thm "case_split"; |
|
128 |
val atomize = thms "induct_atomize"; |
|
18456 | 129 |
val rulify = thms "induct_rulify"; |
130 |
val rulify_fallback = thms "induct_rulify_fallback"; |
|
11678 | 131 |
end); |
132 |
*} |
|
133 |
||
134 |
setup InductMethod.setup |
|
135 |
||
4854 | 136 |
end |