31974
|
1 |
(* Title: FOL/ex/Intuitionistic.thy
|
14239
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
3 |
Copyright 1991 University of Cambridge
|
|
4 |
*)
|
|
5 |
|
31974
|
6 |
header {* Intuitionistic First-Order Logic *}
|
14239
|
7 |
|
31974
|
8 |
theory Intuitionistic
|
|
9 |
imports IFOL
|
|
10 |
begin
|
14239
|
11 |
|
|
12 |
(*
|
|
13 |
Single-step ML commands:
|
|
14 |
by (IntPr.step_tac 1)
|
|
15 |
by (biresolve_tac safe_brls 1);
|
|
16 |
by (biresolve_tac haz_brls 1);
|
|
17 |
by (assume_tac 1);
|
|
18 |
by (IntPr.safe_tac 1);
|
|
19 |
by (IntPr.mp_tac 1);
|
|
20 |
by (IntPr.fast_tac 1);
|
|
21 |
*)
|
|
22 |
|
|
23 |
|
|
24 |
text{*Metatheorem (for \emph{propositional} formulae):
|
|
25 |
$P$ is classically provable iff $\neg\neg P$ is intuitionistically provable.
|
|
26 |
Therefore $\neg P$ is classically provable iff it is intuitionistically
|
|
27 |
provable.
|
|
28 |
|
|
29 |
Proof: Let $Q$ be the conjuction of the propositions $A\vee\neg A$, one for
|
|
30 |
each atom $A$ in $P$. Now $\neg\neg Q$ is intuitionistically provable because
|
|
31 |
$\neg\neg(A\vee\neg A)$ is and because double-negation distributes over
|
|
32 |
conjunction. If $P$ is provable classically, then clearly $Q\rightarrow P$ is
|
|
33 |
provable intuitionistically, so $\neg\neg(Q\rightarrow P)$ is also provable
|
|
34 |
intuitionistically. The latter is intuitionistically equivalent to $\neg\neg
|
|
35 |
Q\rightarrow\neg\neg P$, hence to $\neg\neg P$, since $\neg\neg Q$ is
|
|
36 |
intuitionistically provable. Finally, if $P$ is a negation then $\neg\neg P$
|
|
37 |
is intuitionstically equivalent to $P$. [Andy Pitts] *}
|
|
38 |
|
|
39 |
lemma "~~(P&Q) <-> ~~P & ~~Q"
|
|
40 |
by (tactic{*IntPr.fast_tac 1*})
|
|
41 |
|
|
42 |
lemma "~~ ((~P --> Q) --> (~P --> ~Q) --> P)"
|
|
43 |
by (tactic{*IntPr.fast_tac 1*})
|
|
44 |
|
|
45 |
text{*Double-negation does NOT distribute over disjunction*}
|
|
46 |
|
|
47 |
lemma "~~(P-->Q) <-> (~~P --> ~~Q)"
|
|
48 |
by (tactic{*IntPr.fast_tac 1*})
|
|
49 |
|
|
50 |
lemma "~~~P <-> ~P"
|
|
51 |
by (tactic{*IntPr.fast_tac 1*})
|
|
52 |
|
|
53 |
lemma "~~((P --> Q | R) --> (P-->Q) | (P-->R))"
|
|
54 |
by (tactic{*IntPr.fast_tac 1*})
|
|
55 |
|
|
56 |
lemma "(P<->Q) <-> (Q<->P)"
|
|
57 |
by (tactic{*IntPr.fast_tac 1*})
|
|
58 |
|
|
59 |
lemma "((P --> (Q | (Q-->R))) --> R) --> R"
|
|
60 |
by (tactic{*IntPr.fast_tac 1*})
|
|
61 |
|
|
62 |
lemma "(((G-->A) --> J) --> D --> E) --> (((H-->B)-->I)-->C-->J)
|
|
63 |
--> (A-->H) --> F --> G --> (((C-->B)-->I)-->D)-->(A-->C)
|
|
64 |
--> (((F-->A)-->B) --> I) --> E"
|
|
65 |
by (tactic{*IntPr.fast_tac 1*})
|
|
66 |
|
|
67 |
|
|
68 |
text{*Lemmas for the propositional double-negation translation*}
|
|
69 |
|
|
70 |
lemma "P --> ~~P"
|
|
71 |
by (tactic{*IntPr.fast_tac 1*})
|
|
72 |
|
|
73 |
lemma "~~(~~P --> P)"
|
|
74 |
by (tactic{*IntPr.fast_tac 1*})
|
|
75 |
|
|
76 |
lemma "~~P & ~~(P --> Q) --> ~~Q"
|
|
77 |
by (tactic{*IntPr.fast_tac 1*})
|
|
78 |
|
|
79 |
|
|
80 |
text{*The following are classically but not constructively valid.
|
|
81 |
The attempt to prove them terminates quickly!*}
|
|
82 |
lemma "((P-->Q) --> P) --> P"
|
|
83 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
84 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
85 |
oops
|
|
86 |
|
|
87 |
lemma "(P&Q-->R) --> (P-->R) | (Q-->R)"
|
|
88 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
89 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
90 |
oops
|
|
91 |
|
|
92 |
|
|
93 |
subsection{*de Bruijn formulae*}
|
|
94 |
|
|
95 |
text{*de Bruijn formula with three predicates*}
|
|
96 |
lemma "((P<->Q) --> P&Q&R) &
|
|
97 |
((Q<->R) --> P&Q&R) &
|
|
98 |
((R<->P) --> P&Q&R) --> P&Q&R"
|
|
99 |
by (tactic{*IntPr.fast_tac 1*})
|
|
100 |
|
|
101 |
|
|
102 |
text{*de Bruijn formula with five predicates*}
|
|
103 |
lemma "((P<->Q) --> P&Q&R&S&T) &
|
|
104 |
((Q<->R) --> P&Q&R&S&T) &
|
|
105 |
((R<->S) --> P&Q&R&S&T) &
|
|
106 |
((S<->T) --> P&Q&R&S&T) &
|
|
107 |
((T<->P) --> P&Q&R&S&T) --> P&Q&R&S&T"
|
|
108 |
by (tactic{*IntPr.fast_tac 1*})
|
|
109 |
|
|
110 |
|
|
111 |
(*** Problems from of Sahlin, Franzen and Haridi,
|
|
112 |
An Intuitionistic Predicate Logic Theorem Prover.
|
|
113 |
J. Logic and Comp. 2 (5), October 1992, 619-656.
|
|
114 |
***)
|
|
115 |
|
|
116 |
text{*Problem 1.1*}
|
|
117 |
lemma "(ALL x. EX y. ALL z. p(x) & q(y) & r(z)) <->
|
|
118 |
(ALL z. EX y. ALL x. p(x) & q(y) & r(z))"
|
|
119 |
by (tactic{*IntPr.best_dup_tac 1*}) --{*SLOW*}
|
|
120 |
|
|
121 |
text{*Problem 3.1*}
|
|
122 |
lemma "~ (EX x. ALL y. mem(y,x) <-> ~ mem(x,x))"
|
|
123 |
by (tactic{*IntPr.fast_tac 1*})
|
|
124 |
|
|
125 |
text{*Problem 4.1: hopeless!*}
|
|
126 |
lemma "(ALL x. p(x) --> p(h(x)) | p(g(x))) & (EX x. p(x)) & (ALL x. ~p(h(x)))
|
|
127 |
--> (EX x. p(g(g(g(g(g(x)))))))"
|
|
128 |
oops
|
|
129 |
|
|
130 |
|
|
131 |
subsection{*Intuitionistic FOL: propositional problems based on Pelletier.*}
|
|
132 |
|
|
133 |
text{*~~1*}
|
|
134 |
lemma "~~((P-->Q) <-> (~Q --> ~P))"
|
|
135 |
by (tactic{*IntPr.fast_tac 1*})
|
|
136 |
|
|
137 |
text{*~~2*}
|
|
138 |
lemma "~~(~~P <-> P)"
|
|
139 |
by (tactic{*IntPr.fast_tac 1*})
|
|
140 |
|
|
141 |
text{*3*}
|
|
142 |
lemma "~(P-->Q) --> (Q-->P)"
|
|
143 |
by (tactic{*IntPr.fast_tac 1*})
|
|
144 |
|
|
145 |
text{*~~4*}
|
|
146 |
lemma "~~((~P-->Q) <-> (~Q --> P))"
|
|
147 |
by (tactic{*IntPr.fast_tac 1*})
|
|
148 |
|
|
149 |
text{*~~5*}
|
|
150 |
lemma "~~((P|Q-->P|R) --> P|(Q-->R))"
|
|
151 |
by (tactic{*IntPr.fast_tac 1*})
|
|
152 |
|
|
153 |
text{*~~6*}
|
|
154 |
lemma "~~(P | ~P)"
|
|
155 |
by (tactic{*IntPr.fast_tac 1*})
|
|
156 |
|
|
157 |
text{*~~7*}
|
|
158 |
lemma "~~(P | ~~~P)"
|
|
159 |
by (tactic{*IntPr.fast_tac 1*})
|
|
160 |
|
|
161 |
text{*~~8. Peirce's law*}
|
|
162 |
lemma "~~(((P-->Q) --> P) --> P)"
|
|
163 |
by (tactic{*IntPr.fast_tac 1*})
|
|
164 |
|
|
165 |
text{*9*}
|
|
166 |
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
|
|
167 |
by (tactic{*IntPr.fast_tac 1*})
|
|
168 |
|
|
169 |
text{*10*}
|
|
170 |
lemma "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
|
|
171 |
by (tactic{*IntPr.fast_tac 1*})
|
|
172 |
|
|
173 |
subsection{*11. Proved in each direction (incorrectly, says Pelletier!!) *}
|
|
174 |
lemma "P<->P"
|
|
175 |
by (tactic{*IntPr.fast_tac 1*})
|
|
176 |
|
|
177 |
text{*~~12. Dijkstra's law *}
|
|
178 |
lemma "~~(((P <-> Q) <-> R) <-> (P <-> (Q <-> R)))"
|
|
179 |
by (tactic{*IntPr.fast_tac 1*})
|
|
180 |
|
|
181 |
lemma "((P <-> Q) <-> R) --> ~~(P <-> (Q <-> R))"
|
|
182 |
by (tactic{*IntPr.fast_tac 1*})
|
|
183 |
|
|
184 |
text{*13. Distributive law*}
|
|
185 |
lemma "P | (Q & R) <-> (P | Q) & (P | R)"
|
|
186 |
by (tactic{*IntPr.fast_tac 1*})
|
|
187 |
|
|
188 |
text{*~~14*}
|
|
189 |
lemma "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
|
|
190 |
by (tactic{*IntPr.fast_tac 1*})
|
|
191 |
|
|
192 |
text{*~~15*}
|
|
193 |
lemma "~~((P --> Q) <-> (~P | Q))"
|
|
194 |
by (tactic{*IntPr.fast_tac 1*})
|
|
195 |
|
|
196 |
text{*~~16*}
|
|
197 |
lemma "~~((P-->Q) | (Q-->P))"
|
|
198 |
by (tactic{*IntPr.fast_tac 1*})
|
|
199 |
|
|
200 |
text{*~~17*}
|
|
201 |
lemma "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
|
|
202 |
by (tactic{*IntPr.fast_tac 1*})
|
|
203 |
|
|
204 |
text{*Dijkstra's "Golden Rule"*}
|
|
205 |
lemma "(P&Q) <-> P <-> Q <-> (P|Q)"
|
|
206 |
by (tactic{*IntPr.fast_tac 1*})
|
|
207 |
|
|
208 |
|
|
209 |
subsection{*****Examples with quantifiers*****}
|
|
210 |
|
|
211 |
|
|
212 |
subsection{*The converse is classical in the following implications...*}
|
|
213 |
|
|
214 |
lemma "(EX x. P(x)-->Q) --> (ALL x. P(x)) --> Q"
|
|
215 |
by (tactic{*IntPr.fast_tac 1*})
|
|
216 |
|
|
217 |
lemma "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
|
|
218 |
by (tactic{*IntPr.fast_tac 1*})
|
|
219 |
|
|
220 |
lemma "((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"
|
|
221 |
by (tactic{*IntPr.fast_tac 1*})
|
|
222 |
|
|
223 |
lemma "(ALL x. P(x)) | Q --> (ALL x. P(x) | Q)"
|
|
224 |
by (tactic{*IntPr.fast_tac 1*})
|
|
225 |
|
|
226 |
lemma "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
|
|
227 |
by (tactic{*IntPr.fast_tac 1*})
|
|
228 |
|
|
229 |
|
|
230 |
|
|
231 |
|
|
232 |
subsection{*The following are not constructively valid!*}
|
|
233 |
text{*The attempt to prove them terminates quickly!*}
|
|
234 |
|
|
235 |
lemma "((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
|
|
236 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
237 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
238 |
oops
|
|
239 |
|
|
240 |
lemma "(P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
|
|
241 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
242 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
243 |
oops
|
|
244 |
|
|
245 |
lemma "(ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
|
|
246 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
247 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
248 |
oops
|
|
249 |
|
|
250 |
lemma "(ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
|
|
251 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
252 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
253 |
oops
|
|
254 |
|
|
255 |
text{*Classically but not intuitionistically valid. Proved by a bug in 1986!*}
|
|
256 |
lemma "EX x. Q(x) --> (ALL x. Q(x))"
|
|
257 |
apply (tactic{*IntPr.fast_tac 1*} | -)
|
|
258 |
apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
|
|
259 |
oops
|
|
260 |
|
|
261 |
|
|
262 |
subsection{*Hard examples with quantifiers*}
|
|
263 |
|
|
264 |
text{*The ones that have not been proved are not known to be valid!
|
|
265 |
Some will require quantifier duplication -- not currently available*}
|
|
266 |
|
|
267 |
text{*~~18*}
|
|
268 |
lemma "~~(EX y. ALL x. P(y)-->P(x))"
|
|
269 |
oops --{*NOT PROVED*}
|
|
270 |
|
|
271 |
text{*~~19*}
|
|
272 |
lemma "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"
|
|
273 |
oops --{*NOT PROVED*}
|
|
274 |
|
|
275 |
text{*20*}
|
|
276 |
lemma "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))
|
|
277 |
--> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
|
|
278 |
by (tactic{*IntPr.fast_tac 1*})
|
|
279 |
|
|
280 |
text{*21*}
|
|
281 |
lemma "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))"
|
|
282 |
oops --{*NOT PROVED; needs quantifier duplication*}
|
|
283 |
|
|
284 |
text{*22*}
|
|
285 |
lemma "(ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"
|
|
286 |
by (tactic{*IntPr.fast_tac 1*})
|
|
287 |
|
|
288 |
text{*~~23*}
|
|
289 |
lemma "~~ ((ALL x. P | Q(x)) <-> (P | (ALL x. Q(x))))"
|
|
290 |
by (tactic{*IntPr.fast_tac 1*})
|
|
291 |
|
|
292 |
text{*24*}
|
|
293 |
lemma "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &
|
|
294 |
(~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))
|
|
295 |
--> ~~(EX x. P(x)&R(x))"
|
|
296 |
txt{*Not clear why @{text fast_tac}, @{text best_tac}, @{text ASTAR} and
|
|
297 |
@{text ITER_DEEPEN} all take forever*}
|
|
298 |
apply (tactic{* IntPr.safe_tac*})
|
|
299 |
apply (erule impE)
|
|
300 |
apply (tactic{*IntPr.fast_tac 1*})
|
|
301 |
by (tactic{*IntPr.fast_tac 1*})
|
|
302 |
|
|
303 |
text{*25*}
|
|
304 |
lemma "(EX x. P(x)) &
|
|
305 |
(ALL x. L(x) --> ~ (M(x) & R(x))) &
|
|
306 |
(ALL x. P(x) --> (M(x) & L(x))) &
|
|
307 |
((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))
|
|
308 |
--> (EX x. Q(x)&P(x))"
|
|
309 |
by (tactic{*IntPr.fast_tac 1*})
|
|
310 |
|
|
311 |
text{*~~26*}
|
|
312 |
lemma "(~~(EX x. p(x)) <-> ~~(EX x. q(x))) &
|
|
313 |
(ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y)))
|
|
314 |
--> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"
|
|
315 |
oops --{*NOT PROVED*}
|
|
316 |
|
|
317 |
text{*27*}
|
|
318 |
lemma "(EX x. P(x) & ~Q(x)) &
|
|
319 |
(ALL x. P(x) --> R(x)) &
|
|
320 |
(ALL x. M(x) & L(x) --> P(x)) &
|
|
321 |
((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))
|
|
322 |
--> (ALL x. M(x) --> ~L(x))"
|
|
323 |
by (tactic{*IntPr.fast_tac 1*})
|
|
324 |
|
|
325 |
text{*~~28. AMENDED*}
|
|
326 |
lemma "(ALL x. P(x) --> (ALL x. Q(x))) &
|
|
327 |
(~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) &
|
|
328 |
(~~(EX x. S(x)) --> (ALL x. L(x) --> M(x)))
|
|
329 |
--> (ALL x. P(x) & L(x) --> M(x))"
|
|
330 |
by (tactic{*IntPr.fast_tac 1*})
|
|
331 |
|
|
332 |
text{*29. Essentially the same as Principia Mathematica *11.71*}
|
|
333 |
lemma "(EX x. P(x)) & (EX y. Q(y))
|
|
334 |
--> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <->
|
|
335 |
(ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
|
|
336 |
by (tactic{*IntPr.fast_tac 1*})
|
|
337 |
|
|
338 |
text{*~~30*}
|
|
339 |
lemma "(ALL x. (P(x) | Q(x)) --> ~ R(x)) &
|
|
340 |
(ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
|
|
341 |
--> (ALL x. ~~S(x))"
|
|
342 |
by (tactic{*IntPr.fast_tac 1*})
|
|
343 |
|
|
344 |
text{*31*}
|
|
345 |
lemma "~(EX x. P(x) & (Q(x) | R(x))) &
|
|
346 |
(EX x. L(x) & P(x)) &
|
|
347 |
(ALL x. ~ R(x) --> M(x))
|
|
348 |
--> (EX x. L(x) & M(x))"
|
|
349 |
by (tactic{*IntPr.fast_tac 1*})
|
|
350 |
|
|
351 |
text{*32*}
|
|
352 |
lemma "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) &
|
|
353 |
(ALL x. S(x) & R(x) --> L(x)) &
|
|
354 |
(ALL x. M(x) --> R(x))
|
|
355 |
--> (ALL x. P(x) & M(x) --> L(x))"
|
|
356 |
by (tactic{*IntPr.fast_tac 1*})
|
|
357 |
|
|
358 |
text{*~~33*}
|
|
359 |
lemma "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c))) <->
|
|
360 |
(ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))"
|
|
361 |
apply (tactic{*IntPr.best_tac 1*})
|
|
362 |
done
|
|
363 |
|
|
364 |
|
|
365 |
text{*36*}
|
|
366 |
lemma "(ALL x. EX y. J(x,y)) &
|
|
367 |
(ALL x. EX y. G(x,y)) &
|
|
368 |
(ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z)))
|
|
369 |
--> (ALL x. EX y. H(x,y))"
|
|
370 |
by (tactic{*IntPr.fast_tac 1*})
|
|
371 |
|
|
372 |
text{*37*}
|
|
373 |
lemma "(ALL z. EX w. ALL x. EX y.
|
|
374 |
~~(P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u. Q(u,w)))) &
|
|
375 |
(ALL x z. ~P(x,z) --> (EX y. Q(y,z))) &
|
|
376 |
(~~(EX x y. Q(x,y)) --> (ALL x. R(x,x)))
|
|
377 |
--> ~~(ALL x. EX y. R(x,y))"
|
|
378 |
oops --{*NOT PROVED*}
|
|
379 |
|
|
380 |
text{*39*}
|
|
381 |
lemma "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
|
|
382 |
by (tactic{*IntPr.fast_tac 1*})
|
|
383 |
|
|
384 |
text{*40. AMENDED*}
|
|
385 |
lemma "(EX y. ALL x. F(x,y) <-> F(x,x)) -->
|
|
386 |
~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
|
|
387 |
by (tactic{*IntPr.fast_tac 1*})
|
|
388 |
|
|
389 |
text{*44*}
|
|
390 |
lemma "(ALL x. f(x) -->
|
|
391 |
(EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) &
|
|
392 |
(EX x. j(x) & (ALL y. g(y) --> h(x,y)))
|
|
393 |
--> (EX x. j(x) & ~f(x))"
|
|
394 |
by (tactic{*IntPr.fast_tac 1*})
|
|
395 |
|
|
396 |
text{*48*}
|
|
397 |
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
|
|
398 |
by (tactic{*IntPr.fast_tac 1*})
|
|
399 |
|
|
400 |
text{*51*}
|
|
401 |
lemma "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) -->
|
|
402 |
(EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
|
|
403 |
by (tactic{*IntPr.fast_tac 1*})
|
|
404 |
|
|
405 |
text{*52*}
|
|
406 |
text{*Almost the same as 51. *}
|
|
407 |
lemma "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) -->
|
|
408 |
(EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)"
|
|
409 |
by (tactic{*IntPr.fast_tac 1*})
|
|
410 |
|
|
411 |
text{*56*}
|
|
412 |
lemma "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
|
|
413 |
by (tactic{*IntPr.fast_tac 1*})
|
|
414 |
|
|
415 |
text{*57*}
|
|
416 |
lemma "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &
|
|
417 |
(ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"
|
|
418 |
by (tactic{*IntPr.fast_tac 1*})
|
|
419 |
|
|
420 |
text{*60*}
|
|
421 |
lemma "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
|
|
422 |
by (tactic{*IntPr.fast_tac 1*})
|
|
423 |
|
|
424 |
end
|
|
425 |
|