author | wenzelm |
Tue, 04 Dec 2001 02:00:14 +0100 | |
changeset 12352 | 92c48cc45e78 |
parent 12349 | 94e812f9683e |
child 12368 | 2af9ad81ea56 |
permissions | -rw-r--r-- |
1268 | 1 |
(* Title: FOL/IFOL.thy |
35 | 2 |
ID: $Id$ |
11677 | 3 |
Author: Lawrence C Paulson and Markus Wenzel |
4 |
*) |
|
35 | 5 |
|
11677 | 6 |
header {* Intuitionistic first-order logic *} |
35 | 7 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
8 |
theory IFOL = Pure |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
9 |
files ("IFOL_lemmas.ML") ("fologic.ML") ("hypsubstdata.ML") ("intprover.ML"): |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
10 |
|
0 | 11 |
|
11677 | 12 |
subsection {* Syntax and axiomatic basis *} |
13 |
||
3906 | 14 |
global |
15 |
||
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
16 |
classes "term" < logic |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
17 |
defaultsort "term" |
0 | 18 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
19 |
typedecl o |
79 | 20 |
|
11747 | 21 |
judgment |
22 |
Trueprop :: "o => prop" ("(_)" 5) |
|
0 | 23 |
|
11747 | 24 |
consts |
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
25 |
True :: o |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
26 |
False :: o |
79 | 27 |
|
28 |
(* Connectives *) |
|
29 |
||
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
30 |
"=" :: "['a, 'a] => o" (infixl 50) |
35 | 31 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
32 |
Not :: "o => o" ("~ _" [40] 40) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
33 |
& :: "[o, o] => o" (infixr 35) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
34 |
"|" :: "[o, o] => o" (infixr 30) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
35 |
--> :: "[o, o] => o" (infixr 25) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
36 |
<-> :: "[o, o] => o" (infixr 25) |
79 | 37 |
|
38 |
(* Quantifiers *) |
|
39 |
||
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
40 |
All :: "('a => o) => o" (binder "ALL " 10) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
41 |
Ex :: "('a => o) => o" (binder "EX " 10) |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
42 |
Ex1 :: "('a => o) => o" (binder "EX! " 10) |
79 | 43 |
|
0 | 44 |
|
928 | 45 |
syntax |
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
46 |
"~=" :: "['a, 'a] => o" (infixl 50) |
35 | 47 |
translations |
79 | 48 |
"x ~= y" == "~ (x = y)" |
49 |
||
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12019
diff
changeset
|
50 |
syntax (xsymbols) |
11677 | 51 |
Not :: "o => o" ("\<not> _" [40] 40) |
52 |
"op &" :: "[o, o] => o" (infixr "\<and>" 35) |
|
53 |
"op |" :: "[o, o] => o" (infixr "\<or>" 30) |
|
54 |
"ALL " :: "[idts, o] => o" ("(3\<forall>_./ _)" [0, 10] 10) |
|
55 |
"EX " :: "[idts, o] => o" ("(3\<exists>_./ _)" [0, 10] 10) |
|
56 |
"EX! " :: "[idts, o] => o" ("(3\<exists>!_./ _)" [0, 10] 10) |
|
57 |
"op ~=" :: "['a, 'a] => o" (infixl "\<noteq>" 50) |
|
58 |
"op -->" :: "[o, o] => o" (infixr "\<longrightarrow>" 25) |
|
59 |
"op <->" :: "[o, o] => o" (infixr "\<longleftrightarrow>" 25) |
|
35 | 60 |
|
6340 | 61 |
syntax (HTML output) |
11677 | 62 |
Not :: "o => o" ("\<not> _" [40] 40) |
6340 | 63 |
|
64 |
||
3932 | 65 |
local |
3906 | 66 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
67 |
axioms |
0 | 68 |
|
79 | 69 |
(* Equality *) |
0 | 70 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
71 |
refl: "a=a" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
72 |
subst: "[| a=b; P(a) |] ==> P(b)" |
0 | 73 |
|
79 | 74 |
(* Propositional logic *) |
0 | 75 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
76 |
conjI: "[| P; Q |] ==> P&Q" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
77 |
conjunct1: "P&Q ==> P" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
78 |
conjunct2: "P&Q ==> Q" |
0 | 79 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
80 |
disjI1: "P ==> P|Q" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
81 |
disjI2: "Q ==> P|Q" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
82 |
disjE: "[| P|Q; P ==> R; Q ==> R |] ==> R" |
0 | 83 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
84 |
impI: "(P ==> Q) ==> P-->Q" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
85 |
mp: "[| P-->Q; P |] ==> Q" |
0 | 86 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
87 |
FalseE: "False ==> P" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
88 |
|
0 | 89 |
|
79 | 90 |
(* Definitions *) |
0 | 91 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
92 |
True_def: "True == False-->False" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
93 |
not_def: "~P == P-->False" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
94 |
iff_def: "P<->Q == (P-->Q) & (Q-->P)" |
79 | 95 |
|
96 |
(* Unique existence *) |
|
0 | 97 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
98 |
ex1_def: "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
99 |
|
0 | 100 |
|
79 | 101 |
(* Quantifiers *) |
0 | 102 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
103 |
allI: "(!!x. P(x)) ==> (ALL x. P(x))" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
104 |
spec: "(ALL x. P(x)) ==> P(x)" |
0 | 105 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
106 |
exI: "P(x) ==> (EX x. P(x))" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
107 |
exE: "[| EX x. P(x); !!x. P(x) ==> R |] ==> R" |
0 | 108 |
|
109 |
(* Reflection *) |
|
110 |
||
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
111 |
eq_reflection: "(x=y) ==> (x==y)" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
112 |
iff_reflection: "(P<->Q) ==> (P==Q)" |
0 | 113 |
|
4092 | 114 |
|
11677 | 115 |
subsection {* Lemmas and proof tools *} |
116 |
||
9886 | 117 |
setup Simplifier.setup |
118 |
use "IFOL_lemmas.ML" |
|
11734 | 119 |
|
12349 | 120 |
declare impE [Pure.elim?] iffD1 [Pure.elim?] iffD2 [Pure.elim?] |
11734 | 121 |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
122 |
use "fologic.ML" |
9886 | 123 |
use "hypsubstdata.ML" |
124 |
setup hypsubst_setup |
|
7355
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
125 |
use "intprover.ML" |
4c43090659ca
proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents:
6340
diff
changeset
|
126 |
|
4092 | 127 |
|
12349 | 128 |
lemma impE': |
129 |
(assumes 1: "P --> Q" and 2: "Q ==> R" and 3: "P --> Q ==> P") R |
|
130 |
proof - |
|
131 |
from 3 and 1 have P . |
|
132 |
with 1 have Q .. |
|
133 |
with 2 show R . |
|
134 |
qed |
|
135 |
||
136 |
lemma allE': |
|
137 |
(assumes 1: "ALL x. P(x)" and 2: "P(x) ==> ALL x. P(x) ==> Q") Q |
|
138 |
proof - |
|
139 |
from 1 have "P(x)" by (rule spec) |
|
140 |
from this and 1 show Q by (rule 2) |
|
141 |
qed |
|
142 |
||
143 |
lemma notE': (assumes 1: "~ P" and 2: "~ P ==> P") R |
|
144 |
proof - |
|
145 |
from 2 and 1 have P . |
|
146 |
with 1 show R by (rule notE) |
|
147 |
qed |
|
148 |
||
149 |
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE |
|
150 |
and [Pure.intro!] = iffI conjI impI TrueI notI allI refl |
|
151 |
and [Pure.elim 2] = allE notE' impE' |
|
152 |
and [Pure.intro] = exI disjI2 disjI1 |
|
153 |
||
154 |
ML_setup {* |
|
12352 | 155 |
Context.>> (ContextRules.addSWrapper (fn tac => hyp_subst_tac ORELSE' tac)); |
12349 | 156 |
*} |
157 |
||
158 |
||
11677 | 159 |
subsection {* Atomizing meta-level rules *} |
160 |
||
11747 | 161 |
lemma atomize_all [atomize]: "(!!x. P(x)) == Trueprop (ALL x. P(x))" |
11976 | 162 |
proof |
11677 | 163 |
assume "!!x. P(x)" |
164 |
show "ALL x. P(x)" by (rule allI) |
|
165 |
next |
|
166 |
assume "ALL x. P(x)" |
|
167 |
thus "!!x. P(x)" by (rule allE) |
|
168 |
qed |
|
169 |
||
11747 | 170 |
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)" |
11976 | 171 |
proof |
11677 | 172 |
assume r: "A ==> B" |
173 |
show "A --> B" by (rule impI) (rule r) |
|
174 |
next |
|
175 |
assume "A --> B" and A |
|
176 |
thus B by (rule mp) |
|
177 |
qed |
|
178 |
||
11747 | 179 |
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)" |
11976 | 180 |
proof |
11677 | 181 |
assume "x == y" |
182 |
show "x = y" by (unfold prems) (rule refl) |
|
183 |
next |
|
184 |
assume "x = y" |
|
185 |
thus "x == y" by (rule eq_reflection) |
|
186 |
qed |
|
187 |
||
11953 | 188 |
lemma atomize_conj [atomize]: "(!!C. (A ==> B ==> PROP C) ==> PROP C) == Trueprop (A & B)" |
11976 | 189 |
proof |
11953 | 190 |
assume "!!C. (A ==> B ==> PROP C) ==> PROP C" |
191 |
show "A & B" by (rule conjI) |
|
192 |
next |
|
193 |
fix C |
|
194 |
assume "A & B" |
|
195 |
assume "A ==> B ==> PROP C" |
|
196 |
thus "PROP C" |
|
197 |
proof this |
|
198 |
show A by (rule conjunct1) |
|
199 |
show B by (rule conjunct2) |
|
200 |
qed |
|
201 |
qed |
|
202 |
||
11771 | 203 |
declare atomize_all [symmetric, rulify] atomize_imp [symmetric, rulify] |
204 |
||
11848 | 205 |
|
206 |
subsection {* Calculational rules *} |
|
207 |
||
208 |
lemma forw_subst: "a = b ==> P(b) ==> P(a)" |
|
209 |
by (rule ssubst) |
|
210 |
||
211 |
lemma back_subst: "P(a) ==> a = b ==> P(b)" |
|
212 |
by (rule subst) |
|
213 |
||
214 |
text {* |
|
215 |
Note that this list of rules is in reverse order of priorities. |
|
216 |
*} |
|
217 |
||
12019 | 218 |
lemmas basic_trans_rules [trans] = |
11848 | 219 |
forw_subst |
220 |
back_subst |
|
221 |
rev_mp |
|
222 |
mp |
|
223 |
trans |
|
224 |
||
12349 | 225 |
lemmas [Pure.elim?] = sym |
11848 | 226 |
|
4854 | 227 |
end |