berghofe@12450
|
1 |
(* Title: HOL/ex/Intuitionistic.thy
|
berghofe@12450
|
2 |
ID: $Id$
|
berghofe@12450
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
berghofe@12450
|
4 |
Copyright 1991 University of Cambridge
|
berghofe@12450
|
5 |
|
berghofe@12450
|
6 |
Higher-Order Logic: Intuitionistic predicate calculus problems
|
berghofe@12450
|
7 |
|
berghofe@12450
|
8 |
Taken from FOL/ex/int.ML
|
berghofe@12450
|
9 |
*)
|
berghofe@12450
|
10 |
|
berghofe@12450
|
11 |
theory Intuitionistic = Main:
|
berghofe@12450
|
12 |
|
berghofe@12450
|
13 |
|
berghofe@12450
|
14 |
(*Metatheorem (for PROPOSITIONAL formulae...):
|
berghofe@12450
|
15 |
P is classically provable iff ~~P is intuitionistically provable.
|
berghofe@12450
|
16 |
Therefore ~P is classically provable iff it is intuitionistically provable.
|
berghofe@12450
|
17 |
|
berghofe@12450
|
18 |
Proof: Let Q be the conjuction of the propositions A|~A, one for each atom A
|
berghofe@12450
|
19 |
in P. Now ~~Q is intuitionistically provable because ~~(A|~A) is and because
|
berghofe@12450
|
20 |
~~ distributes over &. If P is provable classically, then clearly Q-->P is
|
berghofe@12450
|
21 |
provable intuitionistically, so ~~(Q-->P) is also provable intuitionistically.
|
berghofe@12450
|
22 |
The latter is intuitionistically equivalent to ~~Q-->~~P, hence to ~~P, since
|
berghofe@12450
|
23 |
~~Q is intuitionistically provable. Finally, if P is a negation then ~~P is
|
berghofe@12450
|
24 |
intuitionstically equivalent to P. [Andy Pitts] *)
|
berghofe@12450
|
25 |
|
berghofe@12450
|
26 |
lemma "(~~(P&Q)) = ((~~P) & (~~Q))"
|
berghofe@12450
|
27 |
by rules
|
berghofe@12450
|
28 |
|
berghofe@12450
|
29 |
lemma "~~ ((~P --> Q) --> (~P --> ~Q) --> P)"
|
berghofe@12450
|
30 |
by rules
|
berghofe@12450
|
31 |
|
berghofe@12450
|
32 |
(* ~~ does NOT distribute over | *)
|
berghofe@12450
|
33 |
|
berghofe@12450
|
34 |
lemma "(~~(P-->Q)) = (~~P --> ~~Q)"
|
berghofe@12450
|
35 |
by rules
|
berghofe@12450
|
36 |
|
berghofe@12450
|
37 |
lemma "(~~~P) = (~P)"
|
berghofe@12450
|
38 |
by rules
|
berghofe@12450
|
39 |
|
berghofe@12450
|
40 |
lemma "~~((P --> Q | R) --> (P-->Q) | (P-->R))"
|
berghofe@12450
|
41 |
by rules
|
berghofe@12450
|
42 |
|
berghofe@12450
|
43 |
lemma "(P=Q) = (Q=P)"
|
berghofe@12450
|
44 |
by rules
|
berghofe@12450
|
45 |
|
berghofe@12450
|
46 |
lemma "((P --> (Q | (Q-->R))) --> R) --> R"
|
berghofe@12450
|
47 |
by rules
|
berghofe@12450
|
48 |
|
berghofe@12450
|
49 |
lemma "(((G-->A) --> J) --> D --> E) --> (((H-->B)-->I)-->C-->J)
|
berghofe@12450
|
50 |
--> (A-->H) --> F --> G --> (((C-->B)-->I)-->D)-->(A-->C)
|
berghofe@12450
|
51 |
--> (((F-->A)-->B) --> I) --> E"
|
berghofe@12450
|
52 |
by rules
|
berghofe@12450
|
53 |
|
berghofe@12450
|
54 |
|
berghofe@12450
|
55 |
(* Lemmas for the propositional double-negation translation *)
|
berghofe@12450
|
56 |
|
berghofe@12450
|
57 |
lemma "P --> ~~P"
|
berghofe@12450
|
58 |
by rules
|
berghofe@12450
|
59 |
|
berghofe@12450
|
60 |
lemma "~~(~~P --> P)"
|
berghofe@12450
|
61 |
by rules
|
berghofe@12450
|
62 |
|
berghofe@12450
|
63 |
lemma "~~P & ~~(P --> Q) --> ~~Q"
|
berghofe@12450
|
64 |
by rules
|
berghofe@12450
|
65 |
|
berghofe@12450
|
66 |
|
berghofe@12450
|
67 |
(* de Bruijn formulae *)
|
berghofe@12450
|
68 |
|
berghofe@12450
|
69 |
(*de Bruijn formula with three predicates*)
|
berghofe@12450
|
70 |
lemma "((P=Q) --> P&Q&R) &
|
berghofe@12450
|
71 |
((Q=R) --> P&Q&R) &
|
berghofe@12450
|
72 |
((R=P) --> P&Q&R) --> P&Q&R"
|
berghofe@12450
|
73 |
by rules
|
berghofe@12450
|
74 |
|
berghofe@12450
|
75 |
(*de Bruijn formula with five predicates*)
|
berghofe@12450
|
76 |
lemma "((P=Q) --> P&Q&R&S&T) &
|
berghofe@12450
|
77 |
((Q=R) --> P&Q&R&S&T) &
|
berghofe@12450
|
78 |
((R=S) --> P&Q&R&S&T) &
|
berghofe@12450
|
79 |
((S=T) --> P&Q&R&S&T) &
|
berghofe@12450
|
80 |
((T=P) --> P&Q&R&S&T) --> P&Q&R&S&T"
|
berghofe@12450
|
81 |
by rules
|
berghofe@12450
|
82 |
|
berghofe@12450
|
83 |
|
berghofe@12450
|
84 |
(*** Problems from Sahlin, Franzen and Haridi,
|
berghofe@12450
|
85 |
An Intuitionistic Predicate Logic Theorem Prover.
|
berghofe@12450
|
86 |
J. Logic and Comp. 2 (5), October 1992, 619-656.
|
berghofe@12450
|
87 |
***)
|
berghofe@12450
|
88 |
|
berghofe@12450
|
89 |
(*Problem 1.1*)
|
berghofe@12450
|
90 |
lemma "(ALL x. EX y. ALL z. p(x) & q(y) & r(z)) =
|
berghofe@12450
|
91 |
(ALL z. EX y. ALL x. p(x) & q(y) & r(z))"
|
berghofe@12450
|
92 |
by (rules del: allE elim 2: allE')
|
berghofe@12450
|
93 |
|
berghofe@12450
|
94 |
(*Problem 3.1*)
|
berghofe@12450
|
95 |
lemma "~ (EX x. ALL y. p y x = (~ p x x))"
|
berghofe@12450
|
96 |
by rules
|
berghofe@12450
|
97 |
|
berghofe@12450
|
98 |
|
berghofe@12450
|
99 |
(* Intuitionistic FOL: propositional problems based on Pelletier. *)
|
berghofe@12450
|
100 |
|
berghofe@12450
|
101 |
(* Problem ~~1 *)
|
berghofe@12450
|
102 |
lemma "~~((P-->Q) = (~Q --> ~P))"
|
berghofe@12450
|
103 |
by rules
|
berghofe@12450
|
104 |
|
berghofe@12450
|
105 |
(* Problem ~~2 *)
|
berghofe@12450
|
106 |
lemma "~~(~~P = P)"
|
berghofe@12450
|
107 |
by rules
|
berghofe@12450
|
108 |
|
berghofe@12450
|
109 |
(* Problem 3 *)
|
berghofe@12450
|
110 |
lemma "~(P-->Q) --> (Q-->P)"
|
berghofe@12450
|
111 |
by rules
|
berghofe@12450
|
112 |
|
berghofe@12450
|
113 |
(* Problem ~~4 *)
|
berghofe@12450
|
114 |
lemma "~~((~P-->Q) = (~Q --> P))"
|
berghofe@12450
|
115 |
by rules
|
berghofe@12450
|
116 |
|
berghofe@12450
|
117 |
(* Problem ~~5 *)
|
berghofe@12450
|
118 |
lemma "~~((P|Q-->P|R) --> P|(Q-->R))"
|
berghofe@12450
|
119 |
by rules
|
berghofe@12450
|
120 |
|
berghofe@12450
|
121 |
(* Problem ~~6 *)
|
berghofe@12450
|
122 |
lemma "~~(P | ~P)"
|
berghofe@12450
|
123 |
by rules
|
berghofe@12450
|
124 |
|
berghofe@12450
|
125 |
(* Problem ~~7 *)
|
berghofe@12450
|
126 |
lemma "~~(P | ~~~P)"
|
berghofe@12450
|
127 |
by rules
|
berghofe@12450
|
128 |
|
berghofe@12450
|
129 |
(* Problem ~~8. Peirce's law *)
|
berghofe@12450
|
130 |
lemma "~~(((P-->Q) --> P) --> P)"
|
berghofe@12450
|
131 |
by rules
|
berghofe@12450
|
132 |
|
berghofe@12450
|
133 |
(* Problem 9 *)
|
berghofe@12450
|
134 |
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
|
berghofe@12450
|
135 |
by rules
|
berghofe@12450
|
136 |
|
berghofe@12450
|
137 |
(* Problem 10 *)
|
berghofe@12450
|
138 |
lemma "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P=Q)"
|
berghofe@12450
|
139 |
by rules
|
berghofe@12450
|
140 |
|
berghofe@12450
|
141 |
(* 11. Proved in each direction (incorrectly, says Pelletier!!) *)
|
berghofe@12450
|
142 |
lemma "P=P"
|
berghofe@12450
|
143 |
by rules
|
berghofe@12450
|
144 |
|
berghofe@12450
|
145 |
(* Problem ~~12. Dijkstra's law *)
|
berghofe@12450
|
146 |
lemma "~~(((P = Q) = R) = (P = (Q = R)))"
|
berghofe@12450
|
147 |
by rules
|
berghofe@12450
|
148 |
|
berghofe@12450
|
149 |
lemma "((P = Q) = R) --> ~~(P = (Q = R))"
|
berghofe@12450
|
150 |
by rules
|
berghofe@12450
|
151 |
|
berghofe@12450
|
152 |
(* Problem 13. Distributive law *)
|
berghofe@12450
|
153 |
lemma "(P | (Q & R)) = ((P | Q) & (P | R))"
|
berghofe@12450
|
154 |
by rules
|
berghofe@12450
|
155 |
|
berghofe@12450
|
156 |
(* Problem ~~14 *)
|
berghofe@12450
|
157 |
lemma "~~((P = Q) = ((Q | ~P) & (~Q|P)))"
|
berghofe@12450
|
158 |
by rules
|
berghofe@12450
|
159 |
|
berghofe@12450
|
160 |
(* Problem ~~15 *)
|
berghofe@12450
|
161 |
lemma "~~((P --> Q) = (~P | Q))"
|
berghofe@12450
|
162 |
by rules
|
berghofe@12450
|
163 |
|
berghofe@12450
|
164 |
(* Problem ~~16 *)
|
berghofe@12450
|
165 |
lemma "~~((P-->Q) | (Q-->P))"
|
berghofe@12450
|
166 |
by rules
|
berghofe@12450
|
167 |
|
berghofe@12450
|
168 |
(* Problem ~~17 *)
|
berghofe@12450
|
169 |
lemma "~~(((P & (Q-->R))-->S) = ((~P | Q | S) & (~P | ~R | S)))"
|
berghofe@12450
|
170 |
oops
|
berghofe@12450
|
171 |
|
berghofe@12450
|
172 |
(*Dijkstra's "Golden Rule"*)
|
berghofe@12450
|
173 |
lemma "(P&Q) = (P = (Q = (P|Q)))"
|
berghofe@12450
|
174 |
by rules
|
berghofe@12450
|
175 |
|
berghofe@12450
|
176 |
|
berghofe@12450
|
177 |
(****Examples with quantifiers****)
|
berghofe@12450
|
178 |
|
berghofe@12450
|
179 |
(* The converse is classical in the following implications... *)
|
berghofe@12450
|
180 |
|
berghofe@12450
|
181 |
lemma "(EX x. P(x)-->Q) --> (ALL x. P(x)) --> Q"
|
berghofe@12450
|
182 |
by rules
|
berghofe@12450
|
183 |
|
berghofe@12450
|
184 |
lemma "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
|
berghofe@12450
|
185 |
by rules
|
berghofe@12450
|
186 |
|
berghofe@12450
|
187 |
lemma "((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"
|
berghofe@12450
|
188 |
by rules
|
berghofe@12450
|
189 |
|
berghofe@12450
|
190 |
lemma "(ALL x. P(x)) | Q --> (ALL x. P(x) | Q)"
|
berghofe@12450
|
191 |
by rules
|
berghofe@12450
|
192 |
|
berghofe@12450
|
193 |
lemma "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
|
berghofe@12450
|
194 |
by rules
|
berghofe@12450
|
195 |
|
berghofe@12450
|
196 |
|
berghofe@12450
|
197 |
(* Hard examples with quantifiers *)
|
berghofe@12450
|
198 |
|
berghofe@12450
|
199 |
(*The ones that have not been proved are not known to be valid!
|
berghofe@12450
|
200 |
Some will require quantifier duplication -- not currently available*)
|
berghofe@12450
|
201 |
|
berghofe@12450
|
202 |
(* Problem ~~19 *)
|
berghofe@12450
|
203 |
lemma "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"
|
berghofe@12450
|
204 |
by rules
|
berghofe@12450
|
205 |
|
berghofe@12450
|
206 |
(* Problem 20 *)
|
berghofe@12450
|
207 |
lemma "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))
|
berghofe@12450
|
208 |
--> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
|
berghofe@12450
|
209 |
by rules
|
berghofe@12450
|
210 |
|
berghofe@12450
|
211 |
(* Problem 21 *)
|
berghofe@12450
|
212 |
lemma "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P=Q(x))"
|
berghofe@12450
|
213 |
by rules
|
berghofe@12450
|
214 |
|
berghofe@12450
|
215 |
(* Problem 22 *)
|
berghofe@12450
|
216 |
lemma "(ALL x. P = Q(x)) --> (P = (ALL x. Q(x)))"
|
berghofe@12450
|
217 |
by rules
|
berghofe@12450
|
218 |
|
berghofe@12450
|
219 |
(* Problem ~~23 *)
|
berghofe@12450
|
220 |
lemma "~~ ((ALL x. P | Q(x)) = (P | (ALL x. Q(x))))"
|
berghofe@12450
|
221 |
by rules
|
berghofe@12450
|
222 |
|
berghofe@12450
|
223 |
(* Problem 25 *)
|
berghofe@12450
|
224 |
lemma "(EX x. P(x)) &
|
berghofe@12450
|
225 |
(ALL x. L(x) --> ~ (M(x) & R(x))) &
|
berghofe@12450
|
226 |
(ALL x. P(x) --> (M(x) & L(x))) &
|
berghofe@12450
|
227 |
((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))
|
berghofe@12450
|
228 |
--> (EX x. Q(x)&P(x))"
|
berghofe@12450
|
229 |
by rules
|
berghofe@12450
|
230 |
|
berghofe@12450
|
231 |
(* Problem 27 *)
|
berghofe@12450
|
232 |
lemma "(EX x. P(x) & ~Q(x)) &
|
berghofe@12450
|
233 |
(ALL x. P(x) --> R(x)) &
|
berghofe@12450
|
234 |
(ALL x. M(x) & L(x) --> P(x)) &
|
berghofe@12450
|
235 |
((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))
|
berghofe@12450
|
236 |
--> (ALL x. M(x) --> ~L(x))"
|
berghofe@12450
|
237 |
by rules
|
berghofe@12450
|
238 |
|
berghofe@12450
|
239 |
(* Problem ~~28. AMENDED *)
|
berghofe@12450
|
240 |
lemma "(ALL x. P(x) --> (ALL x. Q(x))) &
|
berghofe@12450
|
241 |
(~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) &
|
berghofe@12450
|
242 |
(~~(EX x. S(x)) --> (ALL x. L(x) --> M(x)))
|
berghofe@12450
|
243 |
--> (ALL x. P(x) & L(x) --> M(x))"
|
berghofe@12450
|
244 |
by rules
|
berghofe@12450
|
245 |
|
berghofe@12450
|
246 |
(* Problem 29. Essentially the same as Principia Mathematica *11.71 *)
|
berghofe@12450
|
247 |
lemma "(((EX x. P(x)) & (EX y. Q(y))) -->
|
berghofe@12450
|
248 |
(((ALL x. (P(x) --> R(x))) & (ALL y. (Q(y) --> S(y)))) =
|
berghofe@12450
|
249 |
(ALL x y. ((P(x) & Q(y)) --> (R(x) & S(y))))))"
|
berghofe@12450
|
250 |
by rules
|
berghofe@12450
|
251 |
|
berghofe@12450
|
252 |
(* Problem ~~30 *)
|
berghofe@12450
|
253 |
lemma "(ALL x. (P(x) | Q(x)) --> ~ R(x)) &
|
berghofe@12450
|
254 |
(ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
|
berghofe@12450
|
255 |
--> (ALL x. ~~S(x))"
|
berghofe@12450
|
256 |
by rules
|
berghofe@12450
|
257 |
|
berghofe@12450
|
258 |
(* Problem 31 *)
|
berghofe@12450
|
259 |
lemma "~(EX x. P(x) & (Q(x) | R(x))) &
|
berghofe@12450
|
260 |
(EX x. L(x) & P(x)) &
|
berghofe@12450
|
261 |
(ALL x. ~ R(x) --> M(x))
|
berghofe@12450
|
262 |
--> (EX x. L(x) & M(x))"
|
berghofe@12450
|
263 |
by rules
|
berghofe@12450
|
264 |
|
berghofe@12450
|
265 |
(* Problem 32 *)
|
berghofe@12450
|
266 |
lemma "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) &
|
berghofe@12450
|
267 |
(ALL x. S(x) & R(x) --> L(x)) &
|
berghofe@12450
|
268 |
(ALL x. M(x) --> R(x))
|
berghofe@12450
|
269 |
--> (ALL x. P(x) & M(x) --> L(x))"
|
berghofe@12450
|
270 |
by rules
|
berghofe@12450
|
271 |
|
berghofe@12450
|
272 |
(* Problem ~~33 *)
|
berghofe@12450
|
273 |
lemma "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c))) =
|
berghofe@12450
|
274 |
(ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))"
|
berghofe@12450
|
275 |
oops
|
berghofe@12450
|
276 |
|
berghofe@12450
|
277 |
(* Problem 36 *)
|
berghofe@12450
|
278 |
lemma
|
berghofe@12450
|
279 |
"(ALL x. EX y. J x y) &
|
berghofe@12450
|
280 |
(ALL x. EX y. G x y) &
|
berghofe@12450
|
281 |
(ALL x y. J x y | G x y --> (ALL z. J y z | G y z --> H x z))
|
berghofe@12450
|
282 |
--> (ALL x. EX y. H x y)"
|
berghofe@12450
|
283 |
by rules
|
berghofe@12450
|
284 |
|
berghofe@12450
|
285 |
(* Problem 39 *)
|
berghofe@12450
|
286 |
lemma "~ (EX x. ALL y. F y x = (~F y y))"
|
berghofe@12450
|
287 |
by rules
|
berghofe@12450
|
288 |
|
berghofe@12450
|
289 |
(* Problem 40. AMENDED *)
|
berghofe@12450
|
290 |
lemma "(EX y. ALL x. F x y = F x x) -->
|
berghofe@12450
|
291 |
~(ALL x. EX y. ALL z. F z y = (~ F z x))"
|
berghofe@12450
|
292 |
by rules
|
berghofe@12450
|
293 |
|
berghofe@12450
|
294 |
(* Problem 44 *)
|
berghofe@12450
|
295 |
lemma "(ALL x. f(x) -->
|
berghofe@12450
|
296 |
(EX y. g(y) & h x y & (EX y. g(y) & ~ h x y))) &
|
berghofe@12450
|
297 |
(EX x. j(x) & (ALL y. g(y) --> h x y))
|
berghofe@12450
|
298 |
--> (EX x. j(x) & ~f(x))"
|
berghofe@12450
|
299 |
by rules
|
berghofe@12450
|
300 |
|
berghofe@12450
|
301 |
(* Problem 48 *)
|
berghofe@12450
|
302 |
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
|
berghofe@12450
|
303 |
by rules
|
berghofe@12450
|
304 |
|
berghofe@12450
|
305 |
(* Problem 51 *)
|
berghofe@12450
|
306 |
lemma "((EX z w. (ALL x y. (P x y = ((x = z) & (y = w))))) -->
|
berghofe@12450
|
307 |
(EX z. (ALL x. (EX w. ((ALL y. (P x y = (y = w))) = (x = z))))))"
|
berghofe@12450
|
308 |
by rules
|
berghofe@12450
|
309 |
|
berghofe@12450
|
310 |
(* Problem 52 *)
|
berghofe@12450
|
311 |
(*Almost the same as 51. *)
|
berghofe@12450
|
312 |
lemma "((EX z w. (ALL x y. (P x y = ((x = z) & (y = w))))) -->
|
berghofe@12450
|
313 |
(EX w. (ALL y. (EX z. ((ALL x. (P x y = (x = z))) = (y = w))))))"
|
berghofe@12450
|
314 |
by rules
|
berghofe@12450
|
315 |
|
berghofe@12450
|
316 |
(* Problem 56 *)
|
berghofe@12450
|
317 |
lemma "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) = (ALL x. P(x) --> P(f(x)))"
|
berghofe@12450
|
318 |
by rules
|
berghofe@12450
|
319 |
|
berghofe@12450
|
320 |
(* Problem 57 *)
|
berghofe@12450
|
321 |
lemma "P (f a b) (f b c) & P (f b c) (f a c) &
|
berghofe@12450
|
322 |
(ALL x y z. P x y & P y z --> P x z) --> P (f a b) (f a c)"
|
berghofe@12450
|
323 |
by rules
|
berghofe@12450
|
324 |
|
berghofe@12450
|
325 |
(* Problem 60 *)
|
berghofe@12450
|
326 |
lemma "ALL x. P x (f x) = (EX y. (ALL z. P z y --> P z (f x)) & P x y)"
|
berghofe@12450
|
327 |
by rules
|
berghofe@12450
|
328 |
|
berghofe@12450
|
329 |
end
|