author | wenzelm |
Tue, 03 Dec 2019 10:50:28 +0100 | |
changeset 71218 | 73b313432d8a |
parent 69104 | f33352dbbf12 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: Sequents/LK/Hard_Quantifiers.thy |
21426 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1992 University of Cambridge |
|
4 |
||
5 |
Hard examples with quantifiers. Can be read to test the LK system. |
|
6 |
From F. J. Pelletier, |
|
7 |
Seventy-Five Problems for Testing Automatic Theorem Provers, |
|
8 |
J. Automated Reasoning 2 (1986), 191-216. |
|
9 |
Errata, JAR 4 (1988), 236-236. |
|
10 |
||
11 |
Uses pc_tac rather than fast_tac when the former is significantly faster. |
|
12 |
*) |
|
13 |
||
14 |
theory Hard_Quantifiers |
|
55229 | 15 |
imports "../LK" |
21426 | 16 |
begin |
17 |
||
61386 | 18 |
lemma "\<turnstile> (\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> (\<forall>x. P(x)) \<and> (\<forall>x. Q(x))" |
21426 | 19 |
by fast |
20 |
||
61386 | 21 |
lemma "\<turnstile> (\<exists>x. P \<longrightarrow> Q(x)) \<longleftrightarrow> (P \<longrightarrow> (\<exists>x. Q(x)))" |
21426 | 22 |
by fast |
23 |
||
61386 | 24 |
lemma "\<turnstile> (\<exists>x. P(x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x)) \<longrightarrow> Q" |
21426 | 25 |
by fast |
26 |
||
61386 | 27 |
lemma "\<turnstile> (\<forall>x. P(x)) \<or> Q \<longleftrightarrow> (\<forall>x. P(x) \<or> Q)" |
21426 | 28 |
by fast |
29 |
||
30 |
||
31 |
text "Problems requiring quantifier duplication" |
|
32 |
||
61385 | 33 |
(*Not provable by fast: needs multiple instantiation of \<forall>*) |
61386 | 34 |
lemma "\<turnstile> (\<forall>x. P(x) \<longrightarrow> P(f(x))) \<and> P(d) \<longrightarrow> P(f(f(f(d))))" |
21426 | 35 |
by best_dup |
36 |
||
37 |
(*Needs double instantiation of the quantifier*) |
|
61386 | 38 |
lemma "\<turnstile> \<exists>x. P(x) \<longrightarrow> P(a) \<and> P(b)" |
21426 | 39 |
by fast_dup |
40 |
||
61386 | 41 |
lemma "\<turnstile> \<exists>z. P(z) \<longrightarrow> (\<forall>x. P(x))" |
21426 | 42 |
by best_dup |
43 |
||
44 |
||
45 |
text "Hard examples with quantifiers" |
|
46 |
||
47 |
text "Problem 18" |
|
61386 | 48 |
lemma "\<turnstile> \<exists>y. \<forall>x. P(y)\<longrightarrow>P(x)" |
21426 | 49 |
by best_dup |
50 |
||
51 |
text "Problem 19" |
|
61386 | 52 |
lemma "\<turnstile> \<exists>x. \<forall>y z. (P(y)\<longrightarrow>Q(z)) \<longrightarrow> (P(x)\<longrightarrow>Q(x))" |
21426 | 53 |
by best_dup |
54 |
||
55 |
text "Problem 20" |
|
61386 | 56 |
lemma "\<turnstile> (\<forall>x y. \<exists>z. \<forall>w. (P(x) \<and> Q(y)\<longrightarrow>R(z) \<and> S(w))) |
61385 | 57 |
\<longrightarrow> (\<exists>x y. P(x) \<and> Q(y)) \<longrightarrow> (\<exists>z. R(z))" |
21426 | 58 |
by fast |
59 |
||
60 |
text "Problem 21" |
|
61386 | 61 |
lemma "\<turnstile> (\<exists>x. P \<longrightarrow> Q(x)) \<and> (\<exists>x. Q(x) \<longrightarrow> P) \<longrightarrow> (\<exists>x. P \<longleftrightarrow> Q(x))" |
21426 | 62 |
by best_dup |
63 |
||
64 |
text "Problem 22" |
|
61386 | 65 |
lemma "\<turnstile> (\<forall>x. P \<longleftrightarrow> Q(x)) \<longrightarrow> (P \<longleftrightarrow> (\<forall>x. Q(x)))" |
21426 | 66 |
by fast |
67 |
||
68 |
text "Problem 23" |
|
61386 | 69 |
lemma "\<turnstile> (\<forall>x. P \<or> Q(x)) \<longleftrightarrow> (P \<or> (\<forall>x. Q(x)))" |
21426 | 70 |
by best |
71 |
||
72 |
text "Problem 24" |
|
61386 | 73 |
lemma "\<turnstile> \<not> (\<exists>x. S(x) \<and> Q(x)) \<and> (\<forall>x. P(x) \<longrightarrow> Q(x) \<or> R(x)) \<and> |
61385 | 74 |
\<not> (\<exists>x. P(x)) \<longrightarrow> (\<exists>x. Q(x)) \<and> (\<forall>x. Q(x) \<or> R(x) \<longrightarrow> S(x)) |
75 |
\<longrightarrow> (\<exists>x. P(x) \<and> R(x))" |
|
55228 | 76 |
by pc |
21426 | 77 |
|
78 |
text "Problem 25" |
|
61386 | 79 |
lemma "\<turnstile> (\<exists>x. P(x)) \<and> |
61385 | 80 |
(\<forall>x. L(x) \<longrightarrow> \<not> (M(x) \<and> R(x))) \<and> |
81 |
(\<forall>x. P(x) \<longrightarrow> (M(x) \<and> L(x))) \<and> |
|
82 |
((\<forall>x. P(x)\<longrightarrow>Q(x)) \<or> (\<exists>x. P(x) \<and> R(x))) |
|
83 |
\<longrightarrow> (\<exists>x. Q(x) \<and> P(x))" |
|
21426 | 84 |
by best |
85 |
||
86 |
text "Problem 26" |
|
61386 | 87 |
lemma "\<turnstile> ((\<exists>x. p(x)) \<longleftrightarrow> (\<exists>x. q(x))) \<and> |
61385 | 88 |
(\<forall>x. \<forall>y. p(x) \<and> q(y) \<longrightarrow> (r(x) \<longleftrightarrow> s(y))) |
89 |
\<longrightarrow> ((\<forall>x. p(x)\<longrightarrow>r(x)) \<longleftrightarrow> (\<forall>x. q(x)\<longrightarrow>s(x)))" |
|
55228 | 90 |
by pc |
21426 | 91 |
|
92 |
text "Problem 27" |
|
61386 | 93 |
lemma "\<turnstile> (\<exists>x. P(x) \<and> \<not> Q(x)) \<and> |
61385 | 94 |
(\<forall>x. P(x) \<longrightarrow> R(x)) \<and> |
95 |
(\<forall>x. M(x) \<and> L(x) \<longrightarrow> P(x)) \<and> |
|
96 |
((\<exists>x. R(x) \<and> \<not> Q(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> \<not> R(x))) |
|
97 |
\<longrightarrow> (\<forall>x. M(x) \<longrightarrow> \<not> L(x))" |
|
55228 | 98 |
by pc |
21426 | 99 |
|
100 |
text "Problem 28. AMENDED" |
|
61386 | 101 |
lemma "\<turnstile> (\<forall>x. P(x) \<longrightarrow> (\<forall>x. Q(x))) \<and> |
61385 | 102 |
((\<forall>x. Q(x) \<or> R(x)) \<longrightarrow> (\<exists>x. Q(x) \<and> S(x))) \<and> |
103 |
((\<exists>x. S(x)) \<longrightarrow> (\<forall>x. L(x) \<longrightarrow> M(x))) |
|
104 |
\<longrightarrow> (\<forall>x. P(x) \<and> L(x) \<longrightarrow> M(x))" |
|
55228 | 105 |
by pc |
21426 | 106 |
|
107 |
text "Problem 29. Essentially the same as Principia Mathematica *11.71" |
|
61386 | 108 |
lemma "\<turnstile> (\<exists>x. P(x)) \<and> (\<exists>y. Q(y)) |
61385 | 109 |
\<longrightarrow> ((\<forall>x. P(x) \<longrightarrow> R(x)) \<and> (\<forall>y. Q(y) \<longrightarrow> S(y)) \<longleftrightarrow> |
110 |
(\<forall>x y. P(x) \<and> Q(y) \<longrightarrow> R(x) \<and> S(y)))" |
|
55228 | 111 |
by pc |
21426 | 112 |
|
113 |
text "Problem 30" |
|
61386 | 114 |
lemma "\<turnstile> (\<forall>x. P(x) \<or> Q(x) \<longrightarrow> \<not> R(x)) \<and> |
61385 | 115 |
(\<forall>x. (Q(x) \<longrightarrow> \<not> S(x)) \<longrightarrow> P(x) \<and> R(x)) |
116 |
\<longrightarrow> (\<forall>x. S(x))" |
|
21426 | 117 |
by fast |
118 |
||
119 |
text "Problem 31" |
|
61386 | 120 |
lemma "\<turnstile> \<not> (\<exists>x. P(x) \<and> (Q(x) \<or> R(x))) \<and> |
61385 | 121 |
(\<exists>x. L(x) \<and> P(x)) \<and> |
122 |
(\<forall>x. \<not> R(x) \<longrightarrow> M(x)) |
|
123 |
\<longrightarrow> (\<exists>x. L(x) \<and> M(x))" |
|
21426 | 124 |
by fast |
125 |
||
126 |
text "Problem 32" |
|
61386 | 127 |
lemma "\<turnstile> (\<forall>x. P(x) \<and> (Q(x) \<or> R(x)) \<longrightarrow> S(x)) \<and> |
61385 | 128 |
(\<forall>x. S(x) \<and> R(x) \<longrightarrow> L(x)) \<and> |
129 |
(\<forall>x. M(x) \<longrightarrow> R(x)) |
|
130 |
\<longrightarrow> (\<forall>x. P(x) \<and> M(x) \<longrightarrow> L(x))" |
|
21426 | 131 |
by best |
132 |
||
133 |
text "Problem 33" |
|
61386 | 134 |
lemma "\<turnstile> (\<forall>x. P(a) \<and> (P(x) \<longrightarrow> P(b)) \<longrightarrow> P(c)) \<longleftrightarrow> |
61385 | 135 |
(\<forall>x. (\<not> P(a) \<or> P(x) \<or> P(c)) \<and> (\<not> P(a) \<or> \<not> P(b) \<or> P(c)))" |
21426 | 136 |
by fast |
137 |
||
138 |
text "Problem 34 AMENDED (TWICE!!)" |
|
139 |
(*Andrews's challenge*) |
|
61386 | 140 |
lemma "\<turnstile> ((\<exists>x. \<forall>y. p(x) \<longleftrightarrow> p(y)) \<longleftrightarrow> |
61385 | 141 |
((\<exists>x. q(x)) \<longleftrightarrow> (\<forall>y. p(y)))) \<longleftrightarrow> |
142 |
((\<exists>x. \<forall>y. q(x) \<longleftrightarrow> q(y)) \<longleftrightarrow> |
|
143 |
((\<exists>x. p(x)) \<longleftrightarrow> (\<forall>y. q(y))))" |
|
21426 | 144 |
by best_dup |
145 |
||
146 |
text "Problem 35" |
|
61386 | 147 |
lemma "\<turnstile> \<exists>x y. P(x,y) \<longrightarrow> (\<forall>u v. P(u,v))" |
21426 | 148 |
by best_dup |
149 |
||
150 |
text "Problem 36" |
|
61386 | 151 |
lemma "\<turnstile> (\<forall>x. \<exists>y. J(x,y)) \<and> |
61385 | 152 |
(\<forall>x. \<exists>y. G(x,y)) \<and> |
153 |
(\<forall>x y. J(x,y) \<or> G(x,y) \<longrightarrow> |
|
154 |
(\<forall>z. J(y,z) \<or> G(y,z) \<longrightarrow> H(x,z))) |
|
155 |
\<longrightarrow> (\<forall>x. \<exists>y. H(x,y))" |
|
21426 | 156 |
by fast |
157 |
||
158 |
text "Problem 37" |
|
61386 | 159 |
lemma "\<turnstile> (\<forall>z. \<exists>w. \<forall>x. \<exists>y. |
61385 | 160 |
(P(x,z)\<longrightarrow>P(y,w)) \<and> P(y,z) \<and> (P(y,w) \<longrightarrow> (\<exists>u. Q(u,w)))) \<and> |
161 |
(\<forall>x z. \<not> P(x,z) \<longrightarrow> (\<exists>y. Q(y,z))) \<and> |
|
162 |
((\<exists>x y. Q(x,y)) \<longrightarrow> (\<forall>x. R(x,x))) |
|
163 |
\<longrightarrow> (\<forall>x. \<exists>y. R(x,y))" |
|
55228 | 164 |
by pc |
21426 | 165 |
|
166 |
text "Problem 38" |
|
61386 | 167 |
lemma "\<turnstile> (\<forall>x. p(a) \<and> (p(x) \<longrightarrow> (\<exists>y. p(y) \<and> r(x,y))) \<longrightarrow> |
61385 | 168 |
(\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<longleftrightarrow> |
169 |
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> (\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))) \<and> |
|
170 |
(\<not> p(a) \<or> \<not> (\<exists>y. p(y) \<and> r(x,y)) \<or> |
|
171 |
(\<exists>z. \<exists>w. p(z) \<and> r(x,w) \<and> r(w,z))))" |
|
55228 | 172 |
by pc |
21426 | 173 |
|
174 |
text "Problem 39" |
|
61386 | 175 |
lemma "\<turnstile> \<not> (\<exists>x. \<forall>y. F(y,x) \<longleftrightarrow> \<not> F(y,y))" |
21426 | 176 |
by fast |
177 |
||
178 |
text "Problem 40. AMENDED" |
|
61386 | 179 |
lemma "\<turnstile> (\<exists>y. \<forall>x. F(x,y) \<longleftrightarrow> F(x,x)) \<longrightarrow> |
61385 | 180 |
\<not> (\<forall>x. \<exists>y. \<forall>z. F(z,y) \<longleftrightarrow> \<not> F(z,x))" |
21426 | 181 |
by fast |
182 |
||
183 |
text "Problem 41" |
|
61386 | 184 |
lemma "\<turnstile> (\<forall>z. \<exists>y. \<forall>x. f(x,y) \<longleftrightarrow> f(x,z) \<and> \<not> f(x,x)) |
61385 | 185 |
\<longrightarrow> \<not> (\<exists>z. \<forall>x. f(x,z))" |
21426 | 186 |
by fast |
187 |
||
188 |
text "Problem 42" |
|
61386 | 189 |
lemma "\<turnstile> \<not> (\<exists>y. \<forall>x. p(x,y) \<longleftrightarrow> \<not> (\<exists>z. p(x,z) \<and> p(z,x)))" |
21426 | 190 |
oops |
191 |
||
192 |
text "Problem 43" |
|
61386 | 193 |
lemma "\<turnstile> (\<forall>x. \<forall>y. q(x,y) \<longleftrightarrow> (\<forall>z. p(z,x) \<longleftrightarrow> p(z,y))) |
61385 | 194 |
\<longrightarrow> (\<forall>x. (\<forall>y. q(x,y) \<longleftrightarrow> q(y,x)))" |
21426 | 195 |
oops |
196 |
||
197 |
text "Problem 44" |
|
61386 | 198 |
lemma "\<turnstile> (\<forall>x. f(x) \<longrightarrow> |
61385 | 199 |
(\<exists>y. g(y) \<and> h(x,y) \<and> (\<exists>y. g(y) \<and> \<not> h(x,y)))) \<and> |
200 |
(\<exists>x. j(x) \<and> (\<forall>y. g(y) \<longrightarrow> h(x,y))) |
|
201 |
\<longrightarrow> (\<exists>x. j(x) \<and> \<not> f(x))" |
|
21426 | 202 |
by fast |
203 |
||
204 |
text "Problem 45" |
|
61386 | 205 |
lemma "\<turnstile> (\<forall>x. f(x) \<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y)) |
61385 | 206 |
\<longrightarrow> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> k(y))) \<and> |
207 |
\<not> (\<exists>y. l(y) \<and> k(y)) \<and> |
|
208 |
(\<exists>x. f(x) \<and> (\<forall>y. h(x,y) \<longrightarrow> l(y)) |
|
209 |
\<and> (\<forall>y. g(y) \<and> h(x,y) \<longrightarrow> j(x,y))) |
|
210 |
\<longrightarrow> (\<exists>x. f(x) \<and> \<not> (\<exists>y. g(y) \<and> h(x,y)))" |
|
21426 | 211 |
by best |
212 |
||
213 |
||
214 |
text "Problems (mainly) involving equality or functions" |
|
215 |
||
216 |
text "Problem 48" |
|
61386 | 217 |
lemma "\<turnstile> (a = b \<or> c = d) \<and> (a = c \<or> b = d) \<longrightarrow> a = d \<or> b = c" |
55228 | 218 |
by (fast add!: subst) |
21426 | 219 |
|
220 |
text "Problem 50" |
|
61386 | 221 |
lemma "\<turnstile> (\<forall>x. P(a,x) \<or> (\<forall>y. P(x,y))) \<longrightarrow> (\<exists>x. \<forall>y. P(x,y))" |
21426 | 222 |
by best_dup |
223 |
||
224 |
text "Problem 51" |
|
61386 | 225 |
lemma "\<turnstile> (\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow> |
61385 | 226 |
(\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) \<longleftrightarrow> y = w) \<longleftrightarrow> x = z)" |
55228 | 227 |
by (fast add!: subst) |
21426 | 228 |
|
229 |
text "Problem 52" (*Almost the same as 51. *) |
|
61386 | 230 |
lemma "\<turnstile> (\<exists>z w. \<forall>x y. P(x,y) \<longleftrightarrow> (x = z \<and> y = w)) \<longrightarrow> |
61385 | 231 |
(\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) \<longleftrightarrow> x = z) \<longleftrightarrow> y = w)" |
55228 | 232 |
by (fast add!: subst) |
21426 | 233 |
|
234 |
text "Problem 56" |
|
61386 | 235 |
lemma "\<turnstile> (\<forall>x.(\<exists>y. P(y) \<and> x = f(y)) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> P(f(x)))" |
55228 | 236 |
by (best add: symL subst) |
21426 | 237 |
(*requires tricker to orient the equality properly*) |
238 |
||
239 |
text "Problem 57" |
|
61386 | 240 |
lemma "\<turnstile> P(f(a,b), f(b,c)) \<and> P(f(b,c), f(a,c)) \<and> |
61385 | 241 |
(\<forall>x y z. P(x,y) \<and> P(y,z) \<longrightarrow> P(x,z)) \<longrightarrow> P(f(a,b), f(a,c))" |
21426 | 242 |
by fast |
243 |
||
244 |
text "Problem 58!" |
|
61386 | 245 |
lemma "\<turnstile> (\<forall>x y. f(x) = g(y)) \<longrightarrow> (\<forall>x y. f(f(x)) = f(g(y)))" |
55228 | 246 |
by (fast add!: subst) |
21426 | 247 |
|
248 |
text "Problem 59" |
|
249 |
(*Unification works poorly here -- the abstraction %sobj prevents efficient |
|
250 |
operation of the occurs check*) |
|
61386 | 251 |
lemma "\<turnstile> (\<forall>x. P(x) \<longleftrightarrow> \<not> P(f(x))) \<longrightarrow> (\<exists>x. P(x) \<and> \<not> P(f(x)))" |
69104
f33352dbbf12
reduce tracing messages to make it work in PIDE session;
wenzelm
parents:
61386
diff
changeset
|
252 |
using [[unify_trace_bound = 50]] |
21426 | 253 |
by best_dup |
254 |
||
255 |
text "Problem 60" |
|
61386 | 256 |
lemma "\<turnstile> \<forall>x. P(x,f(x)) \<longleftrightarrow> (\<exists>y. (\<forall>z. P(z,y) \<longrightarrow> P(z,f(x))) \<and> P(x,y))" |
21426 | 257 |
by fast |
258 |
||
259 |
text "Problem 62 as corrected in JAR 18 (1997), page 135" |
|
61386 | 260 |
lemma "\<turnstile> (\<forall>x. p(a) \<and> (p(x) \<longrightarrow> p(f(x))) \<longrightarrow> p(f(f(x)))) \<longleftrightarrow> |
61385 | 261 |
(\<forall>x. (\<not> p(a) \<or> p(x) \<or> p(f(f(x)))) \<and> |
262 |
(\<not> p(a) \<or> \<not> p(f(x)) \<or> p(f(f(x)))))" |
|
21426 | 263 |
by fast |
264 |
||
265 |
(*18 June 92: loaded in 372 secs*) |
|
266 |
(*19 June 92: loaded in 166 secs except #34, using repeat_goal_tac*) |
|
267 |
(*29 June 92: loaded in 370 secs*) |
|
268 |
(*18 September 2005: loaded in 1.809 secs*) |
|
269 |
||
270 |
end |