author | wenzelm |
Tue, 23 Nov 2021 12:29:09 +0100 | |
changeset 74832 | c299abcf7081 |
parent 71924 | e5df9c8d9d4b |
child 76987 | 4c275405faae |
permissions | -rw-r--r-- |
71924
e5df9c8d9d4b
clarified sessions: "Notable Examples in Isabelle/Pure";
wenzelm
parents:
71831
diff
changeset
|
1 |
(* Title: Pure/Examples/Higher_Order_Logic.thy |
61759 | 2 |
Author: Makarius |
12360 | 3 |
*) |
4 |
||
59031 | 5 |
section \<open>Foundations of HOL\<close> |
12360 | 6 |
|
60769 | 7 |
theory Higher_Order_Logic |
63585 | 8 |
imports Pure |
60769 | 9 |
begin |
12360 | 10 |
|
59031 | 11 |
text \<open> |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
12 |
The following theory development illustrates the foundations of Higher-Order |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
13 |
Logic. The ``HOL'' logic that is given here resembles @{cite |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
14 |
"Gordon:1985:HOL"} and its predecessor @{cite "church40"}, but the order of |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
15 |
axiomatizations and defined connectives has be adapted to modern |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
16 |
presentations of \<open>\<lambda>\<close>-calculus and Constructive Type Theory. Thus it fits |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
17 |
nicely to the underlying Natural Deduction framework of Isabelle/Pure and |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
18 |
Isabelle/Isar. |
59031 | 19 |
\<close> |
12360 | 20 |
|
21 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
22 |
section \<open>HOL syntax within Pure\<close> |
12360 | 23 |
|
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
24 |
class type |
36452 | 25 |
default_sort type |
12360 | 26 |
|
27 |
typedecl o |
|
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
28 |
instance o :: type .. |
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
29 |
instance "fun" :: (type, type) type .. |
12360 | 30 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
31 |
judgment Trueprop :: "o \<Rightarrow> prop" ("_" 5) |
12360 | 32 |
|
33 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
34 |
section \<open>Minimal logic (axiomatization)\<close> |
12360 | 35 |
|
61759 | 36 |
axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25) |
37 |
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" |
|
38 |
and impE [dest, trans]: "A \<longrightarrow> B \<Longrightarrow> A \<Longrightarrow> B" |
|
39 |
||
40 |
axiomatization All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10) |
|
41 |
where allI [intro]: "(\<And>x. P x) \<Longrightarrow> \<forall>x. P x" |
|
42 |
and allE [dest]: "\<forall>x. P x \<Longrightarrow> P a" |
|
12360 | 43 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
44 |
lemma atomize_imp [atomize]: "(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
45 |
by standard (fact impI, fact impE) |
12360 | 46 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
47 |
lemma atomize_all [atomize]: "(\<And>x. P x) \<equiv> Trueprop (\<forall>x. P x)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
48 |
by standard (fact allI, fact allE) |
12360 | 49 |
|
50 |
||
59031 | 51 |
subsubsection \<open>Derived connectives\<close> |
12360 | 52 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
53 |
definition False :: o |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
54 |
where "False \<equiv> \<forall>A. A" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
55 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
56 |
lemma FalseE [elim]: |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
57 |
assumes "False" |
61759 | 58 |
shows A |
59 |
proof - |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
60 |
from \<open>False\<close> have "\<forall>A. A" by (simp only: False_def) |
61759 | 61 |
then show A .. |
62 |
qed |
|
63 |
||
64 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
65 |
definition True :: o |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
66 |
where "True \<equiv> False \<longrightarrow> False" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
67 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
68 |
lemma TrueI [intro]: True |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
69 |
unfolding True_def .. |
61759 | 70 |
|
71 |
||
59031 | 72 |
definition not :: "o \<Rightarrow> o" ("\<not> _" [40] 40) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
73 |
where "not \<equiv> \<lambda>A. A \<longrightarrow> False" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
74 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
75 |
lemma notI [intro]: |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
76 |
assumes "A \<Longrightarrow> False" |
61759 | 77 |
shows "\<not> A" |
78 |
using assms unfolding not_def .. |
|
12360 | 79 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
80 |
lemma notE [elim]: |
61759 | 81 |
assumes "\<not> A" and A |
82 |
shows B |
|
83 |
proof - |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
84 |
from \<open>\<not> A\<close> have "A \<longrightarrow> False" by (simp only: not_def) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
85 |
from this and \<open>A\<close> have "False" .. |
23373 | 86 |
then show B .. |
12360 | 87 |
qed |
88 |
||
89 |
lemma notE': "A \<Longrightarrow> \<not> A \<Longrightarrow> B" |
|
90 |
by (rule notE) |
|
91 |
||
61759 | 92 |
lemmas contradiction = notE notE' \<comment> \<open>proof by contradiction in any order\<close> |
93 |
||
94 |
||
95 |
definition conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
96 |
where "A \<and> B \<equiv> \<forall>C. (A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" |
12360 | 97 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
98 |
lemma conjI [intro]: |
61759 | 99 |
assumes A and B |
100 |
shows "A \<and> B" |
|
101 |
unfolding conj_def |
|
102 |
proof |
|
103 |
fix C |
|
104 |
show "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 105 |
proof |
61759 | 106 |
assume "A \<longrightarrow> B \<longrightarrow> C" |
107 |
also note \<open>A\<close> |
|
108 |
also note \<open>B\<close> |
|
109 |
finally show C . |
|
12360 | 110 |
qed |
111 |
qed |
|
112 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
113 |
lemma conjE [elim]: |
61759 | 114 |
assumes "A \<and> B" |
115 |
obtains A and B |
|
116 |
proof |
|
117 |
from \<open>A \<and> B\<close> have *: "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" for C |
|
118 |
unfolding conj_def .. |
|
119 |
show A |
|
120 |
proof - |
|
121 |
note * [of A] |
|
12360 | 122 |
also have "A \<longrightarrow> B \<longrightarrow> A" |
123 |
proof |
|
124 |
assume A |
|
23373 | 125 |
then show "B \<longrightarrow> A" .. |
12360 | 126 |
qed |
61759 | 127 |
finally show ?thesis . |
128 |
qed |
|
129 |
show B |
|
130 |
proof - |
|
131 |
note * [of B] |
|
12360 | 132 |
also have "A \<longrightarrow> B \<longrightarrow> B" |
133 |
proof |
|
134 |
show "B \<longrightarrow> B" .. |
|
135 |
qed |
|
61759 | 136 |
finally show ?thesis . |
12360 | 137 |
qed |
138 |
qed |
|
139 |
||
61759 | 140 |
|
141 |
definition disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
142 |
where "A \<or> B \<equiv> \<forall>C. (A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
61759 | 143 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
144 |
lemma disjI1 [intro]: |
61759 | 145 |
assumes A |
146 |
shows "A \<or> B" |
|
147 |
unfolding disj_def |
|
148 |
proof |
|
149 |
fix C |
|
150 |
show "(A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 151 |
proof |
61759 | 152 |
assume "A \<longrightarrow> C" |
153 |
from this and \<open>A\<close> have C .. |
|
154 |
then show "(B \<longrightarrow> C) \<longrightarrow> C" .. |
|
155 |
qed |
|
156 |
qed |
|
157 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
158 |
lemma disjI2 [intro]: |
61759 | 159 |
assumes B |
160 |
shows "A \<or> B" |
|
161 |
unfolding disj_def |
|
162 |
proof |
|
163 |
fix C |
|
164 |
show "(A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
|
165 |
proof |
|
166 |
show "(B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 167 |
proof |
61759 | 168 |
assume "B \<longrightarrow> C" |
169 |
from this and \<open>B\<close> show C .. |
|
12360 | 170 |
qed |
171 |
qed |
|
172 |
qed |
|
173 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
174 |
lemma disjE [elim]: |
61759 | 175 |
assumes "A \<or> B" |
176 |
obtains (a) A | (b) B |
|
177 |
proof - |
|
178 |
from \<open>A \<or> B\<close> have "(A \<longrightarrow> thesis) \<longrightarrow> (B \<longrightarrow> thesis) \<longrightarrow> thesis" |
|
179 |
unfolding disj_def .. |
|
180 |
also have "A \<longrightarrow> thesis" |
|
12360 | 181 |
proof |
60769 | 182 |
assume A |
61759 | 183 |
then show thesis by (rule a) |
12360 | 184 |
qed |
61759 | 185 |
also have "B \<longrightarrow> thesis" |
12360 | 186 |
proof |
60769 | 187 |
assume B |
61759 | 188 |
then show thesis by (rule b) |
12360 | 189 |
qed |
61759 | 190 |
finally show thesis . |
12360 | 191 |
qed |
192 |
||
61759 | 193 |
|
194 |
definition Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10) |
|
195 |
where "\<exists>x. P x \<equiv> \<forall>C. (\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C" |
|
196 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
197 |
lemma exI [intro]: "P a \<Longrightarrow> \<exists>x. P x" |
61759 | 198 |
unfolding Ex_def |
199 |
proof |
|
200 |
fix C |
|
12360 | 201 |
assume "P a" |
61759 | 202 |
show "(\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C" |
12360 | 203 |
proof |
61759 | 204 |
assume "\<forall>x. P x \<longrightarrow> C" |
205 |
then have "P a \<longrightarrow> C" .. |
|
206 |
from this and \<open>P a\<close> show C .. |
|
12360 | 207 |
qed |
208 |
qed |
|
209 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
210 |
lemma exE [elim]: |
61759 | 211 |
assumes "\<exists>x. P x" |
212 |
obtains (that) x where "P x" |
|
213 |
proof - |
|
214 |
from \<open>\<exists>x. P x\<close> have "(\<forall>x. P x \<longrightarrow> thesis) \<longrightarrow> thesis" |
|
215 |
unfolding Ex_def .. |
|
216 |
also have "\<forall>x. P x \<longrightarrow> thesis" |
|
12360 | 217 |
proof |
61759 | 218 |
fix x |
219 |
show "P x \<longrightarrow> thesis" |
|
12360 | 220 |
proof |
221 |
assume "P x" |
|
61759 | 222 |
then show thesis by (rule that) |
12360 | 223 |
qed |
224 |
qed |
|
61759 | 225 |
finally show thesis . |
12360 | 226 |
qed |
227 |
||
228 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
229 |
subsubsection \<open>Extensional equality\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
230 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
231 |
axiomatization equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "=" 50) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
232 |
where refl [intro]: "x = x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
233 |
and subst: "x = y \<Longrightarrow> P x \<Longrightarrow> P y" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
234 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
235 |
abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "\<noteq>" 50) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
236 |
where "x \<noteq> y \<equiv> \<not> (x = y)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
237 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
238 |
abbreviation iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
239 |
where "A \<longleftrightarrow> B \<equiv> A = B" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
240 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
241 |
axiomatization |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
242 |
where ext [intro]: "(\<And>x. f x = g x) \<Longrightarrow> f = g" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
243 |
and iff [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> (B \<Longrightarrow> A) \<Longrightarrow> A \<longleftrightarrow> B" |
71831 | 244 |
for f g :: "'a \<Rightarrow> 'b" |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
245 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
246 |
lemma sym [sym]: "y = x" if "x = y" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
247 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
248 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
249 |
lemma [trans]: "x = y \<Longrightarrow> P y \<Longrightarrow> P x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
250 |
by (rule subst) (rule sym) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
251 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
252 |
lemma [trans]: "P x \<Longrightarrow> x = y \<Longrightarrow> P y" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
253 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
254 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
255 |
lemma arg_cong: "f x = f y" if "x = y" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
256 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
257 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
258 |
lemma fun_cong: "f x = g x" if "f = g" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
259 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
260 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
261 |
lemma trans [trans]: "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
262 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
263 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
264 |
lemma iff1 [elim]: "A \<longleftrightarrow> B \<Longrightarrow> A \<Longrightarrow> B" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
265 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
266 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
267 |
lemma iff2 [elim]: "A \<longleftrightarrow> B \<Longrightarrow> B \<Longrightarrow> A" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
268 |
by (rule subst) (rule sym) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
269 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
270 |
|
61936 | 271 |
subsection \<open>Cantor's Theorem\<close> |
272 |
||
273 |
text \<open> |
|
274 |
Cantor's Theorem states that there is no surjection from a set to its |
|
275 |
powerset. The subsequent formulation uses elementary \<open>\<lambda>\<close>-calculus and |
|
276 |
predicate logic, with standard introduction and elimination rules. |
|
277 |
\<close> |
|
278 |
||
279 |
lemma iff_contradiction: |
|
280 |
assumes *: "\<not> A \<longleftrightarrow> A" |
|
281 |
shows C |
|
282 |
proof (rule notE) |
|
283 |
show "\<not> A" |
|
284 |
proof |
|
285 |
assume A |
|
286 |
with * have "\<not> A" .. |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
287 |
from this and \<open>A\<close> show False .. |
61936 | 288 |
qed |
289 |
with * show A .. |
|
290 |
qed |
|
291 |
||
62038 | 292 |
theorem Cantor: "\<not> (\<exists>f :: 'a \<Rightarrow> 'a \<Rightarrow> o. \<forall>A. \<exists>x. A = f x)" |
61936 | 293 |
proof |
294 |
assume "\<exists>f :: 'a \<Rightarrow> 'a \<Rightarrow> o. \<forall>A. \<exists>x. A = f x" |
|
295 |
then obtain f :: "'a \<Rightarrow> 'a \<Rightarrow> o" where *: "\<forall>A. \<exists>x. A = f x" .. |
|
296 |
let ?D = "\<lambda>x. \<not> f x x" |
|
297 |
from * have "\<exists>x. ?D = f x" .. |
|
298 |
then obtain a where "?D = f a" .. |
|
299 |
then have "?D a \<longleftrightarrow> f a a" using refl by (rule subst) |
|
62266 | 300 |
then have "\<not> f a a \<longleftrightarrow> f a a" . |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
301 |
then show False by (rule iff_contradiction) |
61936 | 302 |
qed |
303 |
||
304 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
305 |
subsection \<open>Characterization of Classical Logic\<close> |
12360 | 306 |
|
61759 | 307 |
text \<open> |
308 |
The subsequent rules of classical reasoning are all equivalent. |
|
309 |
\<close> |
|
310 |
||
12360 | 311 |
locale classical = |
312 |
assumes classical: "(\<not> A \<Longrightarrow> A) \<Longrightarrow> A" |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
313 |
\<comment> \<open>predicate definition and hypothetical context\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
314 |
begin |
12360 | 315 |
|
64908 | 316 |
lemma classical_contradiction: |
317 |
assumes "\<not> A \<Longrightarrow> False" |
|
318 |
shows A |
|
319 |
proof (rule classical) |
|
320 |
assume "\<not> A" |
|
321 |
then have False by (rule assms) |
|
322 |
then show A .. |
|
12360 | 323 |
qed |
324 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
325 |
lemma double_negation: |
61759 | 326 |
assumes "\<not> \<not> A" |
327 |
shows A |
|
64908 | 328 |
proof (rule classical_contradiction) |
61759 | 329 |
assume "\<not> A" |
64908 | 330 |
with \<open>\<not> \<not> A\<close> show False by (rule contradiction) |
12360 | 331 |
qed |
332 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
333 |
lemma tertium_non_datur: "A \<or> \<not> A" |
12360 | 334 |
proof (rule double_negation) |
335 |
show "\<not> \<not> (A \<or> \<not> A)" |
|
336 |
proof |
|
337 |
assume "\<not> (A \<or> \<not> A)" |
|
338 |
have "\<not> A" |
|
339 |
proof |
|
23373 | 340 |
assume A then have "A \<or> \<not> A" .. |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
341 |
with \<open>\<not> (A \<or> \<not> A)\<close> show False by (rule contradiction) |
12360 | 342 |
qed |
23373 | 343 |
then have "A \<or> \<not> A" .. |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
344 |
with \<open>\<not> (A \<or> \<not> A)\<close> show False by (rule contradiction) |
12360 | 345 |
qed |
346 |
qed |
|
347 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
348 |
lemma classical_cases: |
61759 | 349 |
obtains A | "\<not> A" |
350 |
using tertium_non_datur |
|
351 |
proof |
|
352 |
assume A |
|
353 |
then show thesis .. |
|
354 |
next |
|
355 |
assume "\<not> A" |
|
356 |
then show thesis .. |
|
357 |
qed |
|
358 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
359 |
end |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
360 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
361 |
lemma classical_if_cases: classical |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
362 |
if cases: "\<And>A C. (A \<Longrightarrow> C) \<Longrightarrow> (\<not> A \<Longrightarrow> C) \<Longrightarrow> C" |
61759 | 363 |
proof |
364 |
fix A |
|
365 |
assume *: "\<not> A \<Longrightarrow> A" |
|
366 |
show A |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
367 |
proof (rule cases) |
12360 | 368 |
assume A |
61759 | 369 |
then show A . |
12360 | 370 |
next |
371 |
assume "\<not> A" |
|
61759 | 372 |
then show A by (rule *) |
12573 | 373 |
qed |
374 |
qed |
|
375 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
376 |
|
64908 | 377 |
section \<open>Peirce's Law\<close> |
378 |
||
379 |
text \<open> |
|
380 |
Peirce's Law is another characterization of classical reasoning. Its |
|
381 |
statement only requires implication. |
|
382 |
\<close> |
|
383 |
||
384 |
theorem (in classical) Peirce's_Law: "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A" |
|
385 |
proof |
|
386 |
assume *: "(A \<longrightarrow> B) \<longrightarrow> A" |
|
387 |
show A |
|
388 |
proof (rule classical) |
|
389 |
assume "\<not> A" |
|
390 |
have "A \<longrightarrow> B" |
|
391 |
proof |
|
392 |
assume A |
|
393 |
with \<open>\<not> A\<close> show B by (rule contradiction) |
|
394 |
qed |
|
395 |
with * show A .. |
|
396 |
qed |
|
397 |
qed |
|
398 |
||
399 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
400 |
section \<open>Hilbert's choice operator (axiomatization)\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
401 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
402 |
axiomatization Eps :: "('a \<Rightarrow> o) \<Rightarrow> 'a" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
403 |
where someI: "P x \<Longrightarrow> P (Eps P)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
404 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
405 |
syntax "_Eps" :: "pttrn \<Rightarrow> o \<Rightarrow> 'a" ("(3SOME _./ _)" [0, 10] 10) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
406 |
translations "SOME x. P" \<rightleftharpoons> "CONST Eps (\<lambda>x. P)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
407 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
408 |
text \<open> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
409 |
\<^medskip> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
410 |
It follows a derivation of the classical law of tertium-non-datur by |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
411 |
means of Hilbert's choice operator (due to Berghofer, Beeson, Harrison, |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
412 |
based on a proof by Diaconescu). |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
413 |
\<^medskip> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
414 |
\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
415 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
416 |
theorem Diaconescu: "A \<or> \<not> A" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
417 |
proof - |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
418 |
let ?P = "\<lambda>x. (A \<and> x) \<or> \<not> x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
419 |
let ?Q = "\<lambda>x. (A \<and> \<not> x) \<or> x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
420 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
421 |
have a: "?P (Eps ?P)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
422 |
proof (rule someI) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
423 |
have "\<not> False" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
424 |
then show "?P False" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
425 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
426 |
have b: "?Q (Eps ?Q)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
427 |
proof (rule someI) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
428 |
have True .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
429 |
then show "?Q True" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
430 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
431 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
432 |
from a show ?thesis |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
433 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
434 |
assume "A \<and> Eps ?P" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
435 |
then have A .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
436 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
437 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
438 |
assume "\<not> Eps ?P" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
439 |
from b show ?thesis |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
440 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
441 |
assume "A \<and> \<not> Eps ?Q" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
442 |
then have A .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
443 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
444 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
445 |
assume "Eps ?Q" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
446 |
have neq: "?P \<noteq> ?Q" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
447 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
448 |
assume "?P = ?Q" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
449 |
then have "Eps ?P \<longleftrightarrow> Eps ?Q" by (rule arg_cong) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
450 |
also note \<open>Eps ?Q\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
451 |
finally have "Eps ?P" . |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
452 |
with \<open>\<not> Eps ?P\<close> show False by (rule contradiction) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
453 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
454 |
have "\<not> A" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
455 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
456 |
assume A |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
457 |
have "?P = ?Q" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
458 |
proof (rule ext) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
459 |
show "?P x \<longleftrightarrow> ?Q x" for x |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
460 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
461 |
assume "?P x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
462 |
then show "?Q x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
463 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
464 |
assume "\<not> x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
465 |
with \<open>A\<close> have "A \<and> \<not> x" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
466 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
467 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
468 |
assume "A \<and> x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
469 |
then have x .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
470 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
471 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
472 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
473 |
assume "?Q x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
474 |
then show "?P x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
475 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
476 |
assume "A \<and> \<not> x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
477 |
then have "\<not> x" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
478 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
479 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
480 |
assume x |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
481 |
with \<open>A\<close> have "A \<and> x" .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
482 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
483 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
484 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
485 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
486 |
with neq show False by (rule contradiction) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
487 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
488 |
then show ?thesis .. |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
489 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
490 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
491 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
492 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
493 |
text \<open> |
69593 | 494 |
This means, the hypothetical predicate \<^const>\<open>classical\<close> always holds |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
495 |
unconditionally (with all consequences). |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
496 |
\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
497 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
498 |
interpretation classical |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
499 |
proof (rule classical_if_cases) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
500 |
fix A C |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
501 |
assume *: "A \<Longrightarrow> C" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
502 |
and **: "\<not> A \<Longrightarrow> C" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
503 |
from Diaconescu [of A] show C |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
504 |
proof |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
505 |
assume A |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
506 |
then show C by (rule *) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
507 |
next |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
508 |
assume "\<not> A" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
509 |
then show C by (rule **) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
510 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
511 |
qed |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
512 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
513 |
thm classical |
64908 | 514 |
classical_contradiction |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
515 |
double_negation |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
516 |
tertium_non_datur |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
517 |
classical_cases |
64908 | 518 |
Peirce's_Law |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
519 |
|
12360 | 520 |
end |