author | wenzelm |
Thu, 24 Jan 2002 16:37:49 +0100 | |
changeset 12844 | b5b15bbca582 |
parent 12367 | 1cee8a0db392 |
child 13550 | 5a176b8dda84 |
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 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
8 |
theory FOL = IFOL |
9487 | 9 |
files |
10 |
("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") |
|
11 |
("simpdata.ML") ("FOL_lemmas2.ML"): |
|
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 |
|
27 |
setup clasetup |
|
28 |
||
9487 | 29 |
use "blastdata.ML" |
30 |
setup Blast.setup |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
31 |
use "FOL_lemmas2.ML" |
9487 | 32 |
|
33 |
use "simpdata.ML" |
|
34 |
setup simpsetup |
|
35 |
setup "Simplifier.method_setup Splitter.split_modifiers" |
|
36 |
setup Splitter.setup |
|
37 |
setup Clasimp.setup |
|
38 |
||
11678 | 39 |
|
40 |
subsection {* Proof by cases and induction *} |
|
41 |
||
42 |
text {* Proper handling of non-atomic rule statements. *} |
|
43 |
||
44 |
constdefs |
|
45 |
induct_forall :: "('a => o) => o" |
|
46 |
"induct_forall(P) == \<forall>x. P(x)" |
|
47 |
induct_implies :: "o => o => o" |
|
48 |
"induct_implies(A, B) == A --> B" |
|
49 |
induct_equal :: "'a => 'a => o" |
|
50 |
"induct_equal(x, y) == x = y" |
|
51 |
||
52 |
lemma induct_forall_eq: "(!!x. P(x)) == Trueprop(induct_forall(\<lambda>x. P(x)))" |
|
53 |
by (simp only: atomize_all induct_forall_def) |
|
54 |
||
55 |
lemma induct_implies_eq: "(A ==> B) == Trueprop(induct_implies(A, B))" |
|
56 |
by (simp only: atomize_imp induct_implies_def) |
|
57 |
||
58 |
lemma induct_equal_eq: "(x == y) == Trueprop(induct_equal(x, y))" |
|
59 |
by (simp only: atomize_eq induct_equal_def) |
|
60 |
||
11988 | 61 |
lemma induct_impliesI: "(A ==> B) ==> induct_implies(A, B)" |
62 |
by (simp add: induct_implies_def) |
|
63 |
||
12164
0b219d9e3384
induct_atomize: include atomize_conj (for mutual induction);
wenzelm
parents:
12160
diff
changeset
|
64 |
lemmas induct_atomize = atomize_conj induct_forall_eq induct_implies_eq induct_equal_eq |
0b219d9e3384
induct_atomize: include atomize_conj (for mutual induction);
wenzelm
parents:
12160
diff
changeset
|
65 |
lemmas induct_rulify1 [symmetric, standard] = induct_forall_eq induct_implies_eq induct_equal_eq |
11678 | 66 |
lemmas induct_rulify2 = induct_forall_def induct_implies_def induct_equal_def |
67 |
||
12240 | 68 |
lemma all_conj_eq: "(ALL x. P(x)) & (ALL y. Q(y)) == (ALL x y. P(x) & Q(y))" |
69 |
by simp |
|
70 |
||
11678 | 71 |
hide const induct_forall induct_implies induct_equal |
72 |
||
73 |
||
74 |
text {* Method setup. *} |
|
75 |
||
76 |
ML {* |
|
77 |
structure InductMethod = InductMethodFun |
|
78 |
(struct |
|
79 |
val dest_concls = FOLogic.dest_concls; |
|
80 |
val cases_default = thm "case_split"; |
|
11988 | 81 |
val local_impI = thm "induct_impliesI"; |
11678 | 82 |
val conjI = thm "conjI"; |
83 |
val atomize = thms "induct_atomize"; |
|
84 |
val rulify1 = thms "induct_rulify1"; |
|
85 |
val rulify2 = thms "induct_rulify2"; |
|
12240 | 86 |
val localize = [Thm.symmetric (thm "induct_implies_def"), |
87 |
Thm.symmetric (thm "atomize_all"), thm "all_conj_eq"]; |
|
11678 | 88 |
end); |
89 |
*} |
|
90 |
||
91 |
setup InductMethod.setup |
|
92 |
||
4854 | 93 |
end |