author | wenzelm |
Thu, 07 Sep 2000 20:47:34 +0200 | |
changeset 9885 | 34494703d283 |
parent 9713 | 2c5b42311eb0 |
child 10130 | 5a2e00bf1e42 |
permissions | -rw-r--r-- |
9487 | 1 |
(* Title: FOL/FOL.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson and Markus Wenzel |
|
4 |
||
5 |
Classical first-order logic. |
|
6 |
||
7 |
This may serve as a good example of initializing all the tools and |
|
8 |
packages required for a reasonable working environment. Please go |
|
9 |
elsewhere to see actual applications! |
|
10 |
*) |
|
4093 | 11 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
12 |
theory FOL = IFOL |
9487 | 13 |
files |
14 |
("FOL_lemmas1.ML") ("cladata.ML") ("blastdata.ML") |
|
15 |
("simpdata.ML") ("FOL_lemmas2.ML"): |
|
16 |
||
17 |
||
18 |
subsection {* The classical axiom *} |
|
4093 | 19 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
20 |
axioms |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
21 |
classical: "(~P ==> P) ==> P" |
4093 | 22 |
|
9487 | 23 |
|
24 |
subsection {* Setup of several proof tools *} |
|
25 |
||
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
26 |
use "FOL_lemmas1.ML" |
9487 | 27 |
use "cladata.ML" |
28 |
setup Cla.setup |
|
29 |
setup clasetup |
|
30 |
||
31 |
lemma all_eq: "(!!x. P(x)) == Trueprop (ALL x. P(x))" |
|
32 |
proof (rule equal_intr_rule) |
|
33 |
assume "!!x. P(x)" |
|
34 |
show "ALL x. P(x)" .. |
|
35 |
next |
|
36 |
assume "ALL x. P(x)" |
|
37 |
thus "!!x. P(x)" .. |
|
38 |
qed |
|
39 |
||
40 |
lemma imp_eq: "(A ==> B) == Trueprop (A --> B)" |
|
41 |
proof (rule equal_intr_rule) |
|
42 |
assume r: "A ==> B" |
|
43 |
show "A --> B" |
|
44 |
by (rule) (rule r) |
|
45 |
next |
|
46 |
assume "A --> B" and A |
|
47 |
thus B .. |
|
48 |
qed |
|
49 |
||
9525 | 50 |
lemmas atomize = all_eq imp_eq |
51 |
||
9487 | 52 |
use "blastdata.ML" |
53 |
setup Blast.setup |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
5887
diff
changeset
|
54 |
use "FOL_lemmas2.ML" |
9487 | 55 |
|
56 |
use "simpdata.ML" |
|
57 |
setup simpsetup |
|
58 |
setup "Simplifier.method_setup Splitter.split_modifiers" |
|
59 |
setup Splitter.setup |
|
60 |
setup Clasimp.setup |
|
9885 | 61 |
setup Rulify.setup |
9487 | 62 |
|
63 |
||
64 |
subsection {* Calculational rules *} |
|
65 |
||
66 |
lemma forw_subst: "a = b ==> P(b) ==> P(a)" |
|
67 |
by (rule ssubst) |
|
68 |
||
69 |
lemma back_subst: "P(a) ==> a = b ==> P(b)" |
|
70 |
by (rule subst) |
|
71 |
||
72 |
text {* |
|
73 |
Note that this list of rules is in reverse order of priorities. |
|
74 |
*} |
|
75 |
||
76 |
lemmas trans_rules [trans] = |
|
77 |
forw_subst |
|
78 |
back_subst |
|
79 |
rev_mp |
|
80 |
mp |
|
81 |
trans |
|
82 |
||
83 |
lemmas [elim?] = sym |
|
4093 | 84 |
|
4854 | 85 |
end |