author | paulson |
Wed, 23 Jan 2002 11:43:53 +0100 | |
changeset 12836 | 5ef96e63fba6 |
parent 9907 | 473a6604da94 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Zorn.ML |
485 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
485 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
516 | 6 |
Proofs from the paper |
485 | 7 |
Abrial & Laffitte, |
8 |
Towards the Mechanization of the Proofs of Some |
|
9 |
Classical Theorems of Set Theory. |
|
10 |
*) |
|
11 |
||
516 | 12 |
(*** Section 1. Mathematical Preamble ***) |
13 |
||
5774 | 14 |
Goal "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)"; |
2925 | 15 |
by (Blast_tac 1); |
760 | 16 |
qed "Union_lemma0"; |
516 | 17 |
|
5774 | 18 |
Goal "[| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B"; |
2925 | 19 |
by (Blast_tac 1); |
760 | 20 |
qed "Inter_lemma0"; |
516 | 21 |
|
22 |
||
23 |
(*** Section 2. The Transfinite Construction ***) |
|
24 |
||
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
25 |
Goalw [increasing_def] "f: increasing(A) ==> f: Pow(A)->Pow(A)"; |
804 | 26 |
by (etac CollectD1 1); |
760 | 27 |
qed "increasingD1"; |
516 | 28 |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
29 |
Goalw [increasing_def] "[| f: increasing(A); x<=A |] ==> x <= f`x"; |
516 | 30 |
by (eresolve_tac [CollectD2 RS spec RS mp] 1); |
31 |
by (assume_tac 1); |
|
760 | 32 |
qed "increasingD2"; |
516 | 33 |
|
485 | 34 |
(*Introduction rules*) |
516 | 35 |
val [TFin_nextI, Pow_TFin_UnionI] = TFin.intrs; |
9907 | 36 |
bind_thm ("TFin_nextI", TFin_nextI); |
37 |
bind_thm ("Pow_TFin_UnionI", Pow_TFin_UnionI); |
|
38 |
bind_thm ("TFin_UnionI", PowI RS Pow_TFin_UnionI); |
|
485 | 39 |
|
9907 | 40 |
bind_thm ("TFin_is_subset", TFin.dom_subset RS subsetD RS PowD); |
485 | 41 |
|
42 |
||
43 |
(** Structural induction on TFin(S,next) **) |
|
44 |
||
5321 | 45 |
val major::prems = Goal |
485 | 46 |
"[| n: TFin(S,next); \ |
47 |
\ !!x. [| x : TFin(S,next); P(x); next: increasing(S) |] ==> P(next`x); \ |
|
48 |
\ !!Y. [| Y <= TFin(S,next); ALL y:Y. P(y) |] ==> P(Union(Y)) \ |
|
49 |
\ |] ==> P(n)"; |
|
516 | 50 |
by (rtac (major RS TFin.induct) 1); |
4091 | 51 |
by (ALLGOALS (fast_tac (claset() addIs prems))); |
760 | 52 |
qed "TFin_induct"; |
485 | 53 |
|
54 |
||
55 |
(*** Section 3. Some Properties of the Transfinite Construction ***) |
|
56 |
||
803
4c8333ab3eae
changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents:
760
diff
changeset
|
57 |
bind_thm ("increasing_trans", |
1461 | 58 |
TFin_is_subset RSN (3, increasingD2 RSN (2,subset_trans))); |
485 | 59 |
|
60 |
(*Lemma 1 of section 3.1*) |
|
5321 | 61 |
Goal "[| n: TFin(S,next); m: TFin(S,next); \ |
62 |
\ ALL x: TFin(S,next) . x<=m --> x=m | next`x<=m \ |
|
63 |
\ |] ==> n<=m | next`m<=n"; |
|
64 |
by (etac TFin_induct 1); |
|
2925 | 65 |
by (etac Union_lemma0 2); (*or just Blast_tac*) |
2929 | 66 |
by (blast_tac (subset_cs addIs [increasing_trans]) 1); |
760 | 67 |
qed "TFin_linear_lemma1"; |
485 | 68 |
|
69 |
(*Lemma 2 of section 3.2. Interesting in its own right! |
|
70 |
Requires next: increasing(S) in the second induction step. *) |
|
9907 | 71 |
val [major,ninc] = goal (the_context ()) |
485 | 72 |
"[| m: TFin(S,next); next: increasing(S) \ |
73 |
\ |] ==> ALL n: TFin(S,next) . n<=m --> n=m | next`n<=m"; |
|
804 | 74 |
by (rtac (major RS TFin_induct) 1); |
75 |
by (rtac (impI RS ballI) 1); |
|
485 | 76 |
(*case split using TFin_linear_lemma1*) |
77 |
by (res_inst_tac [("n1","n"), ("m1","x")] |
|
78 |
(TFin_linear_lemma1 RS disjE) 1 THEN REPEAT (assume_tac 1)); |
|
79 |
by (dres_inst_tac [("x","n")] bspec 1 THEN assume_tac 1); |
|
2929 | 80 |
by (blast_tac (subset_cs addIs [increasing_trans]) 1); |
485 | 81 |
by (REPEAT (ares_tac [disjI1,equalityI] 1)); |
82 |
(*second induction step*) |
|
804 | 83 |
by (rtac (impI RS ballI) 1); |
84 |
by (rtac (Union_lemma0 RS disjE) 1); |
|
85 |
by (etac disjI2 3); |
|
485 | 86 |
by (REPEAT (ares_tac [disjI1,equalityI] 2)); |
804 | 87 |
by (rtac ballI 1); |
485 | 88 |
by (ball_tac 1); |
89 |
by (set_mp_tac 1); |
|
90 |
by (res_inst_tac [("n1","n"), ("m1","x")] |
|
91 |
(TFin_linear_lemma1 RS disjE) 1 THEN REPEAT (assume_tac 1)); |
|
2929 | 92 |
by (blast_tac subset_cs 1); |
804 | 93 |
by (rtac (ninc RS increasingD2 RS subset_trans RS disjI1) 1); |
485 | 94 |
by (REPEAT (ares_tac [TFin_is_subset] 1)); |
760 | 95 |
qed "TFin_linear_lemma2"; |
485 | 96 |
|
97 |
(*a more convenient form for Lemma 2*) |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
98 |
Goal "[| n<=m; m: TFin(S,next); n: TFin(S,next); next: increasing(S) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
99 |
\ ==> n=m | next`n<=m"; |
804 | 100 |
by (rtac (TFin_linear_lemma2 RS bspec RS mp) 1); |
485 | 101 |
by (REPEAT (assume_tac 1)); |
760 | 102 |
qed "TFin_subsetD"; |
485 | 103 |
|
104 |
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*) |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
105 |
Goal "[| m: TFin(S,next); n: TFin(S,next); next: increasing(S) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
106 |
\ ==> n<=m | m<=n"; |
804 | 107 |
by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1); |
485 | 108 |
by (REPEAT (assume_tac 1) THEN etac disjI2 1); |
2929 | 109 |
by (blast_tac (subset_cs addIs [increasingD2 RS subset_trans, |
1461 | 110 |
TFin_is_subset]) 1); |
760 | 111 |
qed "TFin_subset_linear"; |
485 | 112 |
|
113 |
||
114 |
(*Lemma 3 of section 3.3*) |
|
5321 | 115 |
Goal "[| n: TFin(S,next); m: TFin(S,next); m = next`m |] ==> n<=m"; |
116 |
by (etac TFin_induct 1); |
|
804 | 117 |
by (dtac TFin_subsetD 1); |
485 | 118 |
by (REPEAT (assume_tac 1)); |
4091 | 119 |
by (fast_tac (claset() addEs [ssubst]) 1); |
2929 | 120 |
by (blast_tac (subset_cs addIs [TFin_is_subset]) 1); |
760 | 121 |
qed "equal_next_upper"; |
485 | 122 |
|
123 |
(*Property 3.3 of section 3.3*) |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
124 |
Goal "[| m: TFin(S,next); next: increasing(S) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
125 |
\ ==> m = next`m <-> m = Union(TFin(S,next))"; |
804 | 126 |
by (rtac iffI 1); |
127 |
by (rtac (Union_upper RS equalityI) 1); |
|
128 |
by (rtac (equal_next_upper RS Union_least) 2); |
|
485 | 129 |
by (REPEAT (assume_tac 1)); |
804 | 130 |
by (etac ssubst 1); |
485 | 131 |
by (rtac (increasingD2 RS equalityI) 1 THEN assume_tac 1); |
132 |
by (ALLGOALS |
|
2929 | 133 |
(blast_tac (subset_cs addIs [TFin_UnionI, TFin_nextI, TFin_is_subset]))); |
760 | 134 |
qed "equal_next_Union"; |
485 | 135 |
|
136 |
||
137 |
(*** Section 4. Hausdorff's Theorem: every set contains a maximal chain ***) |
|
138 |
(*** NB: We assume the partial ordering is <=, the subset relation! **) |
|
139 |
||
140 |
(** Defining the "next" operation for Hausdorff's Theorem **) |
|
141 |
||
5067 | 142 |
Goalw [chain_def] "chain(A) <= Pow(A)"; |
804 | 143 |
by (rtac Collect_subset 1); |
760 | 144 |
qed "chain_subset_Pow"; |
485 | 145 |
|
5067 | 146 |
Goalw [super_def] "super(A,c) <= chain(A)"; |
804 | 147 |
by (rtac Collect_subset 1); |
760 | 148 |
qed "super_subset_chain"; |
485 | 149 |
|
5067 | 150 |
Goalw [maxchain_def] "maxchain(A) <= chain(A)"; |
804 | 151 |
by (rtac Collect_subset 1); |
760 | 152 |
qed "maxchain_subset_chain"; |
485 | 153 |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
154 |
Goal "[| ch : (PROD X:Pow(chain(S)) - {0}. X); \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
155 |
\ X : chain(S); X ~: maxchain(S) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
156 |
\ ==> ch ` super(S,X) : super(S,X)"; |
804 | 157 |
by (etac apply_type 1); |
485 | 158 |
by (rewrite_goals_tac [super_def, maxchain_def]); |
2925 | 159 |
by (Blast_tac 1); |
760 | 160 |
qed "choice_super"; |
485 | 161 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
162 |
Goal "[| ch : (PROD X:Pow(chain(S)) - {0}. X); \ |
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
163 |
\ X : chain(S); X ~: maxchain(S) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
164 |
\ ==> ch ` super(S,X) ~= X"; |
804 | 165 |
by (rtac notI 1); |
166 |
by (dtac choice_super 1); |
|
485 | 167 |
by (assume_tac 1); |
168 |
by (assume_tac 1); |
|
4091 | 169 |
by (asm_full_simp_tac (simpset() addsimps [super_def]) 1); |
760 | 170 |
qed "choice_not_equals"; |
485 | 171 |
|
172 |
(*This justifies Definition 4.4*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
173 |
Goal "ch: (PROD X: Pow(chain(S))-{0}. X) ==> \ |
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
174 |
\ EX next: increasing(S). ALL X: Pow(S). \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
175 |
\ next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)"; |
485 | 176 |
by (rtac bexI 1); |
177 |
by (rtac ballI 1); |
|
804 | 178 |
by (rtac beta 1); |
485 | 179 |
by (assume_tac 1); |
804 | 180 |
by (rewtac increasing_def); |
485 | 181 |
by (rtac CollectI 1); |
182 |
by (rtac lam_type 1); |
|
5137 | 183 |
by (Asm_simp_tac 1); |
12836 | 184 |
by (blast_tac (claset() addDs [super_subset_chain RS subsetD, |
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
185 |
chain_subset_Pow RS subsetD, |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
186 |
choice_super]) 1); |
485 | 187 |
(*Now, verify that it increases*) |
5137 | 188 |
by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_refl]) 1); |
4152 | 189 |
by Safe_tac; |
804 | 190 |
by (dtac choice_super 1); |
485 | 191 |
by (REPEAT (assume_tac 1)); |
804 | 192 |
by (rewtac super_def); |
2925 | 193 |
by (Blast_tac 1); |
760 | 194 |
qed "Hausdorff_next_exists"; |
485 | 195 |
|
196 |
(*Lemma 4*) |
|
5268 | 197 |
Goal " [| c: TFin(S,next); \ |
1461 | 198 |
\ ch: (PROD X: Pow(chain(S))-{0}. X); \ |
199 |
\ next: increasing(S); \ |
|
200 |
\ ALL X: Pow(S). next`X = \ |
|
201 |
\ if(X: chain(S)-maxchain(S), ch`super(S,X), X) \ |
|
485 | 202 |
\ |] ==> c: chain(S)"; |
804 | 203 |
by (etac TFin_induct 1); |
485 | 204 |
by (asm_simp_tac |
12836 | 205 |
(simpset() addsimps [chain_subset_Pow RS subsetD RS PowD, |
5137 | 206 |
choice_super RS (super_subset_chain RS subsetD)]) 1); |
804 | 207 |
by (rewtac chain_def); |
2925 | 208 |
by (rtac CollectI 1 THEN Blast_tac 1); |
4152 | 209 |
by Safe_tac; |
485 | 210 |
by (res_inst_tac [("m1","B"), ("n1","Ba")] (TFin_subset_linear RS disjE) 1); |
5469 | 211 |
by (ALLGOALS Fast_tac); (*Blast_tac's slow*) |
760 | 212 |
qed "TFin_chain_lemma4"; |
485 | 213 |
|
5067 | 214 |
Goal "EX c. c : maxchain(S)"; |
485 | 215 |
by (rtac (AC_Pi_Pow RS exE) 1); |
216 |
by (rtac (Hausdorff_next_exists RS bexE) 1); |
|
217 |
by (assume_tac 1); |
|
218 |
by (rename_tac "ch next" 1); |
|
219 |
by (subgoal_tac "Union(TFin(S,next)) : chain(S)" 1); |
|
220 |
by (REPEAT (ares_tac [TFin_chain_lemma4, subset_refl RS TFin_UnionI] 2)); |
|
221 |
by (res_inst_tac [("x", "Union(TFin(S,next))")] exI 1); |
|
804 | 222 |
by (rtac classical 1); |
485 | 223 |
by (subgoal_tac "next ` Union(TFin(S,next)) = Union(TFin(S,next))" 1); |
224 |
by (resolve_tac [equal_next_Union RS iffD2 RS sym] 2); |
|
225 |
by (resolve_tac [subset_refl RS TFin_UnionI] 2); |
|
226 |
by (assume_tac 2); |
|
804 | 227 |
by (rtac refl 2); |
485 | 228 |
by (asm_full_simp_tac |
4091 | 229 |
(simpset() addsimps [subset_refl RS TFin_UnionI RS |
12836 | 230 |
(TFin.dom_subset RS subsetD RS PowD)]) 1); |
485 | 231 |
by (eresolve_tac [choice_not_equals RS notE] 1); |
12836 | 232 |
by (REPEAT (assume_tac 1)); |
760 | 233 |
qed "Hausdorff"; |
485 | 234 |
|
235 |
||
236 |
(*** Section 5. Zorn's Lemma: if all chains in S have upper bounds in S |
|
237 |
then S contains a maximal element ***) |
|
238 |
||
239 |
(*Used in the proof of Zorn's Lemma*) |
|
5067 | 240 |
Goalw [chain_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
241 |
"[| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)"; |
2925 | 242 |
by (Blast_tac 1); |
760 | 243 |
qed "chain_extend"; |
485 | 244 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
245 |
Goal "ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z"; |
485 | 246 |
by (resolve_tac [Hausdorff RS exE] 1); |
4091 | 247 |
by (asm_full_simp_tac (simpset() addsimps [maxchain_def]) 1); |
485 | 248 |
by (rename_tac "c" 1); |
249 |
by (res_inst_tac [("x", "Union(c)")] bexI 1); |
|
5469 | 250 |
by (Blast_tac 2); |
4152 | 251 |
by Safe_tac; |
485 | 252 |
by (rename_tac "z" 1); |
804 | 253 |
by (rtac classical 1); |
485 | 254 |
by (subgoal_tac "cons(z,c): super(S,c)" 1); |
4091 | 255 |
by (blast_tac (claset() addEs [equalityE]) 1); |
804 | 256 |
by (rewtac super_def); |
4152 | 257 |
by Safe_tac; |
4091 | 258 |
by (fast_tac (claset() addEs [chain_extend]) 1); |
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
259 |
by (fast_tac (claset() addEs [equalityE]) 1); |
760 | 260 |
qed "Zorn"; |
485 | 261 |
|
262 |
||
263 |
(*** Section 6. Zermelo's Theorem: every set can be well-ordered ***) |
|
264 |
||
265 |
(*Lemma 5*) |
|
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
266 |
Goal "[| n: TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) : Z |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
267 |
\ ==> ALL m:Z. n<=m"; |
5321 | 268 |
by (etac TFin_induct 1); |
2925 | 269 |
by (Blast_tac 2); (*second induction step is easy*) |
804 | 270 |
by (rtac ballI 1); |
271 |
by (rtac (bspec RS TFin_subsetD RS disjE) 1); |
|
485 | 272 |
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD])); |
273 |
by (subgoal_tac "x = Inter(Z)" 1); |
|
2925 | 274 |
by (Blast_tac 1); |
275 |
by (Blast_tac 1); |
|
760 | 276 |
qed "TFin_well_lemma5"; |
485 | 277 |
|
278 |
(*Well-ordering of TFin(S,next)*) |
|
5137 | 279 |
Goal "[| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z"; |
804 | 280 |
by (rtac classical 1); |
485 | 281 |
by (subgoal_tac "Z = {Union(TFin(S,next))}" 1); |
4091 | 282 |
by (asm_simp_tac (simpset() addsimps [Inter_singleton]) 1); |
804 | 283 |
by (etac equal_singleton 1); |
284 |
by (rtac (Union_upper RS equalityI) 1); |
|
285 |
by (rtac (subset_refl RS TFin_UnionI RS TFin_well_lemma5 RS bspec) 2); |
|
485 | 286 |
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD])); |
760 | 287 |
qed "well_ord_TFin_lemma"; |
485 | 288 |
|
289 |
(*This theorem just packages the previous result*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
290 |
Goal "next: increasing(S) ==> \ |
485 | 291 |
\ well_ord(TFin(S,next), Subset_rel(TFin(S,next)))"; |
804 | 292 |
by (rtac well_ordI 1); |
485 | 293 |
by (rewrite_goals_tac [Subset_rel_def, linear_def]); |
294 |
(*Prove the linearity goal first*) |
|
295 |
by (REPEAT (rtac ballI 2)); |
|
296 |
by (excluded_middle_tac "x=y" 2); |
|
2925 | 297 |
by (Blast_tac 3); |
485 | 298 |
(*The x~=y case remains*) |
299 |
by (res_inst_tac [("n1","x"), ("m1","y")] |
|
300 |
(TFin_subset_linear RS disjE) 2 THEN REPEAT (assume_tac 2)); |
|
2925 | 301 |
by (Blast_tac 2); |
302 |
by (Blast_tac 2); |
|
485 | 303 |
(*Now prove the well_foundedness goal*) |
804 | 304 |
by (rtac wf_onI 1); |
7499 | 305 |
by (ftac well_ord_TFin_lemma 1 THEN assume_tac 1); |
485 | 306 |
by (dres_inst_tac [("x","Inter(Z)")] bspec 1 THEN assume_tac 1); |
2925 | 307 |
by (Blast_tac 1); |
760 | 308 |
qed "well_ord_TFin"; |
485 | 309 |
|
310 |
(** Defining the "next" operation for Zermelo's Theorem **) |
|
311 |
||
312 |
(*This justifies Definition 6.1*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
313 |
Goal "ch: (PROD X: Pow(S)-{0}. X) ==> \ |
1461 | 314 |
\ EX next: increasing(S). ALL X: Pow(S). \ |
485 | 315 |
\ next`X = if(X=S, S, cons(ch`(S-X), X))"; |
316 |
by (rtac bexI 1); |
|
317 |
by (rtac ballI 1); |
|
804 | 318 |
by (rtac beta 1); |
485 | 319 |
by (assume_tac 1); |
804 | 320 |
by (rewtac increasing_def); |
485 | 321 |
by (rtac CollectI 1); |
322 |
by (rtac lam_type 1); |
|
323 |
(*Verify that it increases*) |
|
804 | 324 |
by (rtac allI 2); |
325 |
by (rtac impI 2); |
|
5137 | 326 |
by (asm_simp_tac (simpset() addsimps [Pow_iff, subset_consI, subset_refl]) 2); |
485 | 327 |
(*Type checking is surprisingly hard!*) |
5137 | 328 |
by (asm_simp_tac |
329 |
(simpset() addsimps [Pow_iff, cons_subset_iff, subset_refl]) 1); |
|
4091 | 330 |
by (blast_tac (claset() addSIs [choice_Diff RS DiffD1]) 1); |
760 | 331 |
qed "Zermelo_next_exists"; |
485 | 332 |
|
333 |
||
334 |
(*The construction of the injection*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
335 |
Goal "[| ch: (PROD X: Pow(S)-{0}. X); \ |
5666
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
336 |
\ next: increasing(S); \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
337 |
\ ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) |] \ |
822db50b3ec5
fixed some indenting; changed a VERY slow blast_tac to fast_tac
paulson
parents:
5469
diff
changeset
|
338 |
\ ==> (lam x:S. Union({y: TFin(S,next). x~: y})) \ |
485 | 339 |
\ : inj(S, TFin(S,next) - {S})"; |
340 |
by (res_inst_tac [("d", "%y. ch`(S-y)")] lam_injective 1); |
|
341 |
by (rtac DiffI 1); |
|
342 |
by (resolve_tac [Collect_subset RS TFin_UnionI] 1); |
|
4091 | 343 |
by (blast_tac (claset() addSIs [Collect_subset RS TFin_UnionI] |
2469 | 344 |
addEs [equalityE]) 1); |
485 | 345 |
by (subgoal_tac "x ~: Union({y: TFin(S,next). x~: y})" 1); |
4091 | 346 |
by (blast_tac (claset() addEs [equalityE]) 2); |
485 | 347 |
by (subgoal_tac "Union({y: TFin(S,next). x~: y}) ~= S" 1); |
4091 | 348 |
by (blast_tac (claset() addEs [equalityE]) 2); |
485 | 349 |
(*For proving x : next`Union(...); |
350 |
Abrial & Laffitte's justification appears to be faulty.*) |
|
351 |
by (subgoal_tac "~ next ` Union({y: TFin(S,next). x~: y}) <= \ |
|
352 |
\ Union({y: TFin(S,next). x~: y})" 1); |
|
353 |
by (asm_simp_tac |
|
4091 | 354 |
(simpset() delsimps [Union_iff] |
2493 | 355 |
addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset, |
1461 | 356 |
Pow_iff, cons_subset_iff, subset_refl, |
5137 | 357 |
choice_Diff RS DiffD2]) 2); |
485 | 358 |
by (subgoal_tac "x : next ` Union({y: TFin(S,next). x~: y})" 1); |
2929 | 359 |
by (blast_tac (subset_cs addSIs [Collect_subset RS TFin_UnionI, TFin_nextI]) 2); |
485 | 360 |
(*End of the lemmas!*) |
361 |
by (asm_full_simp_tac |
|
4091 | 362 |
(simpset() addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset, |
5137 | 363 |
Pow_iff, cons_subset_iff, subset_refl]) 1); |
485 | 364 |
by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1)); |
760 | 365 |
qed "choice_imp_injection"; |
485 | 366 |
|
367 |
(*The wellordering theorem*) |
|
5067 | 368 |
Goal "EX r. well_ord(S,r)"; |
485 | 369 |
by (rtac (AC_Pi_Pow RS exE) 1); |
370 |
by (rtac (Zermelo_next_exists RS bexE) 1); |
|
371 |
by (assume_tac 1); |
|
804 | 372 |
by (rtac exI 1); |
373 |
by (rtac well_ord_rvimage 1); |
|
374 |
by (etac well_ord_TFin 2); |
|
485 | 375 |
by (resolve_tac [choice_imp_injection RS inj_weaken_type] 1); |
376 |
by (REPEAT (ares_tac [Diff_subset] 1)); |
|
760 | 377 |
qed "AC_well_ord"; |