| author | paulson |
| Thu, 29 Apr 1999 10:51:58 +0200 | |
| changeset 6536 | 281d44905cab |
| parent 6153 | bff90585cce5 |
| child 7499 | 23e090051cb8 |
| permissions | -rw-r--r-- |
| 1461 | 1 |
(* Title: ZF/AC/WO6_WO1.ML |
| 992 | 2 |
ID: $Id$ |
| 1461 | 3 |
Author: Krzysztof Grabczewski |
| 992 | 4 |
|
| 1041 | 5 |
The proof of "WO6 ==> WO1". Simplified by L C Paulson. |
| 992 | 6 |
|
7 |
From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin, |
|
8 |
pages 2-5 |
|
9 |
*) |
|
10 |
||
| 1041 | 11 |
goal OrderType.thy |
12 |
"!!i j k. [| k < i++j; Ord(i); Ord(j) |] ==> \ |
|
13 |
\ k < i | (~ k<i & k = i ++ (k--i) & (k--i)<j)"; |
|
14 |
by (res_inst_tac [("i","k"),("j","i")] Ord_linear2 1);
|
|
15 |
by (dtac odiff_lt_mono2 4 THEN assume_tac 4); |
|
16 |
by (asm_full_simp_tac |
|
| 4091 | 17 |
(simpset() addsimps [oadd_odiff_inverse, odiff_oadd_inverse]) 4); |
18 |
by (safe_tac (claset() addSEs [lt_Ord])); |
|
| 3731 | 19 |
qed "lt_oadd_odiff_disj"; |
| 1041 | 20 |
|
21 |
(*The corresponding elimination rule*) |
|
| 4152 | 22 |
val lt_oadd_odiff_cases = rule_by_tactic Safe_tac |
| 1041 | 23 |
(lt_oadd_odiff_disj RS disjE); |
24 |
||
| 992 | 25 |
(* ********************************************************************** *) |
| 1461 | 26 |
(* The most complicated part of the proof - lemma ii - p. 2-4 *) |
| 992 | 27 |
(* ********************************************************************** *) |
28 |
||
29 |
(* ********************************************************************** *) |
|
| 1461 | 30 |
(* some properties of relation uu(beta, gamma, delta) - p. 2 *) |
| 992 | 31 |
(* ********************************************************************** *) |
32 |
||
| 5068 | 33 |
Goalw [uu_def] "domain(uu(f,b,g,d)) <= f`b"; |
| 3731 | 34 |
by (Blast_tac 1); |
35 |
qed "domain_uu_subset"; |
|
| 992 | 36 |
|
| 5137 | 37 |
Goal "ALL b<a. f`b lepoll m ==> \ |
| 1041 | 38 |
\ ALL b<a. ALL g<a. ALL d<a. domain(uu(f,b,g,d)) lepoll m"; |
| 4091 | 39 |
by (fast_tac (claset() addSEs |
| 1461 | 40 |
[domain_uu_subset RS subset_imp_lepoll RS lepoll_trans]) 1); |
| 3731 | 41 |
qed "quant_domain_uu_lepoll_m"; |
| 992 | 42 |
|
| 5068 | 43 |
Goalw [uu_def] "uu(f,b,g,d) <= f`b * f`g"; |
| 3731 | 44 |
by (Blast_tac 1); |
45 |
qed "uu_subset1"; |
|
| 992 | 46 |
|
| 5068 | 47 |
Goalw [uu_def] "uu(f,b,g,d) <= f`d"; |
| 3731 | 48 |
by (Blast_tac 1); |
49 |
qed "uu_subset2"; |
|
| 992 | 50 |
|
| 5137 | 51 |
Goal "[| ALL b<a. f`b lepoll m; d<a |] ==> uu(f,b,g,d) lepoll m"; |
| 4091 | 52 |
by (fast_tac (claset() |
| 1461 | 53 |
addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1); |
| 3731 | 54 |
qed "uu_lepoll_m"; |
| 992 | 55 |
|
56 |
(* ********************************************************************** *) |
|
| 1461 | 57 |
(* Two cases for lemma ii *) |
| 992 | 58 |
(* ********************************************************************** *) |
| 5068 | 59 |
Goalw [lesspoll_def] |
| 992 | 60 |
"!! a f u. ALL b<a. ALL g<a. ALL d<a. u(f,b,g,d) lepoll m ==> \ |
| 1041 | 61 |
\ (ALL b<a. f`b ~= 0 --> (EX g<a. EX d<a. u(f,b,g,d) ~= 0 & \ |
| 1461 | 62 |
\ u(f,b,g,d) lesspoll m)) | \ |
| 1041 | 63 |
\ (EX b<a. f`b ~= 0 & (ALL g<a. ALL d<a. u(f,b,g,d) ~= 0 --> \ |
| 1461 | 64 |
\ u(f,b,g,d) eqpoll m))"; |
| 2469 | 65 |
by (Asm_simp_tac 1); |
| 4091 | 66 |
by (blast_tac (claset() delrules [equalityI]) 1); |
| 3731 | 67 |
qed "cases"; |
| 992 | 68 |
|
69 |
(* ********************************************************************** *) |
|
| 1461 | 70 |
(* Lemmas used in both cases *) |
| 992 | 71 |
(* ********************************************************************** *) |
| 5137 | 72 |
Goal "Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))"; |
| 4091 | 73 |
by (fast_tac (claset() addSIs [equalityI] addIs [ltI] |
| 1041 | 74 |
addSDs [lt_oadd_disj] |
75 |
addSEs [lt_oadd1, oadd_lt_mono2]) 1); |
|
| 3731 | 76 |
qed "UN_oadd"; |
| 992 | 77 |
|
78 |
||
79 |
(* ********************************************************************** *) |
|
| 1461 | 80 |
(* Case 1 : lemmas *) |
| 992 | 81 |
(* ********************************************************************** *) |
82 |
||
| 5068 | 83 |
Goalw [vv1_def] "vv1(f,m,b) <= f`b"; |
| 1450 | 84 |
by (rtac (LetI RS LetI) 1); |
| 4091 | 85 |
by (simp_tac (simpset() addsimps [domain_uu_subset]) 1); |
| 3731 | 86 |
qed "vv1_subset"; |
| 992 | 87 |
|
88 |
(* ********************************************************************** *) |
|
| 1461 | 89 |
(* Case 1 : Union of images is the whole "y" *) |
| 992 | 90 |
(* ********************************************************************** *) |
| 5068 | 91 |
Goalw [gg1_def] |
| 1461 | 92 |
"!! a f y. [| Ord(a); m:nat |] ==> \ |
93 |
\ (UN b<a++a. gg1(f,a,m)`b) = (UN b<a. f`b)"; |
|
| 1041 | 94 |
by (asm_simp_tac |
| 4091 | 95 |
(simpset() addsimps [UN_oadd, lt_oadd1, |
| 1461 | 96 |
oadd_le_self RS le_imp_not_lt, lt_Ord, |
97 |
odiff_oadd_inverse, ltD, |
|
98 |
vv1_subset RS Diff_partition, ww1_def]) 1); |
|
| 3731 | 99 |
qed "UN_gg1_eq"; |
| 1041 | 100 |
|
| 5068 | 101 |
Goal "domain(gg1(f,a,m)) = a++a"; |
| 4091 | 102 |
by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg1_def]) 1); |
| 3731 | 103 |
qed "domain_gg1"; |
| 992 | 104 |
|
105 |
(* ********************************************************************** *) |
|
| 1461 | 106 |
(* every value of defined function is less than or equipollent to m *) |
| 992 | 107 |
(* ********************************************************************** *) |
| 5137 | 108 |
Goal "[| P(a, b); Ord(a); Ord(b); \ |
| 1461 | 109 |
\ Least_a = (LEAST a. EX x. Ord(x) & P(a, x)) |] \ |
110 |
\ ==> P(Least_a, LEAST b. P(Least_a, b))"; |
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
111 |
by (etac ssubst 1); |
| 992 | 112 |
by (res_inst_tac [("Q","%z. P(z, LEAST b. P(z, b))")] LeastI2 1);
|
| 4091 | 113 |
by (REPEAT (fast_tac (claset() addSEs [LeastI]) 1)); |
| 3731 | 114 |
qed "nested_LeastI"; |
| 992 | 115 |
|
| 1041 | 116 |
val nested_Least_instance = |
117 |
standard |
|
118 |
(read_instantiate |
|
| 1461 | 119 |
[("P","%g d. domain(uu(f,b,g,d)) ~= 0 & \
|
120 |
\ domain(uu(f,b,g,d)) lepoll m")] nested_LeastI); |
|
| 992 | 121 |
|
| 5068 | 122 |
Goalw [gg1_def] |
| 1041 | 123 |
"!!a. [| Ord(a); m:nat; \ |
| 1461 | 124 |
\ ALL b<a. f`b ~=0 --> \ |
125 |
\ (EX g<a. EX d<a. domain(uu(f,b,g,d)) ~= 0 & \ |
|
126 |
\ domain(uu(f,b,g,d)) lepoll m); \ |
|
127 |
\ ALL b<a. f`b lepoll succ(m); b<a++a \ |
|
128 |
\ |] ==> gg1(f,a,m)`b lepoll m"; |
|
| 2469 | 129 |
by (Asm_simp_tac 1); |
| 4091 | 130 |
by (safe_tac (claset() addSEs [lt_oadd_odiff_cases])); |
| 1041 | 131 |
(*Case b<a : show vv1(f,m,b) lepoll m *) |
| 5137 | 132 |
by (asm_simp_tac (simpset() addsimps [vv1_def, Let_def]) 1); |
| 4091 | 133 |
by (fast_tac (claset() addIs [nested_Least_instance RS conjunct2] |
| 1461 | 134 |
addSEs [lt_Ord] |
135 |
addSIs [empty_lepollI]) 1); |
|
| 1041 | 136 |
(*Case a le b: show ww1(f,m,b--a) lepoll m *) |
| 4091 | 137 |
by (asm_simp_tac (simpset() addsimps [ww1_def]) 1); |
| 1041 | 138 |
by (excluded_middle_tac "f`(b--a) = 0" 1); |
| 4091 | 139 |
by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
140 |
by (rtac Diff_lepoll 1); |
| 3731 | 141 |
by (Blast_tac 1); |
| 1041 | 142 |
by (rtac vv1_subset 1); |
143 |
by (dtac (ospec RS mp) 1); |
|
| 6153 | 144 |
by (REPEAT (eresolve_tac [asm_rl, oexE, conjE] 1)); |
| 4091 | 145 |
by (asm_simp_tac (simpset() |
| 1461 | 146 |
addsimps [vv1_def, Let_def, lt_Ord, |
147 |
nested_Least_instance RS conjunct1]) 1); |
|
| 3731 | 148 |
qed "gg1_lepoll_m"; |
| 992 | 149 |
|
150 |
(* ********************************************************************** *) |
|
| 1461 | 151 |
(* Case 2 : lemmas *) |
| 992 | 152 |
(* ********************************************************************** *) |
153 |
||
154 |
(* ********************************************************************** *) |
|
| 1461 | 155 |
(* Case 2 : vv2_subset *) |
| 992 | 156 |
(* ********************************************************************** *) |
157 |
||
| 5137 | 158 |
Goalw [uu_def] "[| b<a; g<a; f`b~=0; f`g~=0; \ |
| 5315 | 159 |
\ y*y <= y; (UN b<a. f`b)=y \ |
160 |
\ |] ==> EX d<a. uu(f,b,g,d) ~= 0"; |
|
| 4091 | 161 |
by (fast_tac (claset() addSIs [not_emptyI] |
| 1461 | 162 |
addSDs [SigmaI RSN (2, subsetD)] |
163 |
addSEs [not_emptyE]) 1); |
|
| 3731 | 164 |
qed "ex_d_uu_not_empty"; |
| 992 | 165 |
|
| 5137 | 166 |
Goal "[| b<a; g<a; f`b~=0; f`g~=0; \ |
| 5315 | 167 |
\ y*y<=y; (UN b<a. f`b)=y |] \ |
168 |
\ ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0"; |
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
169 |
by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1)); |
| 4091 | 170 |
by (fast_tac (claset() addSEs [LeastI, lt_Ord]) 1); |
| 3731 | 171 |
qed "uu_not_empty"; |
| 992 | 172 |
|
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
173 |
Goal "[| r<=A*B; r~=0 |] ==> domain(r)~=0"; |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
174 |
by (Blast_tac 1); |
| 3731 | 175 |
qed "not_empty_rel_imp_domain"; |
| 992 | 176 |
|
| 5137 | 177 |
Goal "[| b<a; g<a; f`b~=0; f`g~=0; \ |
| 5315 | 178 |
\ y*y <= y; (UN b<a. f`b)=y |] \ |
179 |
\ ==> (LEAST d. uu(f,b,g,d) ~= 0) < a"; |
|
| 992 | 180 |
by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1 |
| 1461 | 181 |
THEN REPEAT (assume_tac 1)); |
| 992 | 182 |
by (resolve_tac [Least_le RS lt_trans1] 1 |
| 1461 | 183 |
THEN (REPEAT (ares_tac [lt_Ord] 1))); |
| 3731 | 184 |
qed "Least_uu_not_empty_lt_a"; |
| 992 | 185 |
|
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
186 |
Goal "[| B<=A; a~:B |] ==> B <= A-{a}";
|
| 3731 | 187 |
by (Blast_tac 1); |
188 |
qed "subset_Diff_sing"; |
|
| 992 | 189 |
|
| 1041 | 190 |
(*Could this be proved more directly?*) |
| 5137 | 191 |
Goal "[| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
192 |
by (etac natE 1); |
| 4091 | 193 |
by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [equalityI]) 1); |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
194 |
by (safe_tac (claset() addSIs [equalityI])); |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
195 |
by (rtac ccontr 1); |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
196 |
by (etac (subset_Diff_sing RS subset_imp_lepoll |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
197 |
RSN (2, Diff_sing_lepoll RSN (3, lepoll_trans RS lepoll_trans)) RS |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5315
diff
changeset
|
198 |
succ_lepoll_natE) 1 |
| 1461 | 199 |
THEN REPEAT (assume_tac 1)); |
| 3731 | 200 |
qed "supset_lepoll_imp_eq"; |
| 992 | 201 |
|
| 5268 | 202 |
Goal "[| ALL g<a. ALL d<a. domain(uu(f, b, g, d))~=0 --> \ |
| 1461 | 203 |
\ domain(uu(f, b, g, d)) eqpoll succ(m); \ |
204 |
\ ALL b<a. f`b lepoll succ(m); y*y <= y; \ |
|
205 |
\ (UN b<a. f`b)=y; b<a; g<a; d<a; \ |
|
206 |
\ f`b~=0; f`g~=0; m:nat; s:f`b \ |
|
| 1071 | 207 |
\ |] ==> uu(f, b, g, LEAST d. uu(f,b,g,d)~=0) : f`b -> f`g"; |
208 |
by (dres_inst_tac [("x2","g")] (ospec RS ospec RS mp) 1);
|
|
209 |
by (rtac ([uu_subset1, uu_not_empty] MRS not_empty_rel_imp_domain) 3); |
|
210 |
by (rtac Least_uu_not_empty_lt_a 2 THEN TRYALL assume_tac); |
|
| 992 | 211 |
by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS |
| 1461 | 212 |
(Least_uu_not_empty_lt_a RSN (2, uu_lepoll_m) RSN (2, |
213 |
uu_subset1 RSN (4, rel_is_fun)))] 1 |
|
214 |
THEN TRYALL assume_tac); |
|
| 1071 | 215 |
by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1); |
| 5137 | 216 |
by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset]) 1)); |
| 3731 | 217 |
qed "uu_Least_is_fun"; |
| 992 | 218 |
|
| 5068 | 219 |
Goalw [vv2_def] |
| 1461 | 220 |
"!!a. [| ALL g<a. ALL d<a. domain(uu(f, b, g, d))~=0 --> \ |
221 |
\ domain(uu(f, b, g, d)) eqpoll succ(m); \ |
|
222 |
\ ALL b<a. f`b lepoll succ(m); y*y <= y; \ |
|
223 |
\ (UN b<a. f`b)=y; b<a; g<a; m:nat; s:f`b \ |
|
224 |
\ |] ==> vv2(f,b,g,s) <= f`g"; |
|
|
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5068
diff
changeset
|
225 |
by (split_tac [split_if] 1); |
| 3731 | 226 |
by Safe_tac; |
| 2493 | 227 |
by (etac (uu_Least_is_fun RS apply_type) 1); |
| 4091 | 228 |
by (REPEAT_SOME (fast_tac (claset() addSIs [not_emptyI, singleton_subsetI]))); |
| 3731 | 229 |
qed "vv2_subset"; |
| 992 | 230 |
|
231 |
(* ********************************************************************** *) |
|
| 1461 | 232 |
(* Case 2 : Union of images is the whole "y" *) |
| 992 | 233 |
(* ********************************************************************** *) |
| 5068 | 234 |
Goalw [gg2_def] |
| 1461 | 235 |
"!!a. [| ALL g<a. ALL d<a. domain(uu(f,b,g,d)) ~= 0 --> \ |
236 |
\ domain(uu(f,b,g,d)) eqpoll succ(m); \ |
|
237 |
\ ALL b<a. f`b lepoll succ(m); y*y<=y; \ |
|
| 3840 | 238 |
\ (UN b<a. f`b)=y; Ord(a); m:nat; s:f`b; b<a \ |
| 1461 | 239 |
\ |] ==> (UN g<a++a. gg2(f,a,b,s) ` g) = y"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
240 |
by (dtac sym 1); |
| 1041 | 241 |
by (asm_simp_tac |
| 4091 | 242 |
(simpset() addsimps [UN_oadd, lt_oadd1, |
| 1461 | 243 |
oadd_le_self RS le_imp_not_lt, lt_Ord, |
244 |
odiff_oadd_inverse, ww2_def, |
|
245 |
vv2_subset RS Diff_partition]) 1); |
|
| 3731 | 246 |
qed "UN_gg2_eq"; |
| 1041 | 247 |
|
| 5068 | 248 |
Goal "domain(gg2(f,a,b,s)) = a++a"; |
| 4091 | 249 |
by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg2_def]) 1); |
| 3731 | 250 |
qed "domain_gg2"; |
| 992 | 251 |
|
252 |
(* ********************************************************************** *) |
|
| 1461 | 253 |
(* every value of defined function is less than or equipollent to m *) |
| 992 | 254 |
(* ********************************************************************** *) |
255 |
||
| 5315 | 256 |
Goalw [vv2_def] "[| m:nat; m~=0 |] ==> vv2(f,b,g,s) lepoll m"; |
| 5137 | 257 |
by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 1); |
| 4091 | 258 |
by (fast_tac (claset() |
| 1461 | 259 |
addSDs [le_imp_subset RS subset_imp_lepoll RS lepoll_0_is_0] |
260 |
addSIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, |
|
261 |
not_lt_imp_le RS le_imp_subset RS subset_imp_lepoll, |
|
262 |
nat_into_Ord, nat_1I]) 1); |
|
| 3731 | 263 |
qed "vv2_lepoll"; |
| 992 | 264 |
|
| 5068 | 265 |
Goalw [ww2_def] |
| 5315 | 266 |
"[| ALL b<a. f`b lepoll succ(m); g<a; m:nat; vv2(f,b,g,d) <= f`g |] \ |
267 |
\ ==> ww2(f,b,g,d) lepoll m"; |
|
| 1041 | 268 |
by (excluded_middle_tac "f`g = 0" 1); |
| 4091 | 269 |
by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
270 |
by (dtac ospec 1 THEN (assume_tac 1)); |
| 5137 | 271 |
by (rtac Diff_lepoll 1 THEN (TRYALL assume_tac)); |
272 |
by (asm_simp_tac (simpset() addsimps [vv2_def, not_emptyI]) 1); |
|
| 3731 | 273 |
qed "ww2_lepoll"; |
| 1041 | 274 |
|
| 5068 | 275 |
Goalw [gg2_def] |
| 1461 | 276 |
"!!a. [| ALL g<a. ALL d<a. domain(uu(f,b,g,d)) ~= 0 --> \ |
277 |
\ domain(uu(f,b,g,d)) eqpoll succ(m); \ |
|
278 |
\ ALL b<a. f`b lepoll succ(m); y*y <= y; \ |
|
279 |
\ (UN b<a. f`b)=y; b<a; s:f`b; m:nat; m~= 0; g<a++a \ |
|
| 1041 | 280 |
\ |] ==> gg2(f,a,b,s) ` g lepoll m"; |
| 2469 | 281 |
by (Asm_simp_tac 1); |
| 4091 | 282 |
by (safe_tac (claset() addSEs [lt_oadd_odiff_cases, lt_Ord2])); |
283 |
by (asm_simp_tac (simpset() addsimps [vv2_lepoll]) 1); |
|
284 |
by (asm_simp_tac (simpset() addsimps [ww2_lepoll, vv2_subset]) 1); |
|
| 3731 | 285 |
qed "gg2_lepoll_m"; |
| 992 | 286 |
|
287 |
(* ********************************************************************** *) |
|
| 1461 | 288 |
(* lemma ii *) |
| 992 | 289 |
(* ********************************************************************** *) |
| 5315 | 290 |
Goalw [NN_def] "[| succ(m) : NN(y); y*y <= y; m:nat; m~=0 |] ==> m : NN(y)"; |
| 992 | 291 |
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1)); |
292 |
by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1 |
|
| 1041 | 293 |
THEN (assume_tac 1)); |
| 992 | 294 |
(* case 1 *) |
| 4091 | 295 |
by (asm_full_simp_tac (simpset() addsimps [lesspoll_succ_iff]) 1); |
| 1041 | 296 |
by (res_inst_tac [("x","a++a")] exI 1);
|
| 4091 | 297 |
by (fast_tac (claset() addSIs [Ord_oadd, domain_gg1, UN_gg1_eq, |
| 1461 | 298 |
gg1_lepoll_m]) 1); |
| 992 | 299 |
(* case 2 *) |
300 |
by (REPEAT (eresolve_tac [oexE, conjE] 1)); |
|
| 1041 | 301 |
by (res_inst_tac [("A","f`?B")] not_emptyE 1 THEN (assume_tac 1));
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
302 |
by (rtac CollectI 1); |
|
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
303 |
by (etac succ_natD 1); |
| 992 | 304 |
by (res_inst_tac [("x","a++a")] exI 1);
|
| 1041 | 305 |
by (res_inst_tac [("x","gg2(f,a,b,x)")] exI 1);
|
306 |
(*Calling fast_tac might get rid of the res_inst_tac calls, but it |
|
307 |
is just too slow.*) |
|
| 4091 | 308 |
by (asm_simp_tac (simpset() addsimps |
| 1461 | 309 |
[Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1); |
| 3731 | 310 |
qed "lemma_ii"; |
| 992 | 311 |
|
312 |
||
313 |
(* ********************************************************************** *) |
|
314 |
(* lemma iv - p. 4 : *) |
|
315 |
(* For every set x there is a set y such that x Un (y * y) <= y *) |
|
316 |
(* ********************************************************************** *) |
|
317 |
||
318 |
(* the quantifier ALL looks inelegant but makes the proofs shorter *) |
|
319 |
(* (used only in the following two lemmas) *) |
|
320 |
||
| 5068 | 321 |
Goal "ALL n:nat. rec(n, x, %k r. r Un r*r) <= \ |
| 992 | 322 |
\ rec(succ(n), x, %k r. r Un r*r)"; |
| 4091 | 323 |
by (fast_tac (claset() addIs [rec_succ RS ssubst]) 1); |
| 3731 | 324 |
qed "z_n_subset_z_succ_n"; |
| 992 | 325 |
|
| 5137 | 326 |
Goal "[| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |] \ |
| 5315 | 327 |
\ ==> f(n)<=f(m)"; |
| 2469 | 328 |
by (eres_inst_tac [("P","n le m")] rev_mp 1);
|
| 992 | 329 |
by (res_inst_tac [("P","%z. n le z --> f(n) <= f(z)")] nat_induct 1);
|
| 2469 | 330 |
by (REPEAT (fast_tac le_cs 1)); |
| 3731 | 331 |
qed "le_subsets"; |
| 992 | 332 |
|
| 5137 | 333 |
Goal "[| n le m; m:nat |] ==> \ |
| 1461 | 334 |
\ rec(n, x, %k r. r Un r*r) <= rec(m, x, %k r. r Un r*r)"; |
| 992 | 335 |
by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 |
| 1041 | 336 |
THEN (TRYALL assume_tac)); |
| 992 | 337 |
by (eresolve_tac [Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD] 1 |
| 1041 | 338 |
THEN (assume_tac 1)); |
| 3731 | 339 |
qed "le_imp_rec_subset"; |
| 992 | 340 |
|
| 5068 | 341 |
Goal "EX y. x Un y*y <= y"; |
| 992 | 342 |
by (res_inst_tac [("x","UN n:nat. rec(n, x, %k r. r Un r*r)")] exI 1);
|
| 4152 | 343 |
by Safe_tac; |
| 2493 | 344 |
by (rtac (nat_0I RS UN_I) 1); |
| 2469 | 345 |
by (Asm_simp_tac 1); |
| 992 | 346 |
by (res_inst_tac [("a","succ(n Un na)")] UN_I 1);
|
| 1041 | 347 |
by (eresolve_tac [Un_nat_type RS nat_succI] 1 THEN (assume_tac 1)); |
| 992 | 348 |
by (fast_tac (ZF_cs addIs [le_imp_rec_subset RS subsetD] |
| 1461 | 349 |
addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type] |
| 4091 | 350 |
addSEs [nat_into_Ord] addss (simpset())) 1); |
| 3731 | 351 |
qed "lemma_iv"; |
| 992 | 352 |
|
353 |
(* ********************************************************************** *) |
|
| 1461 | 354 |
(* Rubin & Rubin wrote : *) |
| 992 | 355 |
(* "It follows from (ii) and mathematical induction that if y*y <= y then *) |
| 1461 | 356 |
(* y can be well-ordered" *) |
| 992 | 357 |
|
| 1461 | 358 |
(* In fact we have to prove : *) |
359 |
(* * WO6 ==> NN(y) ~= 0 *) |
|
360 |
(* * reverse induction which lets us infer that 1 : NN(y) *) |
|
361 |
(* * 1 : NN(y) ==> y can be well-ordered *) |
|
| 992 | 362 |
(* ********************************************************************** *) |
363 |
||
364 |
(* ********************************************************************** *) |
|
| 1461 | 365 |
(* WO6 ==> NN(y) ~= 0 *) |
| 992 | 366 |
(* ********************************************************************** *) |
367 |
||
| 5315 | 368 |
Goalw [WO6_def, NN_def] "WO6 ==> NN(y) ~= 0"; |
|
5265
9d1d4c43c76d
Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents:
5241
diff
changeset
|
369 |
by (fast_tac ZF_cs 1); (*SLOW if current claset is used*) |
| 3731 | 370 |
qed "WO6_imp_NN_not_empty"; |
| 992 | 371 |
|
372 |
(* ********************************************************************** *) |
|
| 1461 | 373 |
(* 1 : NN(y) ==> y can be well-ordered *) |
| 992 | 374 |
(* ********************************************************************** *) |
375 |
||
| 5137 | 376 |
Goal "[| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |] \ |
| 5315 | 377 |
\ ==> EX c<a. f`c = {x}";
|
| 4091 | 378 |
by (fast_tac (claset() addSEs [lepoll_1_is_sing]) 1); |
| 992 | 379 |
val lemma1 = result(); |
380 |
||
| 5137 | 381 |
Goal "[| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |] \ |
| 5315 | 382 |
\ ==> f` (LEAST i. f`i = {x}) = {x}";
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
383 |
by (dtac lemma1 1 THEN REPEAT (assume_tac 1)); |
| 4091 | 384 |
by (fast_tac (claset() addSEs [lt_Ord] addIs [LeastI]) 1); |
| 992 | 385 |
val lemma2 = result(); |
386 |
||
| 5137 | 387 |
Goalw [NN_def] "1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
388 |
by (etac CollectE 1); |
| 992 | 389 |
by (REPEAT (eresolve_tac [exE, conjE] 1)); |
390 |
by (res_inst_tac [("x","a")] exI 1);
|
|
391 |
by (res_inst_tac [("x","lam x:y. LEAST i. f`i = {x}")] exI 1);
|
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
392 |
by (rtac conjI 1 THEN (assume_tac 1)); |
| 992 | 393 |
by (res_inst_tac [("d","%i. THE x. x:f`i")] lam_injective 1);
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
394 |
by (dtac lemma1 1 THEN REPEAT (assume_tac 1)); |
| 4091 | 395 |
by (fast_tac (claset() addSEs [Least_le RS lt_trans1 RS ltD, lt_Ord]) 1); |
| 1041 | 396 |
by (resolve_tac [lemma2 RS ssubst] 1 THEN REPEAT (assume_tac 1)); |
| 5505 | 397 |
by (Blast_tac 1); |
| 3731 | 398 |
qed "NN_imp_ex_inj"; |
| 992 | 399 |
|
| 5137 | 400 |
Goal "[| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
401 |
by (dtac NN_imp_ex_inj 1); |
| 4091 | 402 |
by (fast_tac (claset() addSEs [well_ord_Memrel RSN (2, well_ord_rvimage)]) 1); |
| 3731 | 403 |
qed "y_well_ord"; |
| 992 | 404 |
|
405 |
(* ********************************************************************** *) |
|
| 1461 | 406 |
(* reverse induction which lets us infer that 1 : NN(y) *) |
| 992 | 407 |
(* ********************************************************************** *) |
408 |
||
409 |
val [prem1, prem2] = goal thy |
|
| 1461 | 410 |
"[| n:nat; !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ |
411 |
\ ==> n~=0 --> P(n) --> P(1)"; |
|
| 6070 | 412 |
by (rtac (prem1 RS nat_induct) 1); |
| 3731 | 413 |
by (Blast_tac 1); |
| 992 | 414 |
by (excluded_middle_tac "x=0" 1); |
| 3731 | 415 |
by (Blast_tac 2); |
| 4091 | 416 |
by (fast_tac (claset() addSIs [prem2]) 1); |
| 3731 | 417 |
qed "rev_induct_lemma"; |
| 992 | 418 |
|
| 5315 | 419 |
val prems = |
420 |
Goal "[| P(n); n:nat; n~=0; \ |
|
421 |
\ !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ |
|
422 |
\ ==> P(1)"; |
|
| 992 | 423 |
by (resolve_tac [rev_induct_lemma RS impE] 1); |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
424 |
by (etac impE 4 THEN (assume_tac 5)); |
| 992 | 425 |
by (REPEAT (ares_tac prems 1)); |
| 3731 | 426 |
qed "rev_induct"; |
| 992 | 427 |
|
| 5137 | 428 |
Goalw [NN_def] "n:NN(y) ==> n:nat"; |
| 1057 | 429 |
by (etac CollectD1 1); |
| 3731 | 430 |
qed "NN_into_nat"; |
| 992 | 431 |
|
| 5137 | 432 |
Goal "[| n:NN(y); y*y <= y; n~=0 |] ==> 1:NN(y)"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
433 |
by (rtac rev_induct 1 THEN REPEAT (ares_tac [NN_into_nat] 1)); |
|
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
434 |
by (rtac lemma_ii 1 THEN REPEAT (assume_tac 1)); |
| 992 | 435 |
val lemma3 = result(); |
436 |
||
437 |
(* ********************************************************************** *) |
|
| 1461 | 438 |
(* Main theorem "WO6 ==> WO1" *) |
| 992 | 439 |
(* ********************************************************************** *) |
440 |
||
441 |
(* another helpful lemma *) |
|
| 5137 | 442 |
Goalw [NN_def] "0:NN(y) ==> y=0"; |
| 4091 | 443 |
by (fast_tac (claset() addSIs [equalityI] |
| 5315 | 444 |
addSDs [lepoll_0_is_0] addEs [subst]) 1); |
| 3731 | 445 |
qed "NN_y_0"; |
| 992 | 446 |
|
| 5137 | 447 |
Goalw [WO1_def] "WO6 ==> WO1"; |
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
448 |
by (rtac allI 1); |
| 992 | 449 |
by (excluded_middle_tac "A=0" 1); |
| 4091 | 450 |
by (fast_tac (claset() addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2); |
| 992 | 451 |
by (res_inst_tac [("x1","A")] (lemma_iv RS revcut_rl) 1);
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
452 |
by (etac exE 1); |
|
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
1071
diff
changeset
|
453 |
by (dtac WO6_imp_NN_not_empty 1); |
| 992 | 454 |
by (eresolve_tac [Un_subset_iff RS iffD1 RS conjE] 1); |
455 |
by (eres_inst_tac [("A","NN(y)")] not_emptyE 1);
|
|
456 |
by (forward_tac [y_well_ord] 1); |
|
| 4091 | 457 |
by (fast_tac (claset() addEs [well_ord_subset]) 2); |
458 |
by (fast_tac (claset() addSIs [lemma3] addSDs [NN_y_0] addSEs [not_emptyE]) 1); |
|
| 992 | 459 |
qed "WO6_imp_WO1"; |
460 |