author | wenzelm |
Wed, 22 Jan 2025 22:22:19 +0100 | |
changeset 81954 | 6f2bcdfa9a19 |
parent 81182 | fc5066122e68 |
child 82048 | 2ea9f9ed19c6 |
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 |
76987 | 13 |
Logic. The ``HOL'' logic that is given here resembles \<^cite>\<open>"Gordon:1985:HOL"\<close> and its predecessor \<^cite>\<open>"church40"\<close>, but the order of |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
14 |
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
|
15 |
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
|
16 |
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
|
17 |
Isabelle/Isar. |
59031 | 18 |
\<close> |
12360 | 19 |
|
20 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
21 |
section \<open>HOL syntax within Pure\<close> |
12360 | 22 |
|
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
23 |
class type |
36452 | 24 |
default_sort type |
12360 | 25 |
|
26 |
typedecl o |
|
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
27 |
instance o :: type .. |
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
41460
diff
changeset
|
28 |
instance "fun" :: (type, type) type .. |
12360 | 29 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
30 |
judgment Trueprop :: "o \<Rightarrow> prop" (\<open>_\<close> 5) |
12360 | 31 |
|
32 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
33 |
section \<open>Minimal logic (axiomatization)\<close> |
12360 | 34 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
35 |
axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25) |
61759 | 36 |
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" |
37 |
and impE [dest, trans]: "A \<longrightarrow> B \<Longrightarrow> A \<Longrightarrow> B" |
|
38 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
39 |
axiomatization All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10) |
61759 | 40 |
where allI [intro]: "(\<And>x. P x) \<Longrightarrow> \<forall>x. P x" |
41 |
and allE [dest]: "\<forall>x. P x \<Longrightarrow> P a" |
|
12360 | 42 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
43 |
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
|
44 |
by standard (fact impI, fact impE) |
12360 | 45 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
46 |
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
|
47 |
by standard (fact allI, fact allE) |
12360 | 48 |
|
49 |
||
59031 | 50 |
subsubsection \<open>Derived connectives\<close> |
12360 | 51 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
52 |
definition False :: o |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
53 |
where "False \<equiv> \<forall>A. A" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
54 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
55 |
lemma FalseE [elim]: |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
56 |
assumes "False" |
61759 | 57 |
shows A |
58 |
proof - |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
59 |
from \<open>False\<close> have "\<forall>A. A" by (simp only: False_def) |
61759 | 60 |
then show A .. |
61 |
qed |
|
62 |
||
63 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
64 |
definition True :: o |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
65 |
where "True \<equiv> False \<longrightarrow> False" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
66 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
67 |
lemma TrueI [intro]: True |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
68 |
unfolding True_def .. |
61759 | 69 |
|
70 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
71 |
definition not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
72 |
where "not \<equiv> \<lambda>A. A \<longrightarrow> False" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20523
diff
changeset
|
73 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
74 |
lemma notI [intro]: |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
75 |
assumes "A \<Longrightarrow> False" |
61759 | 76 |
shows "\<not> A" |
77 |
using assms unfolding not_def .. |
|
12360 | 78 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
79 |
lemma notE [elim]: |
61759 | 80 |
assumes "\<not> A" and A |
81 |
shows B |
|
82 |
proof - |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
83 |
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
|
84 |
from this and \<open>A\<close> have "False" .. |
23373 | 85 |
then show B .. |
12360 | 86 |
qed |
87 |
||
88 |
lemma notE': "A \<Longrightarrow> \<not> A \<Longrightarrow> B" |
|
89 |
by (rule notE) |
|
90 |
||
61759 | 91 |
lemmas contradiction = notE notE' \<comment> \<open>proof by contradiction in any order\<close> |
92 |
||
93 |
||
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
94 |
definition conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<and>\<close> 35) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
95 |
where "A \<and> B \<equiv> \<forall>C. (A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" |
12360 | 96 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
97 |
lemma conjI [intro]: |
61759 | 98 |
assumes A and B |
99 |
shows "A \<and> B" |
|
100 |
unfolding conj_def |
|
101 |
proof |
|
102 |
fix C |
|
103 |
show "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 104 |
proof |
61759 | 105 |
assume "A \<longrightarrow> B \<longrightarrow> C" |
106 |
also note \<open>A\<close> |
|
107 |
also note \<open>B\<close> |
|
108 |
finally show C . |
|
12360 | 109 |
qed |
110 |
qed |
|
111 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
112 |
lemma conjE [elim]: |
61759 | 113 |
assumes "A \<and> B" |
114 |
obtains A and B |
|
115 |
proof |
|
116 |
from \<open>A \<and> B\<close> have *: "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C" for C |
|
117 |
unfolding conj_def .. |
|
118 |
show A |
|
119 |
proof - |
|
120 |
note * [of A] |
|
12360 | 121 |
also have "A \<longrightarrow> B \<longrightarrow> A" |
122 |
proof |
|
123 |
assume A |
|
23373 | 124 |
then show "B \<longrightarrow> A" .. |
12360 | 125 |
qed |
61759 | 126 |
finally show ?thesis . |
127 |
qed |
|
128 |
show B |
|
129 |
proof - |
|
130 |
note * [of B] |
|
12360 | 131 |
also have "A \<longrightarrow> B \<longrightarrow> B" |
132 |
proof |
|
133 |
show "B \<longrightarrow> B" .. |
|
134 |
qed |
|
61759 | 135 |
finally show ?thesis . |
12360 | 136 |
qed |
137 |
qed |
|
138 |
||
61759 | 139 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
140 |
definition disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<or>\<close> 30) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
141 |
where "A \<or> B \<equiv> \<forall>C. (A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
61759 | 142 |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
143 |
lemma disjI1 [intro]: |
61759 | 144 |
assumes A |
145 |
shows "A \<or> B" |
|
146 |
unfolding disj_def |
|
147 |
proof |
|
148 |
fix C |
|
149 |
show "(A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 150 |
proof |
61759 | 151 |
assume "A \<longrightarrow> C" |
152 |
from this and \<open>A\<close> have C .. |
|
153 |
then show "(B \<longrightarrow> C) \<longrightarrow> C" .. |
|
154 |
qed |
|
155 |
qed |
|
156 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
157 |
lemma disjI2 [intro]: |
61759 | 158 |
assumes B |
159 |
shows "A \<or> B" |
|
160 |
unfolding disj_def |
|
161 |
proof |
|
162 |
fix C |
|
163 |
show "(A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C" |
|
164 |
proof |
|
165 |
show "(B \<longrightarrow> C) \<longrightarrow> C" |
|
12360 | 166 |
proof |
61759 | 167 |
assume "B \<longrightarrow> C" |
168 |
from this and \<open>B\<close> show C .. |
|
12360 | 169 |
qed |
170 |
qed |
|
171 |
qed |
|
172 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
173 |
lemma disjE [elim]: |
61759 | 174 |
assumes "A \<or> B" |
175 |
obtains (a) A | (b) B |
|
176 |
proof - |
|
177 |
from \<open>A \<or> B\<close> have "(A \<longrightarrow> thesis) \<longrightarrow> (B \<longrightarrow> thesis) \<longrightarrow> thesis" |
|
178 |
unfolding disj_def .. |
|
179 |
also have "A \<longrightarrow> thesis" |
|
12360 | 180 |
proof |
60769 | 181 |
assume A |
61759 | 182 |
then show thesis by (rule a) |
12360 | 183 |
qed |
61759 | 184 |
also have "B \<longrightarrow> thesis" |
12360 | 185 |
proof |
60769 | 186 |
assume B |
61759 | 187 |
then show thesis by (rule b) |
12360 | 188 |
qed |
61759 | 189 |
finally show thesis . |
12360 | 190 |
qed |
191 |
||
61759 | 192 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
193 |
definition Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10) |
61759 | 194 |
where "\<exists>x. P x \<equiv> \<forall>C. (\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C" |
195 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
196 |
lemma exI [intro]: "P a \<Longrightarrow> \<exists>x. P x" |
61759 | 197 |
unfolding Ex_def |
198 |
proof |
|
199 |
fix C |
|
12360 | 200 |
assume "P a" |
61759 | 201 |
show "(\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C" |
12360 | 202 |
proof |
61759 | 203 |
assume "\<forall>x. P x \<longrightarrow> C" |
204 |
then have "P a \<longrightarrow> C" .. |
|
205 |
from this and \<open>P a\<close> show C .. |
|
12360 | 206 |
qed |
207 |
qed |
|
208 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
209 |
lemma exE [elim]: |
61759 | 210 |
assumes "\<exists>x. P x" |
211 |
obtains (that) x where "P x" |
|
212 |
proof - |
|
213 |
from \<open>\<exists>x. P x\<close> have "(\<forall>x. P x \<longrightarrow> thesis) \<longrightarrow> thesis" |
|
214 |
unfolding Ex_def .. |
|
215 |
also have "\<forall>x. P x \<longrightarrow> thesis" |
|
12360 | 216 |
proof |
61759 | 217 |
fix x |
218 |
show "P x \<longrightarrow> thesis" |
|
12360 | 219 |
proof |
220 |
assume "P x" |
|
61759 | 221 |
then show thesis by (rule that) |
12360 | 222 |
qed |
223 |
qed |
|
61759 | 224 |
finally show thesis . |
12360 | 225 |
qed |
226 |
||
227 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
228 |
subsubsection \<open>Extensional equality\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
229 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
230 |
axiomatization equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl \<open>=\<close> 50) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
231 |
where refl [intro]: "x = x" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
232 |
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
|
233 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
234 |
abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl \<open>\<noteq>\<close> 50) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
235 |
where "x \<noteq> y \<equiv> \<not> (x = y)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
236 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
237 |
abbreviation iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longleftrightarrow>\<close> 25) |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
238 |
where "A \<longleftrightarrow> B \<equiv> A = B" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
239 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
240 |
axiomatization |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
241 |
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
|
242 |
and iff [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> (B \<Longrightarrow> A) \<Longrightarrow> A \<longleftrightarrow> B" |
71831 | 243 |
for f g :: "'a \<Rightarrow> 'b" |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
244 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
245 |
lemma sym [sym]: "y = x" if "x = y" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
246 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
247 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
248 |
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
|
249 |
by (rule subst) (rule sym) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
250 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
251 |
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
|
252 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
253 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
254 |
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
|
255 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
256 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
257 |
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
|
258 |
using that by (rule subst) (rule refl) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
259 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
260 |
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
|
261 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
262 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
263 |
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
|
264 |
by (rule subst) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
265 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
266 |
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
|
267 |
by (rule subst) (rule sym) |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
268 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
269 |
|
61936 | 270 |
subsection \<open>Cantor's Theorem\<close> |
271 |
||
272 |
text \<open> |
|
273 |
Cantor's Theorem states that there is no surjection from a set to its |
|
274 |
powerset. The subsequent formulation uses elementary \<open>\<lambda>\<close>-calculus and |
|
275 |
predicate logic, with standard introduction and elimination rules. |
|
276 |
\<close> |
|
277 |
||
278 |
lemma iff_contradiction: |
|
279 |
assumes *: "\<not> A \<longleftrightarrow> A" |
|
280 |
shows C |
|
281 |
proof (rule notE) |
|
282 |
show "\<not> A" |
|
283 |
proof |
|
284 |
assume A |
|
285 |
with * have "\<not> A" .. |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
286 |
from this and \<open>A\<close> show False .. |
61936 | 287 |
qed |
288 |
with * show A .. |
|
289 |
qed |
|
290 |
||
62038 | 291 |
theorem Cantor: "\<not> (\<exists>f :: 'a \<Rightarrow> 'a \<Rightarrow> o. \<forall>A. \<exists>x. A = f x)" |
61936 | 292 |
proof |
293 |
assume "\<exists>f :: 'a \<Rightarrow> 'a \<Rightarrow> o. \<forall>A. \<exists>x. A = f x" |
|
294 |
then obtain f :: "'a \<Rightarrow> 'a \<Rightarrow> o" where *: "\<forall>A. \<exists>x. A = f x" .. |
|
295 |
let ?D = "\<lambda>x. \<not> f x x" |
|
296 |
from * have "\<exists>x. ?D = f x" .. |
|
297 |
then obtain a where "?D = f a" .. |
|
298 |
then have "?D a \<longleftrightarrow> f a a" using refl by (rule subst) |
|
62266 | 299 |
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
|
300 |
then show False by (rule iff_contradiction) |
61936 | 301 |
qed |
302 |
||
303 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
304 |
subsection \<open>Characterization of Classical Logic\<close> |
12360 | 305 |
|
61759 | 306 |
text \<open> |
307 |
The subsequent rules of classical reasoning are all equivalent. |
|
308 |
\<close> |
|
309 |
||
12360 | 310 |
locale classical = |
311 |
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
|
312 |
\<comment> \<open>predicate definition and hypothetical context\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
313 |
begin |
12360 | 314 |
|
64908 | 315 |
lemma classical_contradiction: |
316 |
assumes "\<not> A \<Longrightarrow> False" |
|
317 |
shows A |
|
318 |
proof (rule classical) |
|
319 |
assume "\<not> A" |
|
320 |
then have False by (rule assms) |
|
321 |
then show A .. |
|
12360 | 322 |
qed |
323 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
324 |
lemma double_negation: |
61759 | 325 |
assumes "\<not> \<not> A" |
326 |
shows A |
|
64908 | 327 |
proof (rule classical_contradiction) |
61759 | 328 |
assume "\<not> A" |
64908 | 329 |
with \<open>\<not> \<not> A\<close> show False by (rule contradiction) |
12360 | 330 |
qed |
331 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
332 |
lemma tertium_non_datur: "A \<or> \<not> A" |
12360 | 333 |
proof (rule double_negation) |
334 |
show "\<not> \<not> (A \<or> \<not> A)" |
|
335 |
proof |
|
336 |
assume "\<not> (A \<or> \<not> A)" |
|
337 |
have "\<not> A" |
|
338 |
proof |
|
23373 | 339 |
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
|
340 |
with \<open>\<not> (A \<or> \<not> A)\<close> show False by (rule contradiction) |
12360 | 341 |
qed |
23373 | 342 |
then have "A \<or> \<not> A" .. |
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
343 |
with \<open>\<not> (A \<or> \<not> A)\<close> show False by (rule contradiction) |
12360 | 344 |
qed |
345 |
qed |
|
346 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
347 |
lemma classical_cases: |
61759 | 348 |
obtains A | "\<not> A" |
349 |
using tertium_non_datur |
|
350 |
proof |
|
351 |
assume A |
|
352 |
then show thesis .. |
|
353 |
next |
|
354 |
assume "\<not> A" |
|
355 |
then show thesis .. |
|
356 |
qed |
|
357 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
358 |
end |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
359 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
360 |
lemma classical_if_cases: classical |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
361 |
if cases: "\<And>A C. (A \<Longrightarrow> C) \<Longrightarrow> (\<not> A \<Longrightarrow> C) \<Longrightarrow> C" |
61759 | 362 |
proof |
363 |
fix A |
|
364 |
assume *: "\<not> A \<Longrightarrow> A" |
|
365 |
show A |
|
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
366 |
proof (rule cases) |
12360 | 367 |
assume A |
61759 | 368 |
then show A . |
12360 | 369 |
next |
370 |
assume "\<not> A" |
|
61759 | 371 |
then show A by (rule *) |
12573 | 372 |
qed |
373 |
qed |
|
374 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
375 |
|
64908 | 376 |
section \<open>Peirce's Law\<close> |
377 |
||
378 |
text \<open> |
|
379 |
Peirce's Law is another characterization of classical reasoning. Its |
|
380 |
statement only requires implication. |
|
381 |
\<close> |
|
382 |
||
383 |
theorem (in classical) Peirce's_Law: "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A" |
|
384 |
proof |
|
385 |
assume *: "(A \<longrightarrow> B) \<longrightarrow> A" |
|
386 |
show A |
|
387 |
proof (rule classical) |
|
388 |
assume "\<not> A" |
|
389 |
have "A \<longrightarrow> B" |
|
390 |
proof |
|
391 |
assume A |
|
392 |
with \<open>\<not> A\<close> show B by (rule contradiction) |
|
393 |
qed |
|
394 |
with * show A .. |
|
395 |
qed |
|
396 |
qed |
|
397 |
||
398 |
||
64907
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
399 |
section \<open>Hilbert's choice operator (axiomatization)\<close> |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
400 |
|
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
401 |
axiomatization Eps :: "('a \<Rightarrow> o) \<Rightarrow> 'a" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
402 |
where someI: "P x \<Longrightarrow> P (Eps P)" |
354bfbb27fbb
misc tuning and updates according to Curry-Club Dec-2016;
wenzelm
parents:
64475
diff
changeset
|
403 |
|
81182 | 404 |
syntax "_Eps" :: "pttrn \<Rightarrow> o \<Rightarrow> 'a" (\<open>(\<open>indent=3 notation=\<open>binder SOME\<close>\<close>SOME _./ _)\<close> [0, 10] 10) |
80768 | 405 |
syntax_consts "_Eps" \<rightleftharpoons> Eps |
64907
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 |