author | blanchet |
Fri, 30 Apr 2010 09:36:45 +0200 | |
changeset 36569 | 3a29eb7606c3 |
parent 36566 | f91342f218a9 |
child 36573 | badb32303d1e |
permissions | -rw-r--r-- |
33027 | 1 |
(* Title: HOL/Metis_Examples/set.thy |
23449 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
||
33027 | 4 |
Testing the metis method. |
23449 | 5 |
*) |
6 |
||
33027 | 7 |
theory set |
8 |
imports Main |
|
23449 | 9 |
begin |
10 |
||
36566 | 11 |
sledgehammer_params [isar_proof, debug, overlord] |
12 |
||
23449 | 13 |
lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & |
14 |
(S(x,y) | ~S(y,z) | Q(Z,Z)) & |
|
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
23519
diff
changeset
|
15 |
(Q(X,y) | ~Q(y,Z) | S(X,X))" |
23519 | 16 |
by metis |
23449 | 17 |
|
18 |
lemma "P(n::nat) ==> ~P(0) ==> n ~= 0" |
|
19 |
by metis |
|
20 |
||
36407 | 21 |
sledgehammer_params [shrink_factor = 1] |
23449 | 22 |
|
23 |
(*multiple versions of this example*) |
|
24 |
lemma (*equal_union: *) |
|
36566 | 25 |
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
23449 | 26 |
proof (neg_clausify) |
27 |
fix x |
|
28 |
assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z" |
|
29 |
assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z" |
|
30 |
assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
31 |
assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
32 |
assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
33 |
assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z" |
23449 | 34 |
have 6: "sup Y Z = X \<or> Y \<subseteq> X" |
32685 | 35 |
by (metis 0) |
23449 | 36 |
have 7: "sup Y Z = X \<or> Z \<subseteq> X" |
32685 | 37 |
by (metis 1) |
23449 | 38 |
have 8: "\<And>X3. sup Y Z = X \<or> X \<subseteq> X3 \<or> \<not> Y \<subseteq> X3 \<or> \<not> Z \<subseteq> X3" |
32685 | 39 |
by (metis 5) |
23449 | 40 |
have 9: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 41 |
by (metis 2) |
23449 | 42 |
have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 43 |
by (metis 3) |
23449 | 44 |
have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 45 |
by (metis 4) |
23449 | 46 |
have 12: "Z \<subseteq> X" |
32685 | 47 |
by (metis Un_upper2 7) |
23449 | 48 |
have 13: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z" |
32685 | 49 |
by (metis 8 Un_upper2) |
23449 | 50 |
have 14: "Y \<subseteq> X" |
32685 | 51 |
by (metis Un_upper1 6) |
23449 | 52 |
have 15: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X" |
53 |
by (metis 10 12) |
|
54 |
have 16: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X" |
|
55 |
by (metis 9 12) |
|
56 |
have 17: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X" |
|
57 |
by (metis 11 12) |
|
58 |
have 18: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x" |
|
59 |
by (metis 17 14) |
|
60 |
have 19: "Z \<subseteq> x \<or> sup Y Z \<noteq> X" |
|
61 |
by (metis 15 14) |
|
62 |
have 20: "Y \<subseteq> x \<or> sup Y Z \<noteq> X" |
|
63 |
by (metis 16 14) |
|
64 |
have 21: "sup Y Z = X \<or> X \<subseteq> sup Y Z" |
|
32685 | 65 |
by (metis 13 Un_upper1) |
23449 | 66 |
have 22: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X" |
67 |
by (metis equalityI 21) |
|
68 |
have 23: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X" |
|
32685 | 69 |
by (metis 22 Un_least) |
23449 | 70 |
have 24: "sup Y Z = X \<or> \<not> Y \<subseteq> X" |
71 |
by (metis 23 12) |
|
72 |
have 25: "sup Y Z = X" |
|
73 |
by (metis 24 14) |
|
74 |
have 26: "\<And>X3. X \<subseteq> X3 \<or> \<not> Z \<subseteq> X3 \<or> \<not> Y \<subseteq> X3" |
|
32685 | 75 |
by (metis Un_least 25) |
23449 | 76 |
have 27: "Y \<subseteq> x" |
77 |
by (metis 20 25) |
|
78 |
have 28: "Z \<subseteq> x" |
|
79 |
by (metis 19 25) |
|
80 |
have 29: "\<not> X \<subseteq> x" |
|
81 |
by (metis 18 25) |
|
82 |
have 30: "X \<subseteq> x \<or> \<not> Y \<subseteq> x" |
|
83 |
by (metis 26 28) |
|
84 |
have 31: "X \<subseteq> x" |
|
85 |
by (metis 30 27) |
|
86 |
show "False" |
|
87 |
by (metis 31 29) |
|
88 |
qed |
|
89 |
||
36407 | 90 |
sledgehammer_params [shrink_factor = 2] |
23449 | 91 |
|
92 |
lemma (*equal_union: *) |
|
93 |
"(X = Y \<union> Z) = |
|
94 |
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
|
95 |
proof (neg_clausify) |
|
96 |
fix x |
|
97 |
assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z" |
|
98 |
assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z" |
|
99 |
assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
100 |
assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
101 |
assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
102 |
assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z" |
23449 | 103 |
have 6: "sup Y Z = X \<or> Y \<subseteq> X" |
32685 | 104 |
by (metis 0) |
23449 | 105 |
have 7: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 106 |
by (metis 2) |
23449 | 107 |
have 8: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 108 |
by (metis 4) |
23449 | 109 |
have 9: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z" |
32685 | 110 |
by (metis 5 Un_upper2) |
23449 | 111 |
have 10: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X" |
32685 | 112 |
by (metis 3 Un_upper2) |
23449 | 113 |
have 11: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X" |
32685 | 114 |
by (metis 8 Un_upper2) |
23449 | 115 |
have 12: "Z \<subseteq> x \<or> sup Y Z \<noteq> X" |
32685 | 116 |
by (metis 10 Un_upper1) |
23449 | 117 |
have 13: "sup Y Z = X \<or> X \<subseteq> sup Y Z" |
32685 | 118 |
by (metis 9 Un_upper1) |
23449 | 119 |
have 14: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X" |
32685 | 120 |
by (metis equalityI 13 Un_least) |
23449 | 121 |
have 15: "sup Y Z = X" |
32685 | 122 |
by (metis 14 1 6) |
23449 | 123 |
have 16: "Y \<subseteq> x" |
32685 | 124 |
by (metis 7 Un_upper2 Un_upper1 15) |
23449 | 125 |
have 17: "\<not> X \<subseteq> x" |
32685 | 126 |
by (metis 11 Un_upper1 15) |
23449 | 127 |
have 18: "X \<subseteq> x" |
32685 | 128 |
by (metis Un_least 15 12 15 16) |
23449 | 129 |
show "False" |
130 |
by (metis 18 17) |
|
131 |
qed |
|
132 |
||
36407 | 133 |
sledgehammer_params [shrink_factor = 3] |
23449 | 134 |
|
135 |
lemma (*equal_union: *) |
|
136 |
"(X = Y \<union> Z) = |
|
137 |
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
|
138 |
proof (neg_clausify) |
|
139 |
fix x |
|
140 |
assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z" |
|
141 |
assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z" |
|
142 |
assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
143 |
assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
144 |
assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
145 |
assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z" |
23449 | 146 |
have 6: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 147 |
by (metis 3) |
23449 | 148 |
have 7: "\<And>X3. sup Y Z = X \<or> X \<subseteq> sup X3 Z \<or> \<not> Y \<subseteq> sup X3 Z" |
32685 | 149 |
by (metis 5 Un_upper2) |
23449 | 150 |
have 8: "Y \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X" |
32685 | 151 |
by (metis 2 Un_upper2) |
23449 | 152 |
have 9: "Z \<subseteq> x \<or> sup Y Z \<noteq> X" |
32685 | 153 |
by (metis 6 Un_upper2 Un_upper1) |
23449 | 154 |
have 10: "sup Y Z = X \<or> \<not> sup Y Z \<subseteq> X" |
32685 | 155 |
by (metis equalityI 7 Un_upper1) |
23449 | 156 |
have 11: "sup Y Z = X" |
32685 | 157 |
by (metis 10 Un_least 1 0) |
23449 | 158 |
have 12: "Z \<subseteq> x" |
159 |
by (metis 9 11) |
|
160 |
have 13: "X \<subseteq> x" |
|
32685 | 161 |
by (metis Un_least 11 12 8 Un_upper1 11) |
23449 | 162 |
show "False" |
32685 | 163 |
by (metis 13 4 Un_upper2 Un_upper1 11) |
23449 | 164 |
qed |
165 |
||
166 |
(*Example included in TPHOLs paper*) |
|
167 |
||
36407 | 168 |
sledgehammer_params [shrink_factor = 4] |
23449 | 169 |
|
170 |
lemma (*equal_union: *) |
|
171 |
"(X = Y \<union> Z) = |
|
172 |
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
|
173 |
proof (neg_clausify) |
|
174 |
fix x |
|
175 |
assume 0: "Y \<subseteq> X \<or> X = Y \<union> Z" |
|
176 |
assume 1: "Z \<subseteq> X \<or> X = Y \<union> Z" |
|
177 |
assume 2: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Y \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
178 |
assume 3: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> Z \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
179 |
assume 4: "(\<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X \<or> \<not> X \<subseteq> x) \<or> X \<noteq> Y \<union> Z" |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
180 |
assume 5: "\<And>V. ((\<not> Y \<subseteq> V \<or> \<not> Z \<subseteq> V) \<or> X \<subseteq> V) \<or> X = Y \<union> Z" |
23449 | 181 |
have 6: "sup Y Z \<noteq> X \<or> \<not> X \<subseteq> x \<or> \<not> Y \<subseteq> X \<or> \<not> Z \<subseteq> X" |
32685 | 182 |
by (metis 4) |
23449 | 183 |
have 7: "Z \<subseteq> x \<or> sup Y Z \<noteq> X \<or> \<not> Y \<subseteq> X" |
32685 | 184 |
by (metis 3 Un_upper2) |
23449 | 185 |
have 8: "Z \<subseteq> x \<or> sup Y Z \<noteq> X" |
32685 | 186 |
by (metis 7 Un_upper1) |
23449 | 187 |
have 9: "sup Y Z = X \<or> \<not> Z \<subseteq> X \<or> \<not> Y \<subseteq> X" |
32685 | 188 |
by (metis equalityI 5 Un_upper2 Un_upper1 Un_least) |
23449 | 189 |
have 10: "Y \<subseteq> x" |
32685 | 190 |
by (metis 2 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) |
23449 | 191 |
have 11: "X \<subseteq> x" |
32685 | 192 |
by (metis Un_least 9 Un_upper2 1 Un_upper1 0 8 9 Un_upper2 1 Un_upper1 0 10) |
23449 | 193 |
show "False" |
32685 | 194 |
by (metis 11 6 Un_upper2 1 Un_upper1 0 9 Un_upper2 1 Un_upper1 0) |
23449 | 195 |
qed |
196 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32685
diff
changeset
|
197 |
declare [[ atp_problem_prefix = "set__equal_union" ]] |
23449 | 198 |
lemma (*equal_union: *) |
199 |
"(X = Y \<union> Z) = |
|
200 |
(Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" |
|
201 |
(*One shot proof: hand-reduced. Metis can't do the full proof any more.*) |
|
202 |
by (metis Un_least Un_upper1 Un_upper2 set_eq_subset) |
|
203 |
||
204 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32685
diff
changeset
|
205 |
declare [[ atp_problem_prefix = "set__equal_inter" ]] |
23449 | 206 |
lemma "(X = Y \<inter> Z) = |
207 |
(X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))" |
|
208 |
by (metis Int_greatest Int_lower1 Int_lower2 set_eq_subset) |
|
209 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32685
diff
changeset
|
210 |
declare [[ atp_problem_prefix = "set__fixedpoint" ]] |
23449 | 211 |
lemma fixedpoint: |
212 |
"\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y" |
|
213 |
by metis |
|
214 |
||
26312 | 215 |
lemma (*fixedpoint:*) |
23449 | 216 |
"\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y" |
217 |
proof (neg_clausify) |
|
218 |
fix x xa |
|
219 |
assume 0: "f (g x) = x" |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
220 |
assume 1: "\<And>y. y = x \<or> f (g y) \<noteq> y" |
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
221 |
assume 2: "\<And>x. g (f (xa x)) = xa x \<or> g (f x) \<noteq> x" |
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
222 |
assume 3: "\<And>x. g (f x) \<noteq> x \<or> xa x \<noteq> x" |
23449 | 223 |
have 4: "\<And>X1. g (f X1) \<noteq> X1 \<or> g x \<noteq> X1" |
23519 | 224 |
by (metis 3 1 2) |
23449 | 225 |
show "False" |
226 |
by (metis 4 0) |
|
227 |
qed |
|
228 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32685
diff
changeset
|
229 |
declare [[ atp_problem_prefix = "set__singleton_example" ]] |
23449 | 230 |
lemma (*singleton_example_2:*) |
231 |
"\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}" |
|
232 |
by (metis Set.subsetI Union_upper insertCI set_eq_subset) |
|
233 |
--{*found by SPASS*} |
|
234 |
||
235 |
lemma (*singleton_example_2:*) |
|
236 |
"\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}" |
|
32685 | 237 |
by (metis Set.subsetI Union_upper insert_iff set_eq_subset) |
23449 | 238 |
|
239 |
lemma singleton_example_2: |
|
240 |
"\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}" |
|
241 |
proof (neg_clausify) |
|
24937
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
242 |
assume 0: "\<And>x. \<not> S \<subseteq> {x}" |
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents:
24855
diff
changeset
|
243 |
assume 1: "\<And>x. x \<notin> S \<or> \<Union>S \<subseteq> x" |
23449 | 244 |
have 2: "\<And>X3. X3 = \<Union>S \<or> \<not> X3 \<subseteq> \<Union>S \<or> X3 \<notin> S" |
24855 | 245 |
by (metis set_eq_subset 1) |
23449 | 246 |
have 3: "\<And>X3. S \<subseteq> insert (\<Union>S) X3" |
24855 | 247 |
by (metis insert_iff Set.subsetI Union_upper 2 Set.subsetI) |
23449 | 248 |
show "False" |
24855 | 249 |
by (metis 3 0) |
23449 | 250 |
qed |
251 |
||
252 |
||
253 |
||
254 |
text {* |
|
255 |
From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages |
|
256 |
293-314. |
|
257 |
*} |
|
258 |
||
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32685
diff
changeset
|
259 |
declare [[ atp_problem_prefix = "set__Bledsoe_Fung" ]] |
23449 | 260 |
(*Notes: 1, the numbering doesn't completely agree with the paper. |
261 |
2, we must rename set variables to avoid type clashes.*) |
|
262 |
lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))" |
|
263 |
"D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B" |
|
264 |
"P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)" |
|
265 |
"a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B" |
|
266 |
"P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" |
|
267 |
"P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" |
|
268 |
"\<exists>A. a \<notin> A" |
|
36566 | 269 |
"(\<forall>C. (0, 0) \<in> C \<and> (\<forall>x y. (x, y) \<in> C \<longrightarrow> (Suc x, Suc y) \<in> C) \<longrightarrow> (n, m) \<in> C) \<and> Q n \<longrightarrow> Q m" |
270 |
apply (metis all_not_in_conv)+ |
|
271 |
apply (metis mem_def) |
|
23449 | 272 |
apply (metis insertCI singletonE zless_le) |
32519 | 273 |
apply (metis Collect_def Collect_mem_eq) |
274 |
apply (metis Collect_def Collect_mem_eq) |
|
23449 | 275 |
apply (metis DiffE) |
36566 | 276 |
apply (metis pair_in_Id_conv) |
23449 | 277 |
done |
278 |
||
279 |
end |