31974
|
1 |
(* Title: FOL/ex/Classical.thy
|
14236
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
3 |
Copyright 1994 University of Cambridge
|
|
4 |
*)
|
|
5 |
|
61489
|
6 |
section \<open>Classical Predicate Calculus Problems\<close>
|
14236
|
7 |
|
61489
|
8 |
theory Classical
|
|
9 |
imports FOL
|
|
10 |
begin
|
14236
|
11 |
|
61489
|
12 |
lemma "(P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longrightarrow> Q) \<or> (P \<longrightarrow> R)"
|
|
13 |
by blast
|
14236
|
14 |
|
|
15 |
|
61489
|
16 |
subsubsection \<open>If and only if\<close>
|
|
17 |
|
|
18 |
lemma "(P \<longleftrightarrow> Q) \<longleftrightarrow> (Q \<longleftrightarrow> P)"
|
|
19 |
by blast
|
|
20 |
|
|
21 |
lemma "\<not> (P \<longleftrightarrow> \<not> P)"
|
|
22 |
by blast
|
|
23 |
|
|
24 |
|
|
25 |
subsection \<open>Pelletier's examples\<close>
|
14236
|
26 |
|
61489
|
27 |
text \<open>
|
|
28 |
Sample problems from
|
|
29 |
|
|
30 |
\<^item> F. J. Pelletier,
|
|
31 |
Seventy-Five Problems for Testing Automatic Theorem Provers,
|
|
32 |
J. Automated Reasoning 2 (1986), 191-216.
|
|
33 |
Errata, JAR 4 (1988), 236-236.
|
|
34 |
|
|
35 |
The hardest problems -- judging by experience with several theorem
|
|
36 |
provers, including matrix ones -- are 34 and 43.
|
60770
|
37 |
\<close>
|
14236
|
38 |
|
60770
|
39 |
text\<open>1\<close>
|
61489
|
40 |
lemma "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> \<not> P)"
|
|
41 |
by blast
|
14236
|
42 |
|
60770
|
43 |
text\<open>2\<close>
|
61489
|
44 |
lemma "\<not> \<not> P \<longleftrightarrow> P"
|
|
45 |
by blast
|
14236
|
46 |
|
60770
|
47 |
text\<open>3\<close>
|
61489
|
48 |
lemma "\<not> (P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P)"
|
|
49 |
by blast
|
14236
|
50 |
|
60770
|
51 |
text\<open>4\<close>
|
61489
|
52 |
lemma "(\<not> P \<longrightarrow> Q) \<longleftrightarrow> (\<not> Q \<longrightarrow> P)"
|
|
53 |
by blast
|
14236
|
54 |
|
60770
|
55 |
text\<open>5\<close>
|
61489
|
56 |
lemma "((P \<or> Q) \<longrightarrow> (P \<or> R)) \<longrightarrow> (P \<or> (Q \<longrightarrow> R))"
|
|
57 |
by blast
|
14236
|
58 |
|
60770
|
59 |
text\<open>6\<close>
|
61489
|
60 |
lemma "P \<or> \<not> P"
|
|
61 |
by blast
|
14236
|
62 |
|
60770
|
63 |
text\<open>7\<close>
|
61489
|
64 |
lemma "P \<or> \<not> \<not> \<not> P"
|
|
65 |
by blast
|
14236
|
66 |
|
61489
|
67 |
text\<open>8. Peirce's law\<close>
|
|
68 |
lemma "((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P"
|
|
69 |
by blast
|
14236
|
70 |
|
60770
|
71 |
text\<open>9\<close>
|
61489
|
72 |
lemma "((P \<or> Q) \<and> (\<not> P \<or> Q) \<and> (P \<or> \<not> Q)) \<longrightarrow> \<not> (\<not> P \<or> \<not> Q)"
|
|
73 |
by blast
|
14236
|
74 |
|
60770
|
75 |
text\<open>10\<close>
|
61489
|
76 |
lemma "(Q \<longrightarrow> R) \<and> (R \<longrightarrow> P \<and> Q) \<and> (P \<longrightarrow> Q \<or> R) \<longrightarrow> (P \<longleftrightarrow> Q)"
|
|
77 |
by blast
|
14236
|
78 |
|
61489
|
79 |
text\<open>11. Proved in each direction (incorrectly, says Pelletier!!)\<close>
|
|
80 |
lemma "P \<longleftrightarrow> P"
|
|
81 |
by blast
|
14236
|
82 |
|
61489
|
83 |
text\<open>12. "Dijkstra's law"\<close>
|
|
84 |
lemma "((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))"
|
|
85 |
by blast
|
14236
|
86 |
|
61489
|
87 |
text\<open>13. Distributive law\<close>
|
|
88 |
lemma "P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)"
|
|
89 |
by blast
|
14236
|
90 |
|
60770
|
91 |
text\<open>14\<close>
|
61489
|
92 |
lemma "(P \<longleftrightarrow> Q) \<longleftrightarrow> ((Q \<or> \<not> P) \<and> (\<not> Q \<or> P))"
|
|
93 |
by blast
|
14236
|
94 |
|
60770
|
95 |
text\<open>15\<close>
|
61489
|
96 |
lemma "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"
|
|
97 |
by blast
|
14236
|
98 |
|
60770
|
99 |
text\<open>16\<close>
|
61489
|
100 |
lemma "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"
|
|
101 |
by blast
|
14236
|
102 |
|
60770
|
103 |
text\<open>17\<close>
|
61489
|
104 |
lemma "((P \<and> (Q \<longrightarrow> R)) \<longrightarrow> S) \<longleftrightarrow> ((\<not> P \<or> Q \<or> S) \<and> (\<not> P \<or> \<not> R \<or> S))"
|
|
105 |
by blast
|
|
106 |
|
14236
|
107 |
|
61489
|
108 |
subsection \<open>Classical Logic: examples with quantifiers\<close>
|
14236
|
109 |
|
61489
|
110 |
lemma "(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))"
|
|
111 |
by blast
|
14236
|
112 |
|
61489
|
113 |
lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))"
|
|
114 |
by blast
|
14236
|
115 |
|
61489
|
116 |
lemma "(\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q"
|
|
117 |
by blast
|
14236
|
118 |
|
61489
|
119 |
lemma "(\<forall>x. P(x)) \<or> Q \<longleftrightarrow> (\<forall>x. P(x) \<or> Q)"
|
|
120 |
by blast
|
14236
|
121 |
|
60770
|
122 |
text\<open>Discussed in Avron, Gentzen-Type Systems, Resolution and Tableaux,
|
|
123 |
JAR 10 (265-281), 1993. Proof is trivial!\<close>
|
61489
|
124 |
lemma "\<not> ((\<exists>x. \<not> P(x)) \<and> ((\<exists>x. P(x)) \<or> (\<exists>x. P(x) \<and> Q(x))) \<and> \<not> (\<exists>x. P(x)))"
|
|
125 |
by blast
|
14236
|
126 |
|
|
127 |
|
61489
|
128 |
subsection \<open>Problems requiring quantifier duplication\<close>
|
|
129 |
|
|
130 |
text\<open>Theorem B of Peter Andrews, Theorem Proving via General Matings,
|
60770
|
131 |
JACM 28 (1981).\<close>
|
61489
|
132 |
lemma "(\<exists>x. \<forall>y. P(x) \<longleftrightarrow> P(y)) \<longrightarrow> ((\<exists>x. P(x)) \<longleftrightarrow> (\<forall>y. P(y)))"
|
|
133 |
by blast
|
14236
|
134 |
|
60770
|
135 |
text\<open>Needs multiple instantiation of ALL.\<close>
|
61489
|
136 |
lemma "(\<forall>x. P(x) \<longrightarrow> P(f(x))) \<and> P(d) \<longrightarrow> P(f(f(f(d))))"
|
|
137 |
by blast
|
14236
|
138 |
|
60770
|
139 |
text\<open>Needs double instantiation of the quantifier\<close>
|
61489
|
140 |
lemma "\<exists>x. P(x) \<longrightarrow> P(a) \<and> P(b)"
|
|
141 |
by blast
|
14236
|
142 |
|
61489
|
143 |
lemma "\<exists>z. P(z) \<longrightarrow> (\<forall>x. P(x))"
|
|
144 |
by blast
|
14236
|
145 |
|
61489
|
146 |
lemma "\<exists>x. (\<exists>y. P(y)) \<longrightarrow> P(x)"
|
|
147 |
by blast
|
14236
|
148 |
|
61489
|
149 |
text\<open>V. Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23. NOT PROVED.\<close>
|
|
150 |
lemma
|
|
151 |
"\<exists>x x'. \<forall>y. \<exists>z z'.
|
|
152 |
(\<not> P(y,y) \<or> P(x,x) \<or> \<not> S(z,x)) \<and>
|
|
153 |
(S(x,y) \<or> \<not> S(y,z) \<or> Q(z',z')) \<and>
|
|
154 |
(Q(x',y) \<or> \<not> Q(y,z') \<or> S(x',x'))"
|
|
155 |
oops
|
14236
|
156 |
|
|
157 |
|
61489
|
158 |
subsection \<open>Hard examples with quantifiers\<close>
|
14236
|
159 |
|
60770
|
160 |
text\<open>18\<close>
|
61489
|
161 |
lemma "\<exists>y. \<forall>x. P(y) \<longrightarrow> P(x)"
|
|
162 |
by blast
|
14236
|
163 |
|
60770
|
164 |
text\<open>19\<close>
|
61489
|
165 |
lemma "\<exists>x. \<forall>y z. (P(y) \<longrightarrow> Q(z)) \<longrightarrow> (P(x) \<longrightarrow> Q(x))"
|
|
166 |
by blast
|
14236
|
167 |
|
60770
|
168 |
text\<open>20\<close>
|
61489
|
169 |
lemma "(\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y) \<longrightarrow> R(z) \<and> S(w)))
|
|
170 |
\<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))"
|
|
171 |
by blast
|
14236
|
172 |
|
60770
|
173 |
text\<open>21\<close>
|
61489
|
174 |
lemma "(\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> (\<exists>x. P \<longleftrightarrow> Q(x))"
|
|
175 |
by blast
|
14236
|
176 |
|
60770
|
177 |
text\<open>22\<close>
|
61489
|
178 |
lemma "(\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))"
|
|
179 |
by blast
|
14236
|
180 |
|
60770
|
181 |
text\<open>23\<close>
|
61489
|
182 |
lemma "(\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x)))"
|
|
183 |
by blast
|
14236
|
184 |
|
60770
|
185 |
text\<open>24\<close>
|
61489
|
186 |
lemma
|
|
187 |
"\<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and>
|
|
188 |
(\<not> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x))) \<and> (\<forall>x. Q(x) \<or> R(x) \<longrightarrow> S(x))
|
|
189 |
\<longrightarrow> (\<exists>x. P(x) \<and> R(x))"
|
|
190 |
by blast
|
14236
|
191 |
|
60770
|
192 |
text\<open>25\<close>
|
61489
|
193 |
lemma
|
|
194 |
"(\<exists>x. P(x)) \<and>
|
|
195 |
(\<forall>x. L(x) \<longrightarrow> \<not> (M(x) \<and> R(x))) \<and>
|
|
196 |
(\<forall>x. P(x) \<longrightarrow> (M(x) \<and> L(x))) \<and>
|
|
197 |
((\<forall>x. P(x) \<longrightarrow> Q(x)) \<or> (\<exists>x. P(x) \<and> R(x)))
|
|
198 |
\<longrightarrow> (\<exists>x. Q(x) \<and> P(x))"
|
|
199 |
by blast
|
14236
|
200 |
|
60770
|
201 |
text\<open>26\<close>
|
61489
|
202 |
lemma
|
|
203 |
"((\<exists>x. p(x)) \<longleftrightarrow> (\<exists>x. q(x))) \<and>
|
|
204 |
(\<forall>x. \<forall>y. p(x) \<and> q(y) \<longrightarrow> (r(x) \<longleftrightarrow> s(y)))
|
|
205 |
\<longrightarrow> ((\<forall>x. p(x) \<longrightarrow> r(x)) \<longleftrightarrow> (\<forall>x. q(x) \<longrightarrow> s(x)))"
|
|
206 |
by blast
|
14236
|
207 |
|
60770
|
208 |
text\<open>27\<close>
|
61489
|
209 |
lemma
|
|
210 |
"(\<exists>x. P(x) \<and> \<not> Q(x)) \<and>
|
|
211 |
(\<forall>x. P(x) \<longrightarrow> R(x)) \<and>
|
|
212 |
(\<forall>x. M(x) \<and> L(x) \<longrightarrow> P(x)) \<and>
|
|
213 |
((\<exists>x. R(x) \<and> \<not> Q(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> \<not> R(x)))
|
|
214 |
\<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))"
|
|
215 |
by blast
|
14236
|
216 |
|
61489
|
217 |
text\<open>28. AMENDED\<close>
|
|
218 |
lemma
|
|
219 |
"(\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and>
|
|
220 |
((\<forall>x. Q(x) \<or> R(x)) \<longrightarrow> (\<exists>x. Q(x) \<and> S(x))) \<and>
|
|
221 |
((\<exists>x. S(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> M(x)))
|
|
222 |
\<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))"
|
|
223 |
by blast
|
14236
|
224 |
|
61489
|
225 |
text\<open>29. Essentially the same as Principia Mathematica *11.71\<close>
|
|
226 |
lemma
|
|
227 |
"(\<exists>x. P(x)) \<and> (\<exists>y. Q(y))
|
|
228 |
\<longrightarrow> ((\<forall>x. P(x) \<longrightarrow> R(x)) \<and> (\<forall>y. Q(y) \<longrightarrow> S(y)) \<longleftrightarrow>
|
|
229 |
(\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))"
|
|
230 |
by blast
|
14236
|
231 |
|
60770
|
232 |
text\<open>30\<close>
|
61489
|
233 |
lemma
|
|
234 |
"(\<forall>x. P(x) \<or> Q(x) \<longrightarrow> \<not> R(x)) \<and>
|
|
235 |
(\<forall>x. (Q(x) \<longrightarrow> \<not> S(x)) \<longrightarrow> P(x) \<and> R(x))
|
|
236 |
\<longrightarrow> (\<forall>x. S(x))"
|
|
237 |
by blast
|
14236
|
238 |
|
60770
|
239 |
text\<open>31\<close>
|
61489
|
240 |
lemma
|
|
241 |
"\<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and>
|
|
242 |
(\<exists>x. L(x) \<and> P(x)) \<and>
|
|
243 |
(\<forall>x. \<not> R(x) \<longrightarrow> M(x))
|
|
244 |
\<longrightarrow> (\<exists>x. L(x) \<and> M(x))"
|
|
245 |
by blast
|
14236
|
246 |
|
60770
|
247 |
text\<open>32\<close>
|
61489
|
248 |
lemma
|
|
249 |
"(\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and>
|
|
250 |
(\<forall>x. S(x) \<and> R(x) \<longrightarrow> L(x)) \<and>
|
|
251 |
(\<forall>x. M(x) \<longrightarrow> R(x))
|
|
252 |
\<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))"
|
|
253 |
by blast
|
14236
|
254 |
|
60770
|
255 |
text\<open>33\<close>
|
61489
|
256 |
lemma
|
|
257 |
"(\<forall>x. P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c)) \<longleftrightarrow>
|
|
258 |
(\<forall>x. (\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c)))"
|
|
259 |
by blast
|
14236
|
260 |
|
61489
|
261 |
text\<open>34. AMENDED (TWICE!!). Andrews's challenge.\<close>
|
|
262 |
lemma
|
|
263 |
"((\<exists>x. \<forall>y. p(x) \<longleftrightarrow> p(y)) \<longleftrightarrow> ((\<exists>x. q(x)) \<longleftrightarrow> (\<forall>y. p(y)))) \<longleftrightarrow>
|
|
264 |
((\<exists>x. \<forall>y. q(x) \<longleftrightarrow> q(y)) \<longleftrightarrow> ((\<exists>x. p(x)) \<longleftrightarrow> (\<forall>y. q(y))))"
|
|
265 |
by blast
|
14236
|
266 |
|
60770
|
267 |
text\<open>35\<close>
|
61489
|
268 |
lemma "\<exists>x y. P(x,y) \<longrightarrow> (\<forall>u v. P(u,v))"
|
|
269 |
by blast
|
14236
|
270 |
|
60770
|
271 |
text\<open>36\<close>
|
61489
|
272 |
lemma
|
|
273 |
"(\<forall>x. \<exists>y. J(x,y)) \<and>
|
|
274 |
(\<forall>x. \<exists>y. G(x,y)) \<and>
|
|
275 |
(\<forall>x y. J(x,y) \<or> G(x,y) \<longrightarrow> (\<forall>z. J(y,z) \<or> G(y,z) \<longrightarrow> H(x,z)))
|
|
276 |
\<longrightarrow> (\<forall>x. \<exists>y. H(x,y))"
|
|
277 |
by blast
|
14236
|
278 |
|
60770
|
279 |
text\<open>37\<close>
|
61489
|
280 |
lemma
|
|
281 |
"(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
|
|
282 |
(P(x,z) \<longrightarrow> P(y,w)) \<and> P(y,z) \<and> (P(y,w) \<longrightarrow> (\<exists>u. Q(u,w)))) \<and>
|
|
283 |
(\<forall>x z. \<not> P(x,z) \<longrightarrow> (\<exists>y. Q(y,z))) \<and>
|
|
284 |
((\<exists>x y. Q(x,y)) \<longrightarrow> (\<forall>x. R(x,x)))
|
|
285 |
\<longrightarrow> (\<forall>x. \<exists>y. R(x,y))"
|
|
286 |
by blast
|
14236
|
287 |
|
60770
|
288 |
text\<open>38\<close>
|
61489
|
289 |
lemma
|
|
290 |
"(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> (\<exists>y. p(y) \<and> r(x,y))) \<longrightarrow>
|
|
291 |
(\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<longleftrightarrow>
|
|
292 |
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> (\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<and>
|
|
293 |
(\<not> p(a) \<or> \<not> (\<exists>y. p(y) \<and> r(x,y)) \<or>
|
|
294 |
(\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))))"
|
|
295 |
by blast
|
14236
|
296 |
|
60770
|
297 |
text\<open>39\<close>
|
61489
|
298 |
lemma "\<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))"
|
|
299 |
by blast
|
14236
|
300 |
|
61489
|
301 |
text\<open>40. AMENDED\<close>
|
|
302 |
lemma
|
|
303 |
"(\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow>
|
|
304 |
\<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))"
|
|
305 |
by blast
|
14236
|
306 |
|
60770
|
307 |
text\<open>41\<close>
|
61489
|
308 |
lemma
|
|
309 |
"(\<forall>z. \<exists>y. \<forall>x. f(x,y) \<longleftrightarrow> f(x,z) \<and> \<not> f(x,x))
|
|
310 |
\<longrightarrow> \<not> (\<exists>z. \<forall>x. f(x,z))"
|
|
311 |
by blast
|
14236
|
312 |
|
60770
|
313 |
text\<open>42\<close>
|
61489
|
314 |
lemma "\<not> (\<exists>y. \<forall>x. p(x,y) \<longleftrightarrow> \<not> (\<exists>z. p(x,z) \<and> p(z,x)))"
|
|
315 |
by blast
|
14236
|
316 |
|
60770
|
317 |
text\<open>43\<close>
|
61489
|
318 |
lemma
|
|
319 |
"(\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> (\<forall>z. p(z,x) \<longleftrightarrow> p(z,y)))
|
|
320 |
\<longrightarrow> (\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> q(y,x))"
|
|
321 |
by blast
|
14236
|
322 |
|
61489
|
323 |
text \<open>
|
62020
|
324 |
Other proofs: Can use \<open>auto\<close>, which cheats by using rewriting!
|
|
325 |
\<open>Deepen_tac\<close> alone requires 253 secs. Or
|
|
326 |
\<open>by (mini_tac 1 THEN Deepen_tac 5 1)\<close>.
|
61489
|
327 |
\<close>
|
14236
|
328 |
|
60770
|
329 |
text\<open>44\<close>
|
61489
|
330 |
lemma
|
|
331 |
"(\<forall>x. f(x) \<longrightarrow> (\<exists>y. g(y) \<and> h(x,y) \<and> (\<exists>y. g(y) \<and> \<not> h(x,y)))) \<and>
|
|
332 |
(\<exists>x. j(x) \<and> (\<forall>y. g(y) \<longrightarrow> h(x,y)))
|
|
333 |
\<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))"
|
|
334 |
by blast
|
14236
|
335 |
|
60770
|
336 |
text\<open>45\<close>
|
61489
|
337 |
lemma
|
|
338 |
"(\<forall>x. f(x) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y))
|
|
339 |
\<longrightarrow> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> k(y))) \<and>
|
|
340 |
\<not> (\<exists>y. l(y) \<and> k(y)) \<and>
|
|
341 |
(\<exists>x. f(x) \<and> (\<forall>y. h(x,y) \<longrightarrow> l(y)) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y)))
|
|
342 |
\<longrightarrow> (\<exists>x. f(x) \<and> \<not> (\<exists>y. g(y) \<and> h(x,y)))"
|
|
343 |
by blast
|
14236
|
344 |
|
|
345 |
|
60770
|
346 |
text\<open>46\<close>
|
61489
|
347 |
lemma
|
|
348 |
"(\<forall>x. f(x) \<and> (\<forall>y. f(y) \<and> h(y,x) \<longrightarrow> g(y)) \<longrightarrow> g(x)) \<and>
|
|
349 |
((\<exists>x. f(x) \<and> \<not> g(x)) \<longrightarrow>
|
|
350 |
(\<exists>x. f(x) \<and> \<not> g(x) \<and> (\<forall>y. f(y) \<and> \<not> g(y) \<longrightarrow> j(x,y)))) \<and>
|
|
351 |
(\<forall>x y. f(x) \<and> f(y) \<and> h(x,y) \<longrightarrow> \<not> j(y,x))
|
|
352 |
\<longrightarrow> (\<forall>x. f(x) \<longrightarrow> g(x))"
|
|
353 |
by blast
|
14236
|
354 |
|
|
355 |
|
61489
|
356 |
subsection \<open>Problems (mainly) involving equality or functions\<close>
|
14236
|
357 |
|
60770
|
358 |
text\<open>48\<close>
|
61489
|
359 |
lemma "(a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c"
|
|
360 |
by blast
|
14236
|
361 |
|
61489
|
362 |
text\<open>49. NOT PROVED AUTOMATICALLY. Hard because it involves substitution for
|
|
363 |
Vars; the type constraint ensures that x,y,z have the same type as a,b,u.\<close>
|
|
364 |
lemma
|
|
365 |
"(\<exists>x y::'a. \<forall>z. z = x \<or> z = y) \<and> P(a) \<and> P(b) \<and> a \<noteq> b \<longrightarrow> (\<forall>u::'a. P(u))"
|
|
366 |
apply safe
|
|
367 |
apply (rule_tac x = a in allE, assumption)
|
|
368 |
apply (rule_tac x = b in allE, assumption)
|
62020
|
369 |
apply fast \<comment> \<open>blast's treatment of equality can't do it\<close>
|
61489
|
370 |
done
|
14236
|
371 |
|
61489
|
372 |
text\<open>50. (What has this to do with equality?)\<close>
|
|
373 |
lemma "(\<forall>x. P(a,x) \<or> (\<forall>y. P(x,y))) \<longrightarrow> (\<exists>x. \<forall>y. P(x,y))"
|
|
374 |
by blast
|
14236
|
375 |
|
60770
|
376 |
text\<open>51\<close>
|
61489
|
377 |
lemma
|
|
378 |
"(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
|
|
379 |
(\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y=w) \<longleftrightarrow> x = z)"
|
|
380 |
by blast
|
14236
|
381 |
|
60770
|
382 |
text\<open>52\<close>
|
|
383 |
text\<open>Almost the same as 51.\<close>
|
61489
|
384 |
lemma
|
|
385 |
"(\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow>
|
|
386 |
(\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)"
|
|
387 |
by blast
|
14236
|
388 |
|
60770
|
389 |
text\<open>55\<close>
|
|
390 |
text\<open>Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
|
|
391 |
fast DISCOVERS who killed Agatha.\<close>
|
61489
|
392 |
schematic_goal
|
|
393 |
"lives(agatha) \<and> lives(butler) \<and> lives(charles) \<and>
|
|
394 |
(killed(agatha,agatha) \<or> killed(butler,agatha) \<or> killed(charles,agatha)) \<and>
|
|
395 |
(\<forall>x y. killed(x,y) \<longrightarrow> hates(x,y) \<and> \<not> richer(x,y)) \<and>
|
|
396 |
(\<forall>x. hates(agatha,x) \<longrightarrow> \<not> hates(charles,x)) \<and>
|
|
397 |
(hates(agatha,agatha) \<and> hates(agatha,charles)) \<and>
|
|
398 |
(\<forall>x. lives(x) \<and> \<not> richer(x,agatha) \<longrightarrow> hates(butler,x)) \<and>
|
|
399 |
(\<forall>x. hates(agatha,x) \<longrightarrow> hates(butler,x)) \<and>
|
|
400 |
(\<forall>x. \<not> hates(x,agatha) \<or> \<not> hates(x,butler) \<or> \<not> hates(x,charles)) \<longrightarrow>
|
14236
|
401 |
killed(?who,agatha)"
|
62020
|
402 |
by fast \<comment> \<open>MUCH faster than blast\<close>
|
14236
|
403 |
|
|
404 |
|
60770
|
405 |
text\<open>56\<close>
|
61489
|
406 |
lemma "(\<forall>x. (\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))"
|
|
407 |
by blast
|
14236
|
408 |
|
60770
|
409 |
text\<open>57\<close>
|
61489
|
410 |
lemma
|
|
411 |
"P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and>
|
|
412 |
(\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))"
|
|
413 |
by blast
|
14236
|
414 |
|
60770
|
415 |
text\<open>58 NOT PROVED AUTOMATICALLY\<close>
|
61489
|
416 |
lemma "(\<forall>x y. f(x) = g(y)) \<longrightarrow> (\<forall>x y. f(f(x)) = f(g(y)))"
|
|
417 |
by (slow elim: subst_context)
|
14236
|
418 |
|
|
419 |
|
60770
|
420 |
text\<open>59\<close>
|
61489
|
421 |
lemma "(\<forall>x. P(x) \<longleftrightarrow> \<not> P(f(x))) \<longrightarrow> (\<exists>x. P(x) \<and> \<not> P(f(x)))"
|
|
422 |
by blast
|
14236
|
423 |
|
60770
|
424 |
text\<open>60\<close>
|
61489
|
425 |
lemma "\<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))"
|
|
426 |
by blast
|
14236
|
427 |
|
60770
|
428 |
text\<open>62 as corrected in JAR 18 (1997), page 135\<close>
|
61489
|
429 |
lemma
|
|
430 |
"(\<forall>x. p(a) \<and> (p(x) \<longrightarrow> p(f(x))) \<longrightarrow> p(f(f(x)))) \<longleftrightarrow>
|
|
431 |
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> p(f(f(x)))) \<and>
|
|
432 |
(\<not> p(a) \<or> \<not> p(f(x)) \<or> p(f(f(x)))))"
|
|
433 |
by blast
|
14236
|
434 |
|
61489
|
435 |
text \<open>From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
|
60770
|
436 |
fast indeed copes!\<close>
|
61489
|
437 |
lemma
|
|
438 |
"(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
|
|
439 |
(\<exists>x. K(x) \<and> F(x) \<and> (\<forall>y. H(x,y) \<longrightarrow> K(y))) \<and>
|
|
440 |
(\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<and> J(x))"
|
|
441 |
by fast
|
14236
|
442 |
|
61489
|
443 |
text \<open>From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
|
60770
|
444 |
It does seem obvious!\<close>
|
61489
|
445 |
lemma
|
|
446 |
"(\<forall>x. F(x) \<and> \<not> G(x) \<longrightarrow> (\<exists>y. H(x,y) \<and> J(y))) \<and>
|
|
447 |
(\<exists>x. K(x) \<and> F(x) \<and> (\<forall>y. H(x,y) \<longrightarrow> K(y))) \<and>
|
|
448 |
(\<forall>x. K(x) \<longrightarrow> \<not> G(x)) \<longrightarrow> (\<exists>x. K(x) \<longrightarrow> \<not> G(x))"
|
|
449 |
by fast
|
14236
|
450 |
|
61489
|
451 |
text \<open>Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
|
|
452 |
author U. Egly.\<close>
|
|
453 |
lemma
|
|
454 |
"((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
|
|
455 |
(\<exists>w. C(w) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(w,y,z)))))
|
|
456 |
\<and>
|
|
457 |
(\<forall>w. C(w) \<and> (\<forall>u. C(u) \<longrightarrow> (\<forall>v. D(w,u,v))) \<longrightarrow>
|
|
458 |
(\<forall>y z.
|
|
459 |
(C(y) \<and> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,g)) \<and>
|
|
460 |
(C(y) \<and> \<not> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,b))))
|
|
461 |
\<and>
|
|
462 |
(\<forall>w. C(w) \<and>
|
|
463 |
(\<forall>y z.
|
|
464 |
(C(y) \<and> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,g)) \<and>
|
|
465 |
(C(y) \<and> \<not> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,b))) \<longrightarrow>
|
|
466 |
(\<exists>v. C(v) \<and>
|
|
467 |
(\<forall>y. ((C(y) \<and> Q(w,y,y)) \<and> OO(w,g) \<longrightarrow> \<not> P(v,y)) \<and>
|
|
468 |
((C(y) \<and> Q(w,y,y)) \<and> OO(w,b) \<longrightarrow> P(v,y) \<and> OO(v,b)))))
|
|
469 |
\<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))"
|
|
470 |
by (blast 12)
|
62020
|
471 |
\<comment> \<open>Needed because the search for depths below 12 is very slow.\<close>
|
14236
|
472 |
|
|
473 |
|
61489
|
474 |
text \<open>
|
|
475 |
Halting problem II: credited to M. Bruschi by Li Dafa in JAR 18(1),
|
|
476 |
p. 105.
|
|
477 |
\<close>
|
|
478 |
lemma
|
|
479 |
"((\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z)))) \<longrightarrow>
|
|
480 |
(\<exists>w. C(w) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(w,y,z)))))
|
|
481 |
\<and>
|
|
482 |
(\<forall>w. C(w) \<and> (\<forall>u. C(u) \<longrightarrow> (\<forall>v. D(w,u,v))) \<longrightarrow>
|
|
483 |
(\<forall>y z.
|
|
484 |
(C(y) \<and> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,g)) \<and>
|
|
485 |
(C(y) \<and> \<not> P(y,z) \<longrightarrow> Q(w,y,z) \<and> OO(w,b))))
|
|
486 |
\<and>
|
|
487 |
((\<exists>w. C(w) \<and> (\<forall>y. (C(y) \<and> P(y,y) \<longrightarrow> Q(w,y,y) \<and> OO(w,g)) \<and>
|
|
488 |
(C(y) \<and> \<not> P(y,y) \<longrightarrow> Q(w,y,y) \<and> OO(w,b))))
|
|
489 |
\<longrightarrow>
|
|
490 |
(\<exists>v. C(v) \<and> (\<forall>y. (C(y) \<and> P(y,y) \<longrightarrow> P(v,y) \<and> OO(v,g)) \<and>
|
|
491 |
(C(y) \<and> \<not> P(y,y) \<longrightarrow> P(v,y) \<and> OO(v,b)))))
|
|
492 |
\<longrightarrow>
|
|
493 |
((\<exists>v. C(v) \<and> (\<forall>y. (C(y) \<and> P(y,y) \<longrightarrow> P(v,y) \<and> OO(v,g)) \<and>
|
|
494 |
(C(y) \<and> \<not> P(y,y) \<longrightarrow> P(v,y) \<and> OO(v,b))))
|
|
495 |
\<longrightarrow>
|
|
496 |
(\<exists>u. C(u) \<and> (\<forall>y. (C(y) \<and> P(y,y) \<longrightarrow> \<not> P(u,y)) \<and>
|
|
497 |
(C(y) \<and> \<not> P(y,y) \<longrightarrow> P(u,y) \<and> OO(u,b)))))
|
|
498 |
\<longrightarrow> \<not> (\<exists>x. A(x) \<and> (\<forall>y. C(y) \<longrightarrow> (\<forall>z. D(x,y,z))))"
|
|
499 |
by blast
|
14236
|
500 |
|
61489
|
501 |
text \<open>Challenge found on info-hol.\<close>
|
|
502 |
lemma "\<forall>x. \<exists>v w. \<forall>y z. P(x) \<and> Q(y) \<longrightarrow> (P(v) \<or> R(w)) \<and> (R(z) \<longrightarrow> Q(v))"
|
|
503 |
by blast
|
14236
|
504 |
|
61489
|
505 |
text \<open>
|
|
506 |
Attributed to Lewis Carroll by S. G. Pulman. The first or last assumption
|
|
507 |
can be deleted.\<close>
|
|
508 |
lemma
|
|
509 |
"(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and>
|
|
510 |
\<not> (\<exists>x. grocer(x) \<and> healthy(x)) \<and>
|
|
511 |
(\<forall>x. industrious(x) \<and> grocer(x) \<longrightarrow> honest(x)) \<and>
|
|
512 |
(\<forall>x. cyclist(x) \<longrightarrow> industrious(x)) \<and>
|
|
513 |
(\<forall>x. \<not> healthy(x) \<and> cyclist(x) \<longrightarrow> \<not> honest(x))
|
|
514 |
\<longrightarrow> (\<forall>x. grocer(x) \<longrightarrow> \<not> cyclist(x))"
|
|
515 |
by blast
|
14236
|
516 |
|
|
517 |
|
|
518 |
(*Runtimes for old versions of this file:
|
61489
|
519 |
Thu Jul 23 1992: loaded in 467s using iffE [on SPARC2]
|
|
520 |
Mon Nov 14 1994: loaded in 144s [on SPARC10, with deepen_tac]
|
|
521 |
Wed Nov 16 1994: loaded in 138s [after addition of norm_term_skip]
|
|
522 |
Mon Nov 21 1994: loaded in 131s [DEPTH_FIRST suppressing repetitions]
|
14236
|
523 |
|
|
524 |
Further runtimes on a Sun-4
|
61489
|
525 |
Tue Mar 4 1997: loaded in 93s (version 94-7)
|
14236
|
526 |
Tue Mar 4 1997: loaded in 89s
|
|
527 |
Thu Apr 3 1997: loaded in 44s--using mostly Blast_tac
|
|
528 |
Thu Apr 3 1997: loaded in 96s--addition of two Halting Probs
|
|
529 |
Thu Apr 3 1997: loaded in 98s--using lim-1 for all haz rules
|
|
530 |
Tue Dec 2 1997: loaded in 107s--added 46; new equalSubst
|
|
531 |
Fri Dec 12 1997: loaded in 91s--faster proof reconstruction
|
|
532 |
Thu Dec 18 1997: loaded in 94s--two new "obvious theorems" (??)
|
|
533 |
*)
|
|
534 |
|
|
535 |
end
|