Isar version of AC
authorpaulson
Wed, 16 Jan 2002 17:52:06 +0100
changeset 12776 249600a63ba9
parent 12775 1748c16c2df3
child 12777 70b2651af635
Isar version of AC
src/ZF/AC/AC15_WO6.ML
src/ZF/AC/AC15_WO6.thy
src/ZF/AC/AC16_WO4.ML
src/ZF/AC/AC16_WO4.thy
src/ZF/AC/AC16_lemmas.ML
src/ZF/AC/AC16_lemmas.thy
src/ZF/AC/AC17_AC1.ML
src/ZF/AC/AC17_AC1.thy
src/ZF/AC/AC18_AC19.ML
src/ZF/AC/AC18_AC19.thy
src/ZF/AC/AC1_WO2.ML
src/ZF/AC/AC1_WO2.thy
src/ZF/AC/AC7_AC9.ML
src/ZF/AC/AC_Equiv.ML
src/ZF/AC/AC_Equiv.thy
src/ZF/AC/Cardinal_aux.ML
src/ZF/AC/Cardinal_aux.thy
src/ZF/AC/DC.ML
src/ZF/AC/DC.thy
src/ZF/AC/DC_lemmas.ML
src/ZF/AC/DC_lemmas.thy
src/ZF/AC/HH.ML
src/ZF/AC/HH.thy
src/ZF/AC/Hartog.ML
src/ZF/AC/Hartog.thy
src/ZF/AC/ROOT.ML
src/ZF/AC/WO1_AC.ML
src/ZF/AC/WO1_AC.thy
src/ZF/AC/WO1_WO7.ML
src/ZF/AC/WO1_WO7.thy
src/ZF/AC/WO2_AC16.ML
src/ZF/AC/WO2_AC16.thy
src/ZF/AC/WO6_WO1.ML
src/ZF/AC/WO6_WO1.thy
src/ZF/AC/WO_AC.ML
src/ZF/AC/WO_AC.thy
src/ZF/AC/recfunAC16.ML
src/ZF/AC/recfunAC16.thy
src/ZF/AC/rel_is_fun.ML
src/ZF/AC/rel_is_fun.thy
src/ZF/CardinalArith.thy
src/ZF/IsaMakefile
src/ZF/Main.thy
--- a/src/ZF/AC/AC15_WO6.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,292 +0,0 @@
-(*  Title:      ZF/AC/AC15_WO6.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.
-We need the following:
-
-WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6
-
-In order to add the formulations AC13 and AC14 we need:
-
-AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15
-
-or
-
-AC1 ==> AC13(1);  AC13(m) ==> AC13(n) ==> AC14 ==> AC15    (m le n)
-
-So we don't have to prove all implications of both cases.
-Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as
-Rubin & Rubin do.
-*)
-
-(* ********************************************************************** *)
-(* Lemmas used in the proofs in which the conclusion is AC13, AC14        *)
-(* or AC15                                                                *)
-(*  - cons_times_nat_not_Finite                                           *)
-(*  - ex_fun_AC13_AC15                                                    *)
-(* ********************************************************************** *)
-
-Goalw [lepoll_def] "A\\<noteq>0 ==> B lepoll A*B";
-by (etac not_emptyE 1);
-by (res_inst_tac [("x","\\<lambda>z \\<in> B. <x,z>")] exI 1);
-by (fast_tac (claset() addSIs [snd_conv, lam_injective]) 1);
-qed "lepoll_Sigma";
-
-Goal "0\\<notin>A ==> \\<forall>B \\<in> {cons(0,x*nat). x \\<in> A}. ~Finite(B)";
-by (rtac ballI 1);
-by (etac RepFunE 1);
-by (hyp_subst_tac 1);
-by (rtac notI 1);
-by (dresolve_tac [subset_consI RS subset_imp_lepoll RS lepoll_Finite] 1);
-by (resolve_tac [lepoll_Sigma RS lepoll_Finite RS (nat_not_Finite RS notE)] 1
-        THEN (assume_tac 2));
-by (Fast_tac 1);
-qed "cons_times_nat_not_Finite";
-
-Goal "[| Union(C)=A; a \\<in> A |] ==> \\<exists>B \\<in> C. a \\<in> B & B \\<subseteq> A";
-by (Fast_tac 1);
-val lemma1 = result();
-
-Goalw [pairwise_disjoint_def]
-        "[| pairwise_disjoint(A); B \\<in> A; C \\<in> A; a \\<in> B; a \\<in> C |] ==> B=C";
-by (dtac IntI 1 THEN (assume_tac 1));
-by (dres_inst_tac [("A","B Int C")] not_emptyI 1);
-by (Fast_tac 1);
-val lemma2 = result();
-
-Goalw [sets_of_size_between_def]
-        "\\<forall>B \\<in> {cons(0, x*nat). x \\<in> A}. pairwise_disjoint(f`B) &  \
-\               sets_of_size_between(f`B, 2, n) & Union(f`B)=B  \
-\       ==> \\<forall>B \\<in> A. \\<exists>! u. u \\<in> f`cons(0, B*nat) & u \\<subseteq> cons(0, B*nat) &  \
-\               0 \\<in> u & 2 lepoll u & u lepoll n";
-by (rtac ballI 1);
-by (etac ballE 1);
-by (Fast_tac 2);
-by (REPEAT (etac conjE 1));
-by (dresolve_tac [consI1 RSN (2, lemma1)] 1);
-by (etac bexE 1);
-by (rtac ex1I 1);
-by (Fast_tac 1);
-by (REPEAT (etac conjE 1));
-by (rtac lemma2 1 THEN (REPEAT (assume_tac 1)));
-val lemma3 = result();
-
-Goalw [lepoll_def] "[| A lepoll i; Ord(i) |] ==> {P(a). a \\<in> A} lepoll i";
-by (etac exE 1);
-by (res_inst_tac
-        [("x", "\\<lambda>x \\<in> RepFun(A, P). LEAST j. \\<exists>a \\<in> A. x=P(a) & f`a=j")] exI 1);
-by (res_inst_tac [("d", "%y. P(converse(f)`y)")] lam_injective 1);
-by (etac RepFunE 1);
-by (forward_tac [inj_is_fun RS apply_type] 1 THEN (assume_tac 1));
-by (fast_tac (claset() addIs [LeastI2]
-                addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);
-by (etac RepFunE 1);
-by (rtac LeastI2 1);
-by (Fast_tac 1);
-by (fast_tac (claset() addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);
-by (fast_tac (claset() addEs [sym, left_inverse RS ssubst]) 1);
-val lemma4 = result();
-
-Goal "[| n \\<in> nat; B \\<in> A; u(B) \\<subseteq> cons(0, B*nat); 0 \\<in> u(B); 2 lepoll u(B);  \
-\       u(B) lepoll succ(n) |]  \
-\       ==> (\\<lambda>x \\<in> A. {fst(x). x \\<in> u(x)-{0}})`B \\<noteq> 0 &  \
-\               (\\<lambda>x \\<in> A. {fst(x). x \\<in> u(x)-{0}})`B \\<subseteq> B &  \
-\               (\\<lambda>x \\<in> A. {fst(x). x \\<in> u(x)-{0}})`B lepoll n";
-by (Asm_simp_tac 1);
-by (rtac conjI 1);
-by (fast_tac (empty_cs addSDs [RepFun_eq_0_iff RS iffD1]
-                addDs [lepoll_Diff_sing]
-                addEs [lepoll_trans RS succ_lepoll_natE, ssubst]
-                addSIs [notI, lepoll_refl, nat_0I]) 1);
-by (rtac conjI 1);
-by (fast_tac (claset() addSIs [fst_type] addSEs [consE]) 1);
-by (fast_tac (claset() addSEs [equalityE,
-                Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1);
-val lemma5 = result();
-
-Goal "[| \\<exists>f. \\<forall>B \\<in> {cons(0, x*nat). x \\<in> A}.  \
-\               pairwise_disjoint(f`B) &  \
-\               sets_of_size_between(f`B, 2, succ(n)) &  \
-\               Union(f`B)=B; n \\<in> nat |]  \
-\       ==> \\<exists>f. \\<forall>B \\<in> A. f`B \\<noteq> 0 & f`B \\<subseteq> B & f`B lepoll n";
-by (fast_tac (empty_cs addSDs [lemma3, theI] addDs [bspec]
-                addSEs [exE, conjE]
-                addIs [exI, ballI, lemma5]) 1);
-qed "ex_fun_AC13_AC15";
-
-(* ********************************************************************** *)
-(* The target proofs                                                      *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* AC10(n) ==> AC11                                                       *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "[| n \\<in> nat; 1 le n; AC10(n) |] ==> AC11";
-by (rtac bexI 1 THEN (assume_tac 2));
-by (Fast_tac 1);
-qed "AC10_AC11";
-
-(* ********************************************************************** *)
-(* AC11 ==> AC12                                                          *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC11 ==> AC12";
-by (fast_tac (FOL_cs addSEs [bexE] addIs [bexI]) 1);
-qed "AC11_AC12";
-
-(* ********************************************************************** *)
-(* AC12 ==> AC15                                                          *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC12 ==> AC15";
-by Safe_tac;
-by (etac allE 1);
-by (etac impE 1);
-by (etac cons_times_nat_not_Finite 1);
-by (fast_tac (claset() addSIs [ex_fun_AC13_AC15]) 1);
-qed "AC12_AC15";
-
-(* ********************************************************************** *)
-(* AC15 ==> WO6                                                           *)
-(* ********************************************************************** *)
-
-Goal "Ord(x) ==> (\\<Union>a<x. F(a)) = (\\<Union>a \\<in> x. F(a))";
-by (fast_tac (claset() addSIs [ltI] addSDs [ltD]) 1);
-qed "OUN_eq_UN";
-
-val [prem] = goal thy "\\<forall>x \\<in> Pow(A)-{0}. f`x\\<noteq>0 & f`x \\<subseteq> x & f`x lepoll m ==>  \
-\       (\\<Union>i<LEAST x. HH(f,A,x)={A}. HH(f,A,i)) = A";
-by (simp_tac (simpset() addsimps [Ord_Least RS OUN_eq_UN]) 1);
-by (rtac equalityI 1);
-by (fast_tac (claset() addSDs [less_Least_subset_x]) 1);
-by (fast_tac (claset() addSDs [prem RS bspec]
-                addSIs [f_subsets_imp_UN_HH_eq_x RS (Diff_eq_0_iff RS iffD1)]) 1);
-val lemma1 = result();
-
-val [prem] = goal thy "\\<forall>x \\<in> Pow(A)-{0}. f`x\\<noteq>0 & f`x \\<subseteq> x & f`x lepoll m ==>  \
-\       \\<forall>x<LEAST x. HH(f,A,x)={A}. HH(f,A,x) lepoll m";
-by (rtac oallI 1);
-by (dresolve_tac [ltD RS less_Least_subset_x] 1);
-by (ftac HH_subset_imp_eq 1);
-by (etac ssubst 1);
-by (fast_tac (claset() addIs [prem RS ballE]
-                addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
-val lemma2 = result();
-
-Goalw [AC15_def, WO6_def] "AC15 ==> WO6";
-by (rtac allI 1);
-by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
-by (etac impE 1);
-by (Fast_tac 1);
-by (REPEAT (eresolve_tac [bexE,conjE,exE] 1));
-by (rtac bexI 1 THEN (assume_tac 2));
-by (rtac conjI 1 THEN (assume_tac 1));
-by (res_inst_tac [("x","LEAST i. HH(f,A,i)={A}")] exI 1);
-by (res_inst_tac [("x","\\<lambda>j \\<in> (LEAST i. HH(f,A,i)={A}). HH(f,A,j)")] exI 1);
-by (Asm_full_simp_tac 1);
-by (fast_tac (claset() addSIs [Ord_Least, lam_type RS domain_of_fun]
-                addSEs [less_Least_subset_x, lemma1, lemma2]) 1);
-qed "AC15_WO6";
-
-
-(* ********************************************************************** *)
-(* The proof needed in the first case, not in the second                  *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* AC10(n) ==> AC13(n-1)  if 2 le n                                       *)
-(*                                                                        *)
-(* Because of the change to the formal definition of AC10(n) we prove     *)
-(* the following obviously equivalent theorem \\<in>                           *)
-(* AC10(n) implies AC13(n) for (1 le n)                                   *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "[| n \\<in> nat; 1 le n; AC10(n) |] ==> AC13(n)";
-by Safe_tac;
-by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE),
-                                ex_fun_AC13_AC15]) 1);
-qed "AC10_AC13";
-
-(* ********************************************************************** *)
-(* The proofs needed in the second case, not in the first                 *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* AC1 ==> AC13(1)                                                        *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC1 ==> AC13(1)";
-by (rtac allI 1);
-by (etac allE 1);
-by (rtac impI 1);
-by (mp_tac 1);
-by (etac exE 1);
-by (res_inst_tac [("x","\\<lambda>x \\<in> A. {f`x}")] exI 1);
-by (asm_simp_tac (simpset() addsimps
-		  [singleton_eqpoll_1 RS eqpoll_imp_lepoll,
-		   singletonI RS not_emptyI]) 1);
-qed "AC1_AC13";
-
-(* ********************************************************************** *)
-(* AC13(m) ==> AC13(n) for m \\<subseteq> n                                         *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "[| m le n; AC13(m) |] ==> AC13(n)";
-by (dtac le_imp_lepoll 1);
-by (fast_tac (claset() addSEs [lepoll_trans]) 1);
-qed "AC13_mono";
-
-(* ********************************************************************** *)
-(* The proofs necessary for both cases                                    *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* AC13(n) ==> AC14  if 1 \\<subseteq> n                                            *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "[| n \\<in> nat; 1 le n; AC13(n) |] ==> AC14";
-by (fast_tac (FOL_cs addIs [bexI]) 1);
-qed "AC13_AC14";
-
-(* ********************************************************************** *)
-(* AC14 ==> AC15                                                          *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC14 ==> AC15";
-by (Fast_tac 1);
-qed "AC14_AC15";
-
-(* ********************************************************************** *)
-(* The redundant proofs; however cited by Rubin & Rubin                   *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* AC13(1) ==> AC1                                                        *)
-(* ********************************************************************** *)
-
-Goal "[| A\\<noteq>0; A lepoll 1 |] ==> \\<exists>a. A={a}";
-by (fast_tac (claset() addSEs [not_emptyE, lepoll_1_is_sing]) 1);
-qed "lemma_aux";
-
-Goal "\\<forall>B \\<in> A. f(B)\\<noteq>0 & f(B)<=B & f(B) lepoll 1  \
-\     ==> (\\<lambda>x \\<in> A. THE y. f(x)={y}) \\<in> (\\<Pi>X \\<in> A. X)";
-by (rtac lam_type 1);
-by (dtac bspec 1 THEN (assume_tac 1));
-by (REPEAT (etac conjE 1));
-by (eresolve_tac [lemma_aux RS exE] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [the_element]) 1);
-val lemma = result();
-
-Goalw AC_defs "AC13(1) ==> AC1";
-by (fast_tac (claset() addSEs [lemma]) 1);
-qed "AC13_AC1";
-
-(* ********************************************************************** *)
-(* AC11 ==> AC14                                                          *)
-(* ********************************************************************** *)
-
-Goalw [AC11_def, AC14_def] "AC11 ==> AC14";
-by (fast_tac (claset() addSIs [AC10_AC13]) 1);
-qed "AC11_AC14";
--- a/src/ZF/AC/AC15_WO6.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC15_WO6.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,290 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/AC15_WO6.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.
+We need the following:
+
+WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6
+
+In order to add the formulations AC13 and AC14 we need:
+
+AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15
+
+or
+
+AC1 ==> AC13(1);  AC13(m) ==> AC13(n) ==> AC14 ==> AC15    (m\<le>n)
+
+So we don't have to prove all implications of both cases.
+Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as
+Rubin & Rubin do.
+*)
+
+theory AC15_WO6 = HH + Cardinal_aux:
+
+
+(* ********************************************************************** *)
+(* Lemmas used in the proofs in which the conclusion is AC13, AC14        *)
+(* or AC15                                                                *)
+(*  - cons_times_nat_not_Finite                                           *)
+(*  - ex_fun_AC13_AC15                                                    *)
+(* ********************************************************************** *)
+
+lemma lepoll_Sigma: "A\<noteq>0 ==> B \<lesssim> A*B"
+apply (unfold lepoll_def)
+apply (erule not_emptyE)
+apply (rule_tac x = "\<lambda>z \<in> B. <x,z>" in exI)
+apply (fast intro!: snd_conv lam_injective)
+done
+
+lemma cons_times_nat_not_Finite:
+     "0\<notin>A ==> \<forall>B \<in> {cons(0,x*nat). x \<in> A}. ~Finite(B)"
+apply clarify 
+apply (drule subset_consI [THEN subset_imp_lepoll, THEN lepoll_Finite])
+apply (rule nat_not_Finite [THEN notE] )
+apply (subgoal_tac "x ~= 0")
+apply (blast intro: lepoll_Sigma [THEN lepoll_Finite] , blast) 
+done
+
+lemma lemma1: "[| Union(C)=A; a \<in> A |] ==> \<exists>B \<in> C. a \<in> B & B \<subseteq> A"
+by fast
+
+lemma lemma2: 
+        "[| pairwise_disjoint(A); B \<in> A; C \<in> A; a \<in> B; a \<in> C |] ==> B=C"
+by (unfold pairwise_disjoint_def, blast) 
+
+lemma lemma3: 
+        "\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) &   
+                sets_of_size_between(f`B, 2, n) & Union(f`B)=B   
+        ==> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &   
+                0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"
+apply (unfold sets_of_size_between_def)
+apply (rule ballI)
+apply (erule ballE)
+prefer 2 apply blast 
+apply (blast dest: lemma1 intro!: lemma2) 
+done
+
+lemma lemma4: "[| A \<lesssim> i; Ord(i) |] ==> {P(a). a \<in> A} \<lesssim> i"
+apply (unfold lepoll_def)
+apply (erule exE)
+apply (rule_tac x = "\<lambda>x \<in> RepFun(A,P). LEAST j. \<exists>a\<in>A. x=P(a) & f`a=j" 
+       in exI)
+apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective)
+apply (erule RepFunE)
+apply (frule inj_is_fun [THEN apply_type], assumption)
+apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type])
+apply (erule RepFunE)
+apply (rule LeastI2)
+  apply fast
+ apply (fast elim!: Ord_in_Ord inj_is_fun [THEN apply_type])
+apply (fast elim: sym left_inverse [THEN ssubst])
+done
+
+lemma lemma5_1:
+     "[| B \<in> A; 2 \<lesssim> u(B) |] ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<noteq> 0"
+apply simp
+apply (fast dest: lepoll_Diff_sing 
+            elim: lepoll_trans [THEN succ_lepoll_natE] ssubst
+            intro!: lepoll_refl)
+done
+
+lemma lemma5_2:
+     "[|  B \<in> A; u(B) \<subseteq> cons(0, B*nat) |]   
+      ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<subseteq> B"
+apply auto 
+done
+
+lemma lemma5_3:
+     "[| n \<in> nat; B \<in> A; 0 \<in> u(B); u(B) \<lesssim> succ(n) |]   
+      ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<lesssim> n"
+apply simp
+apply (fast elim!: Diff_lepoll [THEN lemma4 [OF _ nat_into_Ord]])
+done
+
+lemma ex_fun_AC13_AC15:
+     "[| \<forall>B \<in> {cons(0, x*nat). x \<in> A}.   
+                pairwise_disjoint(f`B) &   
+                sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B; 
+         n \<in> nat |]   
+      ==> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"
+by (fast del: subsetI notI
+	 dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)
+
+
+(* ********************************************************************** *)
+(* The target proofs                                                      *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* AC10(n) ==> AC11                                                       *)
+(* ********************************************************************** *)
+
+lemma AC10_AC11: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC11"
+by (unfold AC10_def AC11_def, blast)
+
+(* ********************************************************************** *)
+(* AC11 ==> AC12                                                          *)
+(* ********************************************************************** *)
+
+lemma AC11_AC12: "AC11 ==> AC12"
+by (unfold AC10_def AC11_def AC11_def AC12_def, blast)
+
+(* ********************************************************************** *)
+(* AC12 ==> AC15                                                          *)
+(* ********************************************************************** *)
+
+lemma AC12_AC15: "AC12 ==> AC15"
+apply (unfold AC12_def AC15_def)
+apply (blast del: ballI 
+             intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15)
+done
 
-AC15_WO6 = HH
+(* ********************************************************************** *)
+(* AC15 ==> WO6                                                           *)
+(* ********************************************************************** *)
+
+lemma OUN_eq_UN: "Ord(x) ==> (\<Union>a<x. F(a)) = (\<Union>a \<in> x. F(a))"
+by (fast intro!: ltI dest!: ltD)
+
+lemma lemma1:
+     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
+      ==> (\<Union>i<LEAST x. HH(f,A,x)={A}. HH(f,A,i)) = A"
+apply (simp add: Ord_Least [THEN OUN_eq_UN])
+apply (rule equalityI)
+apply (fast dest!: less_Least_subset_x)
+apply (blast del: subsetI 
+           intro!: f_subsets_imp_UN_HH_eq_x [THEN Diff_eq_0_iff [THEN iffD1]])
+done
+
+lemma lemma2:
+     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
+      ==> \<forall>x < (LEAST x. HH(f,A,x)={A}). HH(f,A,x) \<lesssim> m"
+apply (rule oallI)
+apply (drule ltD [THEN less_Least_subset_x])
+apply (frule HH_subset_imp_eq)
+apply (erule ssubst)
+apply (blast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2])
+	(*but can't use del: DiffE despite the obvious conflictc*)
+done
+
+lemma AC15_WO6: "AC15 ==> WO6"
+apply (unfold AC15_def WO6_def)
+apply (rule allI)
+apply (erule_tac x = "Pow (A) -{0}" in allE)
+apply (erule impE, fast)
+apply (elim bexE conjE exE)
+apply (rule bexI)
+apply (rule conjI, assumption)
+apply (rule_tac x = "LEAST i. HH (f,A,i) ={A}" in exI)
+apply (rule_tac x = "\<lambda>j \<in> (LEAST i. HH (f,A,i) ={A}) . HH (f,A,j) " in exI)
+apply simp
+apply (fast intro!: Ord_Least lam_type [THEN domain_of_fun]
+            elim!: less_Least_subset_x lemma1 lemma2, assumption); 
+done
+
+
+(* ********************************************************************** *)
+(* The proof needed in the first case, not in the second                  *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* AC10(n) ==> AC13(n-1)  if 2\<le>n                                       *)
+(*                                                                        *)
+(* Because of the change to the formal definition of AC10(n) we prove     *)
+(* the following obviously equivalent theorem \<in>                           *)
+(* AC10(n) implies AC13(n) for (1\<le>n)                                   *)
+(* ********************************************************************** *)
+
+lemma AC10_AC13: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC13(n)"
+apply (unfold AC10_def AC13_def, safe)
+apply (erule allE) 
+apply (erule impE [OF _ cons_times_nat_not_Finite], assumption); 
+apply (fast elim!: impE [OF _ cons_times_nat_not_Finite] 
+            dest!: ex_fun_AC13_AC15)
+done
+
+(* ********************************************************************** *)
+(* The proofs needed in the second case, not in the first                 *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* AC1 ==> AC13(1)                                                        *)
+(* ********************************************************************** *)
+
+lemma AC1_AC13: "AC1 ==> AC13(1)"
+apply (unfold AC1_def AC13_def)
+apply (rule allI)
+apply (erule allE)
+apply (rule impI)
+apply (drule mp, assumption) 
+apply (elim exE)
+apply (rule_tac x = "\<lambda>x \<in> A. {f`x}" in exI)
+apply (simp add: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll])
+done
+
+(* ********************************************************************** *)
+(* AC13(m) ==> AC13(n) for m \<subseteq> n                                         *)
+(* ********************************************************************** *)
+
+lemma AC13_mono: "[| m\<le>n; AC13(m) |] ==> AC13(n)"
+apply (unfold AC13_def)
+apply (drule le_imp_lepoll)
+apply (fast elim!: lepoll_trans)
+done
+
+(* ********************************************************************** *)
+(* The proofs necessary for both cases                                    *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* AC13(n) ==> AC14  if 1 \<subseteq> n                                            *)
+(* ********************************************************************** *)
+
+lemma AC13_AC14: "[| n \<in> nat; 1\<le>n; AC13(n) |] ==> AC14"
+by (unfold AC13_def AC14_def, auto)
+
+(* ********************************************************************** *)
+(* AC14 ==> AC15                                                          *)
+(* ********************************************************************** *)
+
+lemma AC14_AC15: "AC14 ==> AC15"
+by (unfold AC13_def AC14_def AC15_def, fast)
+
+(* ********************************************************************** *)
+(* The redundant proofs; however cited by Rubin & Rubin                   *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* AC13(1) ==> AC1                                                        *)
+(* ********************************************************************** *)
+
+lemma lemma_aux: "[| A\<noteq>0; A \<lesssim> 1 |] ==> \<exists>a. A={a}"
+by (fast elim!: not_emptyE lepoll_1_is_sing)
+
+lemma AC13_AC1_lemma:
+     "\<forall>B \<in> A. f(B)\<noteq>0 & f(B)<=B & f(B) \<lesssim> 1   
+      ==> (\<lambda>x \<in> A. THE y. f(x)={y}) \<in> (\<Pi>X \<in> A. X)"
+apply (rule lam_type)
+apply (drule bspec, assumption)
+apply (elim conjE)
+apply (erule lemma_aux [THEN exE], assumption)
+apply (simp add: the_element)
+done
+
+lemma AC13_AC1: "AC13(1) ==> AC1"
+apply (unfold AC13_def AC1_def)
+apply (fast elim!: AC13_AC1_lemma)
+done
+
+(* ********************************************************************** *)
+(* AC11 ==> AC14                                                          *)
+(* ********************************************************************** *)
+
+lemma AC11_AC14: "AC11 ==> AC14"
+apply (unfold AC11_def AC14_def)
+apply (fast intro!: AC10_AC13)
+done
+
+end
+
--- a/src/ZF/AC/AC16_WO4.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,611 +0,0 @@
-(*  Title:      ZF/AC/AC16_WO4.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-  The proof of AC16(n, k) ==> WO4(n-k)
-*)
-
-(* ********************************************************************** *)
-(* The case of finite set                                                 *)
-(* ********************************************************************** *)
-
-Goalw [Finite_def] "[| Finite(A); 0<m; m \\<in> nat |] ==>  \
-\       \\<exists>a f. Ord(a) & domain(f) = a &  \
-\               (\\<Union>b<a. f`b) = A & (\\<forall>b<a. f`b lepoll m)";
-by (etac bexE 1);
-by (dresolve_tac [eqpoll_sym RS (eqpoll_def RS def_imp_iff RS iffD1)] 1);
-by (etac exE 1);
-by (res_inst_tac [("x","n")] exI 1);
-by (res_inst_tac [("x","\\<lambda>i \\<in> n. {f`i}")] exI 1);
-by (Asm_full_simp_tac 1);
-by (rewrite_goals_tac [bij_def, surj_def]);
-by (fast_tac (claset() addSIs [ltI, nat_into_Ord, lam_funtype RS domain_of_fun,
-        equalityI, singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
-        nat_1_lepoll_iff RS iffD2]
-        addSEs [apply_type, ltE]) 1);
-val lemma1 = result();
-
-(* ********************************************************************** *)
-(* The case of infinite set                                               *)
-(* ********************************************************************** *)
-
-(* well_ord(x,r) ==> well_ord({{y,z}. y \\<in> x}, Something(x,z))  **)
-bind_thm ("well_ord_paired", (paired_bij RS bij_is_inj RS well_ord_rvimage));
-
-Goal "[| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
-by (fast_tac (claset() addEs [notE, lepoll_trans]) 1);
-qed "lepoll_trans1";
-
-(* ********************************************************************** *)
-(* There exists a well ordered set y such that ...                        *)
-(* ********************************************************************** *)
-
-val lepoll_paired = paired_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll;
-
-Goal "\\<exists>y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)";
-by (res_inst_tac [("x","{{a,x}. a \\<in> nat Un Hartog(z)}")] exI 1);
-by (resolve_tac [transfer thy Ord_nat RS well_ord_Memrel RS
-		 (Ord_Hartog RS
-		  well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1);
-by (fast_tac 
-    (claset() addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired,
-		      HartogI RSN (2, lepoll_trans1),
-		 subset_imp_lepoll RS (lepoll_paired RSN (2, lepoll_trans))]
-              addSEs [nat_not_Finite RS notE] addEs [mem_asym]
-	      addSDs [Un_upper1 RS subset_imp_lepoll RS lepoll_Finite,
-		      lepoll_paired RS lepoll_Finite]) 1);
-val lemma2 = result();
-
-Goal "~Finite(B) ==> ~Finite(A Un B)";
-by (blast_tac (claset() addIs [subset_Finite]) 1);
-qed "infinite_Un";
-
-(* ********************************************************************** *)
-(* There is a v \\<in> s(u) such that k lepoll x Int y (in our case succ(k))    *)
-(* The idea of the proof is the following \\<in>                               *)
-(* Suppose not, i.e. every element of s(u) has exactly k-1 elements of y   *)
-(* Thence y is less than or equipollent to {v \\<in> Pow(x). v eqpoll n#-k}      *)
-(*   We have obtained this result in two steps \\<in>                          *)
-(*   1. y is less than or equipollent to {v \\<in> s(u). a \\<subseteq> v}                  *)
-(*      where a is certain k-2 element subset of y                        *)
-(*   2. {v \\<in> s(u). a \\<subseteq> v} is less than or equipollent                       *)
-(*      to {v \\<in> Pow(x). v eqpoll n-k}                                       *)
-(* ********************************************************************** *)
-
-(*Proof simplified by LCP*)
-Goal "[| ~(\\<exists>x \\<in> A. f`x=y); f \\<in> inj(A, B); y \\<in> B |]  \
-\     ==> (\\<lambda>a \\<in> succ(A). if(a=A, y, f`a)) \\<in> inj(succ(A), B)";
-by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1);
-by (force_tac (claset(), simpset() addsimps [inj_is_fun RS apply_type]) 1);
-(*this preliminary simplification prevents looping somehow*)
-by (Asm_simp_tac 1);  
-by (force_tac (claset(), simpset() addsimps []) 1);  
-qed "succ_not_lepoll_lemma";
-
-Goalw [lepoll_def, eqpoll_def, bij_def, surj_def]
-        "[| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B";
-by (fast_tac (claset() addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1);
-qed "succ_not_lepoll_imp_eqpoll";
-
-
-(* ********************************************************************** *)
-(* There is a k-2 element subset of y                                     *)
-(* ********************************************************************** *)
-
-val ordertype_eqpoll =
-        ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2));
-
-Goal "[| a \\<subseteq> y; b \\<in> y-a; u \\<in> x |] ==> cons(b, cons(u, a)) \\<in> Pow(x Un y)";
-by (Fast_tac 1);
-qed "cons_cons_subset";
-
-Goal "[| a eqpoll k; a \\<subseteq> y; b \\<in> y-a; u \\<in> x; x Int y = 0 |]   \
-\     ==> cons(b, cons(u, a)) eqpoll succ(succ(k))";
-by (fast_tac (claset() addSIs [cons_eqpoll_succ]) 1);
-qed "cons_cons_eqpoll";
-
-Goal "[| succ(k) eqpoll A; k eqpoll B; B \\<subseteq> A; a \\<in> A-B; k \\<in> nat |]   \
-\     ==> A = cons(a, B)";
-by (rtac equalityI 1);
-by (Fast_tac 2);
-by (resolve_tac [Diff_eq_0_iff RS iffD1] 1);
-by (rtac equals0I 1);
-by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1);
-by (dresolve_tac [eqpoll_sym RS cons_eqpoll_succ] 1);
-by (Fast_tac 1);
-by (dtac cons_eqpoll_succ 1);
-by (Fast_tac 1);
-by (fast_tac 
-    (claset() 
-        addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS
-        (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1);
-qed "set_eq_cons";
-
-Goal "[| cons(x,a) = cons(y,a); x\\<notin> a |] ==> x = y ";
-by (fast_tac (claset() addSEs [equalityE]) 1);
-qed "cons_eqE";
-
-Goal "A = B ==> A Int C = B Int C";
-by (Asm_simp_tac 1);
-qed "eq_imp_Int_eq";
-
-(* ********************************************************************** *)
-(* some arithmetic                                                        *)
-(* ********************************************************************** *)
-
-Goal "[| k \\<in> nat; m \\<in> nat |] ==>  \
-\       \\<forall>A B. A eqpoll k #+ m & k lepoll B & B \\<subseteq> A --> A-B lepoll m";
-by (induct_tac "k" 1);
-by (asm_simp_tac (simpset() addsimps [add_0]) 1);
-by (fast_tac (claset() addIs [eqpoll_imp_lepoll RS
-        (Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1);
-by (REPEAT (resolve_tac [allI,impI] 1));
-by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (Fast_tac 1);
-by (eres_inst_tac [("x","A - {xa}")] allE 1);
-by (eres_inst_tac [("x","B - {xa}")] allE 1);
-by (etac impE 1);
-by (asm_full_simp_tac (simpset() addsimps [add_succ]) 1);
-by (fast_tac (claset() addSIs [Diff_sing_eqpoll, lepoll_Diff_sing]) 1);
-by (res_inst_tac [("P","%z. z lepoll m")] subst 1 THEN (assume_tac 2));
-by (Fast_tac 1);
-qed "eqpoll_sum_imp_Diff_lepoll_lemma";
-
-Goal "[| A eqpoll succ(k #+ m); B \\<subseteq> A; succ(k) lepoll B;  k \\<in> nat; m \\<in> nat |]  \
-\     ==> A-B lepoll m";
-by (dresolve_tac [add_succ RS ssubst] 1);
-by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_lepoll_lemma] 1
-        THEN (REPEAT (assume_tac 1)));
-by (Fast_tac 1);
-qed "eqpoll_sum_imp_Diff_lepoll";
-
-(* ********************************************************************** *)
-(* similar properties for eqpoll                                          *)
-(* ********************************************************************** *)
-
-Goal "[| k \\<in> nat; m \\<in> nat |] ==>  \
-\       \\<forall>A B. A eqpoll k #+ m & k eqpoll B & B \\<subseteq> A --> A-B eqpoll m";
-by (induct_tac "k" 1);
-by (fast_tac (claset() addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0]
-        addss (simpset() addsimps [add_0])) 1);
-by (REPEAT (resolve_tac [allI,impI] 1));
-by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1);
-by (eres_inst_tac [("x","A - {xa}")] allE 1);
-by (eres_inst_tac [("x","B - {xa}")] allE 1);
-by (etac impE 1);
-by (fast_tac (claset() addSIs [Diff_sing_eqpoll,
-        eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym]
-        addss (simpset() addsimps [add_succ])) 1);
-by (res_inst_tac [("P","%z. z eqpoll m")] subst 1 THEN (assume_tac 2));
-by (Fast_tac 1);
-qed "eqpoll_sum_imp_Diff_eqpoll_lemma";
-
-Goal "[| A eqpoll succ(k #+ m); B \\<subseteq> A; succ(k) eqpoll B; k \\<in> nat; m \\<in> nat |]  \
-\     ==> A-B eqpoll m";
-by (dresolve_tac [add_succ RS ssubst] 1);
-by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_eqpoll_lemma] 1
-        THEN (REPEAT (assume_tac 1)));
-by (Fast_tac 1);
-qed "eqpoll_sum_imp_Diff_eqpoll";
-
-
-(* ********************************************************************** *)
-(* LL can be well ordered                                                 *)
-(* ********************************************************************** *)
-
-Goal "{x \\<in> Pow(X). x lepoll 0} = {0}";
-by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [lepoll_refl]) 1);
-qed "subsets_lepoll_0_eq_unit";
-
-Goal "n \\<in> nat ==> {z \\<in> Pow(y). z lepoll succ(n)} =  \
-\               {z \\<in> Pow(y). z lepoll n} Un {z \\<in> Pow(y). z eqpoll succ(n)}";
-by (fast_tac (claset() addIs [le_refl, leI, le_imp_lepoll]
-                addSDs [lepoll_succ_disj]
-                addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1);
-qed "subsets_lepoll_succ";
-
-Goal "n \\<in> nat ==> {z \\<in> Pow(y). z lepoll n} Int {z \\<in> Pow(y). z eqpoll succ(n)} = 0";
-by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_imp_lepoll 
-                RS lepoll_trans RS succ_lepoll_natE]
-                addSIs [equals0I]) 1);
-qed "Int_empty";
-
-
-Open_locale "AC16"; 
-
-val all_ex = thm "all_ex";
-val disjoint = thm "disjoint";
-val includes = thm "includes";
-val WO_R = thm "WO_R";
-val k_def = thm "k_def";
-val lnat = thm "lnat";
-val mnat = thm "mnat";
-val mpos = thm "mpos";
-val Infinite = thm "Infinite";
-val noLepoll = thm "noLepoll";
-
-val LL_def = thm "LL_def";
-val MM_def = thm "MM_def";
-val GG_def = thm "GG_def";
-val s_def = thm "s_def";
-
-Addsimps [disjoint, WO_R, lnat, mnat, mpos, Infinite];
-AddSIs [disjoint, WO_R, lnat, mnat, mpos];
-
-Goalw [k_def] "k \\<in> nat";
-by Auto_tac;
-qed "knat";
-Addsimps [knat];  AddSIs [knat];
-
-AddSIs [Infinite];   (*if notI is removed!*)
-AddSEs [Infinite RS notE];
-
-AddEs [[disjoint, IntI] MRS (equals0D RS notE)];
-
-(*use k = succ(l) *)
-val includes_l = simplify (FOL_ss addsimps [k_def]) includes;
-val all_ex_l = simplify (FOL_ss addsimps [k_def]) all_ex;
-
-(* ********************************************************************** *)
-(*   1. y is less than or equipollent to {v \\<in> s(u). a \\<subseteq> v}                  *)
-(*      where a is certain k-2 element subset of y                        *)
-(* ********************************************************************** *)
-
-Goal "[| l eqpoll a; a \\<subseteq> y |] ==> y - a eqpoll y";
-by (cut_facts_tac [WO_R, Infinite, lnat] 1);
-by (fast_tac (empty_cs addIs [lesspoll_trans1]
-        addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans,
-                Card_cardinal RS Card_is_Ord RS nat_le_infinite_Ord
-                RS le_imp_lepoll]
-        addSEs [well_ord_cardinal_eqpoll,
-                well_ord_cardinal_eqpoll RS eqpoll_sym,
-                eqpoll_sym RS eqpoll_imp_lepoll,
-                n_lesspoll_nat RS lesspoll_trans2,
-                well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
-                RS lepoll_infinite]) 1);
-qed "Diff_Finite_eqpoll";
-
-Goalw [s_def] "s(u) \\<subseteq> t_n";
-by (Fast_tac 1);
-qed "s_subset";
-
-Goalw [s_def, succ_def, k_def]
-      "[| w \\<in> t_n; cons(b,cons(u,a)) \\<subseteq> w; a \\<subseteq> y; b \\<in> y-a; l eqpoll a  \
-\      |] ==> w \\<in> s(u)";
-by (fast_tac (claset() addDs [eqpoll_imp_lepoll RS cons_lepoll_cong]
-                addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
-qed "sI";
-
-Goalw [s_def] "v \\<in> s(u) ==> u \\<in> v";
-by (Fast_tac 1);
-qed "in_s_imp_u_in";
-
-
-Goal "[| l eqpoll a;  a \\<subseteq> y;  b \\<in> y - a;  u \\<in> x |]  \
-\     ==> \\<exists>! c. c \\<in> s(u) & a \\<subseteq> c & b \\<in> c";
-by (rtac (all_ex_l RS ballE) 1);
-by (blast_tac (claset() delrules [PowI]
-		        addSIs [cons_cons_subset,
-				eqpoll_sym RS cons_cons_eqpoll]) 2);
-by (etac ex1E 1);
-by (res_inst_tac [("a","w")] ex1I 1);
-by (blast_tac (claset() addIs [sI]) 1);
-by (etac allE 1);
-by (etac impE 1);
-by (assume_tac 2);
-by (fast_tac (claset() addSEs [s_subset RS subsetD, in_s_imp_u_in]) 1);
-qed "ex1_superset_a";
-
-Goal "[| \\<forall>v \\<in> s(u). succ(l) eqpoll v Int y;  \
-\        l eqpoll a;  a \\<subseteq> y;  b \\<in> y - a;  u \\<in> x |]   \
-\     ==> (THE c. c \\<in> s(u) & a \\<subseteq> c & b \\<in> c)  \
-\              Int y = cons(b, a)";
-by (forward_tac [ex1_superset_a RS theI] 1 THEN REPEAT (assume_tac 1));
-by (rtac set_eq_cons 1);
-by (REPEAT (Fast_tac 1));
-qed "the_eq_cons";
-
-Goal "[| \\<forall>v \\<in> s(u). succ(l) eqpoll v Int y;  \
-\        l eqpoll a;  a \\<subseteq> y;  u \\<in> x |]  \
-\     ==> y lepoll {v \\<in> s(u). a \\<subseteq> v}";
-by (resolve_tac [Diff_Finite_eqpoll RS eqpoll_sym RS 
-		 eqpoll_imp_lepoll RS lepoll_trans] 1
-    THEN REPEAT (Fast_tac 1));
-by (res_inst_tac 
-     [("f3", "\\<lambda>b \\<in> y-a. THE c. c \\<in> s(u) & a \\<subseteq> c & b \\<in> c")]
-     (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1);
-by (simp_tac (simpset() addsimps [inj_def]) 1);
-by (rtac conjI 1);
-by (rtac lam_type 1);
-by (forward_tac [ex1_superset_a RS theI] 1 THEN REPEAT (Fast_tac 1));
-by (Asm_simp_tac 1);
-by (Clarify_tac 1);
-by (rtac cons_eqE 1);
-by (Fast_tac 2);
-by (dres_inst_tac [("A","THE c. ?P(c)"), ("C","y")] eq_imp_Int_eq 1);
-by (asm_full_simp_tac (simpset() addsimps [the_eq_cons]) 1);
-qed "y_lepoll_subset_s";
-
-
-(* ********************************************************************** *)
-(* back to the second part                                                *)
-(* ********************************************************************** *)
-
-Goal "w \\<subseteq> x Un y ==> w Int (x - {u}) = w - cons(u, w Int y)";
-by (Fast_tac 1);
-qed "w_Int_eq_w_Diff";
-
-Goal "[| w \\<in> {v \\<in> s(u). a \\<subseteq> v};  \
-\        l eqpoll a;  u \\<in> x;  \
-\        \\<forall>v \\<in> s(u). succ(l) eqpoll v Int y  \
-\     |] ==> w Int (x - {u}) eqpoll m";
-by (etac CollectE 1);
-by (stac w_Int_eq_w_Diff 1);
-by (fast_tac (claset() addSDs [s_subset RS subsetD,
-			       includes_l RS subsetD]) 1);
-by (fast_tac (claset() addSDs [bspec]
-        addDs [s_subset RS subsetD, includes_l RS subsetD]
-        addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_imp_u_in]
-        addSIs [eqpoll_sum_imp_Diff_eqpoll]) 1);
-qed "w_Int_eqpoll_m";
-
-(* ********************************************************************** *)
-(*   2. {v \\<in> s(u). a \\<subseteq> v} is less than or equipollent                       *)
-(*      to {v \\<in> Pow(x). v eqpoll n-k}                                       *)
-(* ********************************************************************** *)
-
-Goal "x eqpoll m ==> x \\<noteq> 0";
-by (cut_facts_tac [mpos] 1);
-by (fast_tac (claset() addSEs [zero_lt_natE]
-		       addSDs [eqpoll_succ_imp_not_empty]) 1);
-qed "eqpoll_m_not_empty";
-
-Goal "[| z \\<in> xa Int (x - {u}); l eqpoll a; a \\<subseteq> y; u \\<in> x |]  \
-\     ==> \\<exists>! w. w \\<in> t_n & cons(z, cons(u, a)) \\<subseteq> w";
-by (rtac (all_ex RS bspec) 1);
-by (rewtac k_def);
-by (fast_tac (claset() addSIs [cons_eqpoll_succ] addEs [eqpoll_sym]) 1);
-qed "cons_cons_in";
-
-Goal "[| \\<forall>v \\<in> s(u). succ(l) eqpoll v Int y;  \
-\        a \\<subseteq> y; l eqpoll a; u \\<in> x |]  \
-\     ==> {v \\<in> s(u). a \\<subseteq> v} lepoll {v \\<in> Pow(x). v eqpoll m}";
-by (res_inst_tac [("f3","\\<lambda>w \\<in> {v \\<in> s(u). a \\<subseteq> v}. w Int (x - {u})")] 
-        (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1);
-by (simp_tac (simpset() addsimps [inj_def]) 1);
-by (rtac conjI 1);
-by (rtac lam_type 1);
-by (rtac CollectI 1);
-by (Fast_tac 1);
-by (rtac w_Int_eqpoll_m 1 THEN REPEAT (assume_tac 1));
-by (REPEAT (resolve_tac [ballI, impI] 1));
-(** LEVEL 8 **)
-by (resolve_tac [w_Int_eqpoll_m RS eqpoll_m_not_empty RS not_emptyE] 1);
-by (EVERY (map Blast_tac [4,3,2,1]));
-by (dresolve_tac [equalityD1 RS subsetD] 1 THEN (assume_tac 1));
-by (ftac cons_cons_in 1 THEN REPEAT (assume_tac 1));
-by (etac ex1_two_eq 1);
-by (REPEAT (blast_tac
-	    (claset() addDs [s_subset RS subsetD, in_s_imp_u_in]) 1));
-qed "subset_s_lepoll_w";
-
-
-(* ********************************************************************** *)
-(* well_ord_subsets_lepoll_n                                              *)
-(* ********************************************************************** *)
-
-Goal "n \\<in> nat ==> \\<exists>S. well_ord({z \\<in> Pow(y) . z eqpoll succ(n)}, S)";
-by (resolve_tac [WO_R RS well_ord_infinite_subsets_eqpoll_X
-		 RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1);
-by (REPEAT (fast_tac (claset() addIs [bij_is_inj RS well_ord_rvimage]) 1));
-qed "well_ord_subsets_eqpoll_n";
-
-Goal "n \\<in> nat ==> \\<exists>R. well_ord({z \\<in> Pow(y). z lepoll n}, R)";
-by (induct_tac "n" 1);
-by (force_tac (claset() addSIs [well_ord_unit],
-	       simpset() addsimps [subsets_lepoll_0_eq_unit]) 1);
-by (etac exE 1);
-by (resolve_tac [well_ord_subsets_eqpoll_n RS exE] 1 THEN assume_tac 1);
-by (asm_simp_tac (simpset() addsimps [subsets_lepoll_succ]) 1);
-by (dtac well_ord_radd 1 THEN (assume_tac 1));
-by (eresolve_tac [Int_empty RS disj_Un_eqpoll_sum RS 
-                (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1);
-by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage]) 1);
-qed "well_ord_subsets_lepoll_n";
-
-
-Goalw [LL_def, MM_def] "LL \\<subseteq> {z \\<in> Pow(y). z lepoll succ(k #+ m)}";
-by (cut_facts_tac [includes] 1);
-by (fast_tac (claset() addIs [subset_imp_lepoll 
-			      RS (eqpoll_imp_lepoll
-				  RSN (2, lepoll_trans))]) 1);
-qed "LL_subset";
-
-Goal "\\<exists>S. well_ord(LL,S)";
-by (rtac (well_ord_subsets_lepoll_n RS exE) 1);
-by (blast_tac (claset() addIs [LL_subset RSN (2, well_ord_subset)]) 2);
-by Auto_tac;
-qed "well_ord_LL";
-
-(* ********************************************************************** *)
-(* every element of LL is a contained in exactly one element of MM        *)
-(* ********************************************************************** *)
-
-Goalw [MM_def, LL_def] "v \\<in> LL ==> \\<exists>! w. w \\<in> MM & v \\<subseteq> w";
-by Safe_tac;
-by (Fast_tac 1);
-by (resolve_tac [lepoll_imp_eqpoll_subset RS exE] 1 THEN (assume_tac 1));
-by (res_inst_tac [("x","x")] (all_ex RS ballE) 1);
-by (fast_tac (claset() addSEs [eqpoll_sym]) 2);
-by (Blast_tac 1);
-qed "unique_superset_in_MM";
-
-val unique_superset1 = unique_superset_in_MM RS theI RS conjunct1;
-val unique_superset2 = unique_superset_in_MM RS the_equality2;
-
-
-(* ********************************************************************** *)
-(* The function GG satisfies the conditions of WO4                        *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* The union of appropriate values is the whole x                         *)
-(* ********************************************************************** *)
-
-Goalw [LL_def] "w \\<in> MM ==> w Int y \\<in> LL";
-by (Fast_tac 1);
-qed "Int_in_LL";
-
-Goalw [LL_def] 
-     "v \\<in> LL ==> v = (THE x. x \\<in> MM & v \\<subseteq> x) Int y";
-by (Clarify_tac 1);
-by (stac unique_superset2 1);
-by (auto_tac (claset(), simpset() addsimps [Int_in_LL]));
-qed "in_LL_eq_Int";
-
-Goal "v \\<in> LL ==> (THE x. x \\<in> MM & v \\<subseteq> x) \\<subseteq> x Un y";
-by (dtac unique_superset1 1);
-by (rewtac MM_def);
-by (fast_tac (claset() addSDs [unique_superset1, includes RS subsetD]) 1);
-qed "the_in_MM_subset";
-
-Goalw [GG_def] "v \\<in> LL ==> GG ` v \\<subseteq> x";
-by (ftac the_in_MM_subset 1);
-by (ftac in_LL_eq_Int 1); 
-by (force_tac (claset() addEs [equalityE], simpset()) 1);
-qed "GG_subset";
-
-
-Goal "n \\<in> nat ==> \\<exists>z. z \\<subseteq> y & n eqpoll z";
-by (etac nat_lepoll_imp_ex_eqpoll_n 1);
-by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
-        RSN (2, lepoll_trans)] 1);
-by (rtac WO_R 2);
-by (fast_tac 
-    (claset() delrules [notI]
-              addSIs [nat_le_infinite_Ord RS le_imp_lepoll]
-	      addIs [Ord_ordertype, 
-		     ordertype_eqpoll RS eqpoll_imp_lepoll
-		     RS lepoll_infinite]) 1);
-qed "ex_subset_eqpoll_n";
-
-
-Goal "u \\<in> x ==> \\<exists>v \\<in> s(u). succ(k) lepoll v Int y";
-by (rtac ccontr 1);
-by (subgoal_tac "\\<forall>v \\<in> s(u). k eqpoll v Int y" 1);
-by (full_simp_tac (simpset() addsimps [s_def]) 2);
-by (blast_tac (claset() addIs [succ_not_lepoll_imp_eqpoll]) 2);
-by (rewtac k_def);
-by (cut_facts_tac [all_ex, includes, lnat] 1);
-by (rtac (ex_subset_eqpoll_n RS exE) 1 THEN assume_tac 1);
-by (rtac (noLepoll RS notE) 1);
-by (blast_tac (claset() addIs
-	   [[y_lepoll_subset_s, subset_s_lepoll_w] MRS lepoll_trans]) 1);
-qed "exists_proper_in_s";
-
-Goal "u \\<in> x ==> \\<exists>w \\<in> MM. u \\<in> w";
-by (eresolve_tac [exists_proper_in_s RS bexE] 1);
-by (rewrite_goals_tac [MM_def, s_def]);
-by (Fast_tac 1);
-qed "exists_in_MM";
-
-Goal "u \\<in> x ==> \\<exists>w \\<in> LL. u \\<in> GG`w";
-by (rtac (exists_in_MM RS bexE) 1);
-by (assume_tac 1);
-by (rtac bexI 1);
-by (etac Int_in_LL 2);
-by (rewtac GG_def);
-by (asm_simp_tac (simpset() addsimps [Int_in_LL]) 1);
-by (stac unique_superset2 1);
-by (REPEAT (fast_tac (claset() addSEs [Int_in_LL]) 1));
-qed "exists_in_LL";
-
-
-Goal "well_ord(LL,S) ==>      \
-\     (\\<Union>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b)) = x";
-by (rtac equalityI 1);
-by (rtac subsetI 1);
-by (etac OUN_E 1);
-by (resolve_tac [GG_subset RS subsetD] 1);
-by (assume_tac 2);
-by (blast_tac (claset() addIs [ordermap_bij RS bij_converse_bij RS
-			       bij_is_fun RS apply_type, ltD]) 1);
-by (rtac subsetI 1);
-by (eresolve_tac [exists_in_LL RS bexE] 1);
-by (force_tac (claset() addIs [Ord_ordertype RSN (2, ltI),
-			       ordermap_type RS apply_type],
-        simpset() addsimps [ordermap_bij RS bij_is_inj RS left_inverse]) 1);
-qed "OUN_eq_x";
-
-(* ********************************************************************** *)
-(* Every element of the family is less than or equipollent to n-k (m)     *)
-(* ********************************************************************** *)
-
-Goalw [MM_def] "w \\<in> MM ==> w eqpoll succ(k #+ m)";
-by (fast_tac (claset() addDs [includes RS subsetD]) 1);
-qed "in_MM_eqpoll_n";
-
-Goalw [LL_def, MM_def] "w \\<in> LL ==> succ(k) lepoll w";
-by (Fast_tac 1);
-qed "in_LL_eqpoll_n";
-
-val in_LL = in_LL_eq_Int RS equalityD1 RS (Int_lower1 RSN (2, subset_trans));
-
-Goalw [GG_def] 
-      "well_ord(LL,S) ==>      \
-\      \\<forall>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b) lepoll m";
-by (rtac oallI 1);
-by (asm_simp_tac 
-    (simpset() addsimps [ltD,
-			 ordermap_bij RS bij_converse_bij RS
-			 bij_is_fun RS apply_type]) 1);
-by (cut_facts_tac [includes] 1);
-by (rtac eqpoll_sum_imp_Diff_lepoll 1);
-by (REPEAT
-    (fast_tac (claset() delrules [subsetI]
-		        addSDs [ltD]
-			addSIs [eqpoll_sum_imp_Diff_lepoll, in_LL_eqpoll_n]
-			addIs [unique_superset1 RS in_MM_eqpoll_n, in_LL,
-			       ordermap_bij RS bij_converse_bij RS 
-			       bij_is_fun RS apply_type]) 1 ));
-qed "all_in_lepoll_m";
-
-
-Goal "\\<exists>a f. Ord(a) & domain(f) = a &  \
-\             (\\<Union>b<a. f ` b) = x & (\\<forall>b<a. f ` b lepoll m)";
-by (resolve_tac [well_ord_LL RS exE] 1 THEN REPEAT (assume_tac 1));
-by (rename_tac "S" 1);
-by (res_inst_tac [("x","ordertype(LL,S)")] exI 1);
-by (res_inst_tac [("x",
-        "\\<lambda>b \\<in> ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b)")] 
-    exI 1);
-by (Simp_tac 1);
-by (REPEAT (ares_tac [conjI, lam_funtype RS domain_of_fun,
-		      Ord_ordertype, 
-		      all_in_lepoll_m, OUN_eq_x] 1));
-qed "conclusion";
-
-Close_locale "AC16";
-
-
-
-(* ********************************************************************** *)
-(* The main theorem AC16(n, k) ==> WO4(n-k)                               *)
-(* ********************************************************************** *)
-
-Goalw [AC16_def,WO4_def]
-        "[| AC16(k #+ m, k); 0 < k; 0 < m; k \\<in> nat; m \\<in> nat |] ==> WO4(m)";
-by (rtac allI 1);
-by (case_tac "Finite(A)" 1);
-by (rtac lemma1 1 THEN REPEAT (assume_tac 1));
-by (cut_facts_tac [lemma2] 1);
-by (REPEAT (eresolve_tac [exE, conjE] 1));
-by (eres_inst_tac [("x","A Un y")] allE 1);
-by (ftac infinite_Un 1 THEN (mp_tac 1));
-by (etac zero_lt_natE 1); 
-by (assume_tac 1);
-by (Clarify_tac 1);
-by (DEPTH_SOLVE (ares_tac [export conclusion] 1));
-qed "AC16_WO4";
-
--- a/src/ZF/AC/AC16_WO4.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC16_WO4.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -2,40 +2,575 @@
     ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-Tidied using locales by LCP
+The proof of AC16(n, k) ==> WO4(n-k)
+
+Tidied (using locales) by LCP
 *)
 
-AC16_WO4 = OrderType + AC16_lemmas + Cardinal_aux +
+theory AC16_WO4 = AC16_lemmas:
+
+(* ********************************************************************** *)
+(* The case of finite set                                                 *)
+(* ********************************************************************** *)
+
+lemma lemma1:
+     "[| Finite(A); 0<m; m \<in> nat |] 
+      ==> \<exists>a f. Ord(a) & domain(f) = a &   
+                (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
+apply (unfold Finite_def)
+apply (erule bexE)
+apply (drule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]])
+apply (erule exE)
+apply (rule_tac x = "n" in exI)
+apply (rule_tac x = "\<lambda>i \<in> n. {f`i}" in exI, simp)
+apply (unfold bij_def surj_def)
+apply (fast intro!: ltI nat_into_Ord lam_funtype [THEN domain_of_fun] 
+               singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] 
+                    nat_1_lepoll_iff [THEN iffD2]
+          elim!: apply_type ltE)
+done
+
+(* ********************************************************************** *)
+(* The case of infinite set                                               *)
+(* ********************************************************************** *)
+
+(* well_ord(x,r) ==> well_ord({{y,z}. y \<in> x}, Something(x,z))  **)
+lemmas well_ord_paired = paired_bij [THEN bij_is_inj, THEN well_ord_rvimage]
+
+lemma lepoll_trans1: "[| A \<lesssim> B; ~ A \<lesssim> C |] ==> ~ B \<lesssim> C"
+by (blast intro: lepoll_trans)
+
+(* ********************************************************************** *)
+(* There exists a well ordered set y such that ...                        *)
+(* ********************************************************************** *)
+
+lemmas lepoll_paired = paired_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll];
+
+lemma lemma2: "\<exists>y R. well_ord(y,R) & x Int y = 0 & ~y \<lesssim> z & ~Finite(y)"
+apply (rule_tac x = "{{a,x}. a \<in> nat Un Hartog (z) }" in exI)
+apply (rule well_ord_Un [OF Ord_nat [THEN well_ord_Memrel] 
+                         Ord_Hartog [THEN well_ord_Memrel], THEN exE])
+apply (blast intro!: Ord_Hartog well_ord_Memrel well_ord_paired
+                      lepoll_trans1 [OF _ not_Hartog_lepoll_self]
+                      lepoll_trans [OF subset_imp_lepoll lepoll_paired]
+       elim!: nat_not_Finite [THEN notE]
+       elim: mem_asym 
+       dest!: Un_upper1 [THEN subset_imp_lepoll, THEN lepoll_Finite]
+              lepoll_paired [THEN lepoll_Finite])
+done
+
+lemma infinite_Un: "~Finite(B) ==> ~Finite(A Un B)"
+by (blast intro: subset_Finite)
+
+(* ********************************************************************** *)
+(* There is a v \<in> s(u) such that k \<lesssim> x Int y (in our case succ(k))    *)
+(* The idea of the proof is the following \<in>                               *)
+(* Suppose not, i.e. every element of s(u) has exactly k-1 elements of y   *)
+(* Thence y is less than or equipollent to {v \<in> Pow(x). v \<approx> n#-k}      *)
+(*   We have obtained this result in two steps \<in>                          *)
+(*   1. y is less than or equipollent to {v \<in> s(u). a \<subseteq> v}                  *)
+(*      where a is certain k-2 element subset of y                        *)
+(*   2. {v \<in> s(u). a \<subseteq> v} is less than or equipollent                       *)
+(*      to {v \<in> Pow(x). v \<approx> n-k}                                       *)
+(* ********************************************************************** *)
+
+(*Proof simplified by LCP*)
+lemma succ_not_lepoll_lemma:
+     "[| ~(\<exists>x \<in> A. f`x=y); f \<in> inj(A, B); y \<in> B |]   
+      ==> (\<lambda>a \<in> succ(A). if(a=A, y, f`a)) \<in> inj(succ(A), B)"
+apply (rule_tac d = "%z. if (z=y, A, converse (f) `z) " in lam_injective)
+apply (force simp add: inj_is_fun [THEN apply_type])
+(*this preliminary simplification prevents looping somehow*)
+apply (simp (no_asm_simp))
+apply force
+done
+
+lemma succ_not_lepoll_imp_eqpoll: "[| ~A \<approx> B; A \<lesssim> B |] ==> succ(A) \<lesssim> B"
+apply (unfold lepoll_def eqpoll_def bij_def surj_def)
+apply (fast elim!: succ_not_lepoll_lemma inj_is_fun)
+done
+
+
+(* ********************************************************************** *)
+(* There is a k-2 element subset of y                                     *)
+(* ********************************************************************** *)
+
+lemmas ordertype_eqpoll =
+       ordermap_bij [THEN exI [THEN eqpoll_def [THEN def_imp_iff, THEN iffD2]]]
+
+lemma cons_cons_subset:
+     "[| a \<subseteq> y; b \<in> y-a; u \<in> x |] ==> cons(b, cons(u, a)) \<in> Pow(x Un y)"
+by fast
+
+lemma cons_cons_eqpoll:
+     "[| a \<approx> k; a \<subseteq> y; b \<in> y-a; u \<in> x; x Int y = 0 |]    
+      ==> cons(b, cons(u, a)) \<approx> succ(succ(k))"
+by (fast intro!: cons_eqpoll_succ)
+
+lemma set_eq_cons:
+     "[| succ(k) \<approx> A; k \<approx> B; B \<subseteq> A; a \<in> A-B; k \<in> nat |] ==> A = cons(a, B)"
+apply (rule equalityI)
+prefer 2 apply fast
+apply (rule Diff_eq_0_iff [THEN iffD1])
+apply (rule equals0I)
+apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll])
+apply (drule eqpoll_sym [THEN cons_eqpoll_succ], fast)
+apply (drule cons_eqpoll_succ, fast)
+apply (fast elim!: lepoll_trans [THEN lepoll_trans, THEN succ_lepoll_natE,
+         OF eqpoll_sym [THEN eqpoll_imp_lepoll] subset_imp_lepoll])
+done
+
+lemma cons_eqE: "[| cons(x,a) = cons(y,a); x \<notin> a |] ==> x = y "
+by (fast elim!: equalityE)
+
+lemma eq_imp_Int_eq: "A = B ==> A Int C = B Int C"
+by blast
+
+(* ********************************************************************** *)
+(* some arithmetic                                                        *)
+(* ********************************************************************** *)
+
+lemma eqpoll_sum_imp_Diff_lepoll_lemma [rule_format]:
+     "[| k \<in> nat; m \<in> nat |] 
+      ==> \<forall>A B. A \<approx> k #+ m & k \<lesssim> B & B \<subseteq> A --> A-B \<lesssim> m"
+apply (induct_tac "k")
+apply (simp add: add_0)
+apply (blast intro: eqpoll_imp_lepoll lepoll_trans
+                    Diff_subset [THEN subset_imp_lepoll])
+apply (intro allI impI)
+apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], fast)
+apply (erule_tac x = "A - {xa}" in allE)
+apply (erule_tac x = "B - {xa}" in allE)
+apply (erule impE)
+apply (simp add: add_succ)
+apply (fast intro!: Diff_sing_eqpoll lepoll_Diff_sing) 
+apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp); 
+apply blast 
+done
+
+lemma eqpoll_sum_imp_Diff_lepoll:
+     "[| A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<lesssim> B;  k \<in> nat; m \<in> nat |]   
+      ==> A-B \<lesssim> m"
+apply (simp only: add_succ [symmetric]) 
+apply (blast intro: eqpoll_sum_imp_Diff_lepoll_lemma) 
+done
+
+(* ********************************************************************** *)
+(* similar properties for \<approx>                                          *)
+(* ********************************************************************** *)
+
+lemma eqpoll_sum_imp_Diff_eqpoll_lemma [rule_format]:
+     "[| k \<in> nat; m \<in> nat |] 
+      ==> \<forall>A B. A \<approx> k #+ m & k \<approx> B & B \<subseteq> A --> A-B \<approx> m"
+apply (induct_tac "k")
+apply (force dest!: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_0_is_0])
+apply (intro allI impI)
+apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE])
+apply (fast elim!: eqpoll_imp_lepoll)
+apply (erule_tac x = "A - {xa}" in allE)
+apply (erule_tac x = "B - {xa}" in allE)
+apply (erule impE)
+apply (force intro: eqpoll_sym intro!: Diff_sing_eqpoll)
+apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp); 
+apply blast 
+done
+
+lemma eqpoll_sum_imp_Diff_eqpoll:
+     "[| A \<approx> succ(k #+ m); B \<subseteq> A; succ(k) \<approx> B; k \<in> nat; m \<in> nat |]   
+      ==> A-B \<approx> m"
+apply (simp only: add_succ [symmetric]) 
+apply (blast intro: eqpoll_sum_imp_Diff_eqpoll_lemma) 
+done
+
+
+(* ********************************************************************** *)
+(* LL can be well ordered                                                 *)
+(* ********************************************************************** *)
+
+lemma subsets_lepoll_0_eq_unit: "{x \<in> Pow(X). x \<lesssim> 0} = {0}"
+by (fast dest!: lepoll_0_is_0 intro!: lepoll_refl)
+
+lemma subsets_lepoll_succ:
+     "n \<in> nat ==> {z \<in> Pow(y). z \<lesssim> succ(n)} =   
+                  {z \<in> Pow(y). z \<lesssim> n} Un {z \<in> Pow(y). z \<approx> succ(n)}"
+by (blast intro: leI le_imp_lepoll nat_into_Ord 
+                    lepoll_trans eqpoll_imp_lepoll
+          dest!: lepoll_succ_disj)
+
+lemma Int_empty:
+     "n \<in> nat ==> {z \<in> Pow(y). z \<lesssim> n} Int {z \<in> Pow(y). z \<approx> succ(n)} = 0"
+by (blast intro: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] 
+                 succ_lepoll_natE)
+
 
 locale AC16 =
-  fixes 
-    x	:: i
-    y	:: i
-    k	:: i
-    l   :: i
-    m	:: i
-    t_n	:: i
-    R	:: i
-    MM  :: i
-    LL  :: i
-    GG  :: i
-    s   :: i=>i
-  assumes
-    all_ex    "\\<forall>z \\<in> {z \\<in> Pow(x Un y) . z eqpoll succ(k)}.
-	         \\<exists>! w. w \\<in> t_n & z \\<subseteq> w "
-    disjoint  "x Int y = 0"
-    includes  "t_n \\<subseteq> {v \\<in> Pow(x Un y). v eqpoll succ(k #+ m)}"
-    WO_R      "well_ord(y,R)"
-    lnat      "l \\<in> nat"
-    mnat      "m \\<in> nat"
-    mpos      "0<m"
-    Infinite  "~ Finite(y)"
-    noLepoll  "~ y lepoll {v \\<in> Pow(x). v eqpoll m}"
-  defines
-    k_def     "k   == succ(l)"
-    MM_def    "MM  == {v \\<in> t_n. succ(k) lepoll v Int y}"
-    LL_def    "LL  == {v Int y. v \\<in> MM}"
-    GG_def    "GG  == \\<lambda>v \\<in> LL. (THE w. w \\<in> MM & v \\<subseteq> w) - v"
-    s_def     "s(u) == {v \\<in> t_n. u \\<in> v & k lepoll v Int y}"
+  fixes x and y and k and l and m and t_n and R and MM and LL and GG and s 
+  defines k_def:     "k   == succ(l)"
+      and MM_def:    "MM  == {v \<in> t_n. succ(k) \<lesssim> v Int y}"
+      and LL_def:    "LL  == {v Int y. v \<in> MM}"
+      and GG_def:    "GG  == \<lambda>v \<in> LL. (THE w. w \<in> MM & v \<subseteq> w) - v"
+      and s_def:     "s(u) == {v \<in> t_n. u \<in> v & k \<lesssim> v Int y}"
+  assumes all_ex:    "\<forall>z \<in> {z \<in> Pow(x Un y) . z \<approx> succ(k)}.
+	               \<exists>! w. w \<in> t_n & z \<subseteq> w "
+    and disjoint[iff]:  "x Int y = 0"
+    and includes:  "t_n \<subseteq> {v \<in> Pow(x Un y). v \<approx> succ(k #+ m)}"
+    and WO_R[iff]:      "well_ord(y,R)"
+    and lnat[iff]:      "l \<in> nat"
+    and mnat[iff]:      "m \<in> nat"
+    and mpos[iff]:      "0<m"
+    and Infinite[iff]:  "~ Finite(y)"
+    and noLepoll:  "~ y \<lesssim> {v \<in> Pow(x). v \<approx> m}"
+
+lemma (in AC16) knat [iff]: "k \<in> nat"
+by (simp add: k_def)
+
+
+(* ********************************************************************** *)
+(*   1. y is less than or equipollent to {v \<in> s(u). a \<subseteq> v}                *)
+(*      where a is certain k-2 element subset of y                        *)
+(* ********************************************************************** *)
+
+lemma (in AC16) Diff_Finite_eqpoll: "[| l \<approx> a; a \<subseteq> y |] ==> y - a \<approx> y"
+apply (insert WO_R Infinite lnat)
+apply (rule eqpoll_trans) 
+apply (rule Diff_lesspoll_eqpoll_Card) 
+apply (erule well_ord_cardinal_eqpoll [THEN eqpoll_sym])
+apply (blast intro: lesspoll_trans1
+            intro!: Card_cardinal  
+                    Card_cardinal [THEN Card_is_Ord, THEN nat_le_infinite_Ord,
+                                   THEN le_imp_lepoll] 
+            dest: well_ord_cardinal_eqpoll 
+		   eqpoll_sym  eqpoll_imp_lepoll
+                   n_lesspoll_nat [THEN lesspoll_trans2]
+                   well_ord_cardinal_eqpoll [THEN eqpoll_sym, 
+                          THEN eqpoll_imp_lepoll, THEN lepoll_infinite])+
+done
+
+
+lemma (in AC16) s_subset: "s(u) \<subseteq> t_n"
+by (unfold s_def, blast)
+
+lemma (in AC16) sI: 
+      "[| w \<in> t_n; cons(b,cons(u,a)) \<subseteq> w; a \<subseteq> y; b \<in> y-a; l \<approx> a |] 
+       ==> w \<in> s(u)"
+apply (unfold s_def succ_def k_def)
+apply (blast intro!: eqpoll_imp_lepoll [THEN cons_lepoll_cong]
+             intro: subset_imp_lepoll lepoll_trans)
+done
+
+lemma (in AC16) in_s_imp_u_in: "v \<in> s(u) ==> u \<in> v"
+by (unfold s_def, blast)
+
+
+lemma (in AC16) ex1_superset_a:
+     "[| l \<approx> a;  a \<subseteq> y;  b \<in> y - a;  u \<in> x |]   
+      ==> \<exists>! c. c \<in> s(u) & a \<subseteq> c & b \<in> c"
+apply (rule all_ex [simplified k_def, THEN ballE])
+ apply (erule ex1E)
+ apply (rule_tac a = "w" in ex1I, blast intro: sI)
+ apply (blast dest: s_subset [THEN subsetD] in_s_imp_u_in)
+apply (blast del: PowI 
+             intro!: cons_cons_subset eqpoll_sym [THEN cons_cons_eqpoll])
+done
+
+lemma (in AC16) the_eq_cons:
+     "[| \<forall>v \<in> s(u). succ(l) \<approx> v Int y;   
+         l \<approx> a;  a \<subseteq> y;  b \<in> y - a;  u \<in> x |]    
+      ==> (THE c. c \<in> s(u) & a \<subseteq> c & b \<in> c) Int y = cons(b, a)"
+apply (frule ex1_superset_a [THEN theI], assumption+)
+apply (rule set_eq_cons)
+apply (fast+)
+done
+
+lemma (in AC16) y_lepoll_subset_s:
+     "[| \<forall>v \<in> s(u). succ(l) \<approx> v Int y;   
+         l \<approx> a;  a \<subseteq> y;  u \<in> x |]   
+      ==> y \<lesssim> {v \<in> s(u). a \<subseteq> v}"
+apply (rule Diff_Finite_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll, 
+                                THEN lepoll_trans],  fast+)
+apply (rule_tac  f3 = "\<lambda>b \<in> y-a. THE c. c \<in> s (u) & a \<subseteq> c & b \<in> c" 
+        in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]])
+apply (simp add: inj_def)
+apply (rule conjI)
+apply (rule lam_type)
+apply (frule ex1_superset_a [THEN theI], fast+, clarify)
+apply (rule cons_eqE [of _ a])
+apply (drule_tac A = "THE c. ?P (c) " and C = "y" in eq_imp_Int_eq)
+apply (simp_all add: the_eq_cons)
+done
+
+
+(* ********************************************************************** *)
+(* back to the second part                                                *)
+(* ********************************************************************** *)
+
+
+(*relies on the disjointness of x, y*)
+lemma (in AC16) x_imp_not_y [dest]: "a \<in> x ==> a \<notin> y"
+by (blast dest:  disjoint [THEN equalityD1, THEN subsetD, OF IntI])
+
+lemma (in AC16) w_Int_eq_w_Diff:
+     "w \<subseteq> x Un y ==> w Int (x - {u}) = w - cons(u, w Int y)" 
+by blast
+
+lemma (in AC16) w_Int_eqpoll_m:
+     "[| w \<in> {v \<in> s(u). a \<subseteq> v};   
+         l \<approx> a;  u \<in> x;   
+         \<forall>v \<in> s(u). succ(l) \<approx> v Int y |] 
+      ==> w Int (x - {u}) \<approx> m"
+apply (erule CollectE)
+apply (subst w_Int_eq_w_Diff)
+apply (fast dest!: s_subset [THEN subsetD] 
+                   includes [simplified k_def, THEN subsetD])
+apply (blast dest: s_subset [THEN subsetD] 
+                   includes [simplified k_def, THEN subsetD] 
+             dest: eqpoll_sym [THEN cons_eqpoll_succ, THEN eqpoll_sym] 
+                   in_s_imp_u_in
+            intro!: eqpoll_sum_imp_Diff_eqpoll)
+done
+
+
+(* ********************************************************************** *)
+(*   2. {v \<in> s(u). a \<subseteq> v} is less than or equipollent                       *)
+(*      to {v \<in> Pow(x). v \<approx> n-k}                                       *)
+(* ********************************************************************** *)
+
+lemma (in AC16) eqpoll_m_not_empty: "a \<approx> m ==> a \<noteq> 0"
+apply (insert mpos)
+apply (fast elim!: zero_lt_natE dest!: eqpoll_succ_imp_not_empty)
+done
+
+lemma (in AC16) cons_cons_in:
+     "[| z \<in> xa Int (x - {u}); l \<approx> a; a \<subseteq> y; u \<in> x |]   
+      ==> \<exists>! w. w \<in> t_n & cons(z, cons(u, a)) \<subseteq> w"
+apply (rule all_ex [THEN bspec])
+apply (unfold k_def)
+apply (fast intro!: cons_eqpoll_succ elim: eqpoll_sym)
+done
+
+lemma (in AC16) subset_s_lepoll_w:
+     "[| \<forall>v \<in> s(u). succ(l) \<approx> v Int y; a \<subseteq> y; l \<approx> a; u \<in> x |]   
+      ==> {v \<in> s(u). a \<subseteq> v} \<lesssim> {v \<in> Pow(x). v \<approx> m}"
+apply (rule_tac f3 = "\<lambda>w \<in> {v \<in> s (u) . a \<subseteq> v}. w Int (x - {u})" 
+       in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]])
+apply (simp add: inj_def)
+apply (intro conjI lam_type CollectI)
+  apply fast
+ apply (blast intro: w_Int_eqpoll_m) 
+apply (intro ballI impI)
+(** LEVEL 8 **)
+apply (rule w_Int_eqpoll_m [THEN eqpoll_m_not_empty, THEN not_emptyE])
+apply (blast, assumption+)
+apply (drule equalityD1 [THEN subsetD], (assumption))
+apply (frule cons_cons_in, assumption+)
+apply (blast dest: ex1_two_eq intro: s_subset [THEN subsetD] in_s_imp_u_in)+
+done
+
+
+(* ********************************************************************** *)
+(* well_ord_subsets_lepoll_n                                              *)
+(* ********************************************************************** *)
+
+lemma (in AC16) well_ord_subsets_eqpoll_n:
+     "n \<in> nat ==> \<exists>S. well_ord({z \<in> Pow(y) . z \<approx> succ(n)}, S)"
+apply (rule WO_R [THEN well_ord_infinite_subsets_eqpoll_X,
+                  THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE])
+apply (fast intro: bij_is_inj [THEN well_ord_rvimage])+
+done
+
+lemma (in AC16) well_ord_subsets_lepoll_n:
+     "n \<in> nat ==> \<exists>R. well_ord({z \<in> Pow(y). z \<lesssim> n}, R)"
+apply (induct_tac "n")
+apply (force intro!: well_ord_unit simp add: subsets_lepoll_0_eq_unit)
+apply (erule exE)
+apply (rule well_ord_subsets_eqpoll_n [THEN exE], assumption)
+apply (simp add: subsets_lepoll_succ)
+apply (drule well_ord_radd, (assumption))
+apply (erule Int_empty [THEN disj_Un_eqpoll_sum,
+                  THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE])
+apply (fast elim!: bij_is_inj [THEN well_ord_rvimage])
+done
+
+
+lemma (in AC16) LL_subset: "LL \<subseteq> {z \<in> Pow(y). z \<lesssim> succ(k #+ m)}"
+apply (unfold LL_def MM_def)
+apply (insert includes)
+apply (blast intro: subset_imp_lepoll eqpoll_imp_lepoll lepoll_trans)
+done
+
+lemma (in AC16) well_ord_LL: "\<exists>S. well_ord(LL,S)"
+apply (rule well_ord_subsets_lepoll_n [THEN exE, of "succ(k#+m)"])
+apply simp 
+apply (blast intro: well_ord_subset [OF _ LL_subset])
+done
+
+(* ********************************************************************** *)
+(* every element of LL is a contained in exactly one element of MM        *)
+(* ********************************************************************** *)
+
+lemma (in AC16) unique_superset_in_MM:
+     "v \<in> LL ==> \<exists>! w. w \<in> MM & v \<subseteq> w"
+apply (unfold MM_def LL_def, safe)
+apply fast
+apply (rule lepoll_imp_eqpoll_subset [THEN exE], (assumption))
+apply (rule_tac x = "x" in all_ex [THEN ballE]) 
+apply (blast intro: eqpoll_sym)+
+done
+
+
+(* ********************************************************************** *)
+(* The function GG satisfies the conditions of WO4                        *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* The union of appropriate values is the whole x                         *)
+(* ********************************************************************** *)
+
+lemma (in AC16) Int_in_LL: "w \<in> MM ==> w Int y \<in> LL"
+by (unfold LL_def, fast)
+
+lemma (in AC16) in_LL_eq_Int: 
+     "v \<in> LL ==> v = (THE x. x \<in> MM & v \<subseteq> x) Int y"
+apply (unfold LL_def, clarify)
+apply (subst unique_superset_in_MM [THEN the_equality2])
+apply (auto simp add: Int_in_LL)
+done
+
+lemma (in AC16) unique_superset1: "a \<in> LL \<Longrightarrow> (THE x. x \<in> MM \<and> a \<subseteq> x) \<in> MM"
+by (erule unique_superset_in_MM [THEN theI, THEN conjunct1]); 
+
+lemma (in AC16) the_in_MM_subset:
+     "v \<in> LL ==> (THE x. x \<in> MM & v \<subseteq> x) \<subseteq> x Un y"
+apply (drule unique_superset1)
+apply (unfold MM_def)
+apply (fast dest!: unique_superset1 includes [THEN subsetD])
+done
+
+lemma (in AC16) GG_subset: "v \<in> LL ==> GG ` v \<subseteq> x"
+apply (unfold GG_def)
+apply (frule the_in_MM_subset)
+apply (frule in_LL_eq_Int)
+apply (force elim: equalityE)
+done
+
+lemma (in AC16) nat_lepoll_ordertype: "nat \<lesssim> ordertype(y, R)"
+apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll]) 
+ apply (rule Ord_ordertype [OF WO_R]) 
+apply (rule ordertype_eqpoll [THEN eqpoll_imp_lepoll, THEN lepoll_infinite]) 
+ apply (rule WO_R) 
+apply (rule Infinite) 
+done
+
+lemma (in AC16) ex_subset_eqpoll_n: "n \<in> nat ==> \<exists>z. z \<subseteq> y & n \<approx> z"
+apply (erule nat_lepoll_imp_ex_eqpoll_n)
+apply (rule lepoll_trans [OF nat_lepoll_ordertype]) 
+apply (rule ordertype_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]) 
+apply (rule WO_R) 
+done
+
+
+lemma (in AC16) exists_proper_in_s: "u \<in> x ==> \<exists>v \<in> s(u). succ(k) \<lesssim> v Int y"
+apply (rule ccontr)
+apply (subgoal_tac "\<forall>v \<in> s (u) . k \<approx> v Int y")
+prefer 2 apply (simp add: s_def, blast intro: succ_not_lepoll_imp_eqpoll)
+apply (unfold k_def)
+apply (insert all_ex includes lnat)
+apply (rule ex_subset_eqpoll_n [THEN exE], assumption)
+apply (rule noLepoll [THEN notE])
+apply (blast intro: lepoll_trans [OF y_lepoll_subset_s subset_s_lepoll_w])
+done
+
+lemma (in AC16) exists_in_MM: "u \<in> x ==> \<exists>w \<in> MM. u \<in> w"
+apply (erule exists_proper_in_s [THEN bexE])
+apply (unfold MM_def s_def, fast)
+done
+
+lemma (in AC16) exists_in_LL: "u \<in> x ==> \<exists>w \<in> LL. u \<in> GG`w"
+apply (rule exists_in_MM [THEN bexE], assumption)
+apply (rule bexI)
+apply (erule_tac [2] Int_in_LL)
+apply (unfold GG_def)
+apply (simp add: Int_in_LL)
+apply (subst unique_superset_in_MM [THEN the_equality2])
+apply (fast elim!: Int_in_LL)+
+done
+
+lemma (in AC16) OUN_eq_x: "well_ord(LL,S) ==>       
+      (\<Union>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b)) = x"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule OUN_E)
+apply (rule GG_subset [THEN subsetD])
+prefer 2 apply assumption
+apply (blast intro: ltD  ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun,
+                                       THEN apply_type])
+apply (rule subsetI)
+apply (erule exists_in_LL [THEN bexE])
+apply (force intro: ltI [OF _ Ord_ordertype]
+                    ordermap_type [THEN apply_type]
+             simp add: ordermap_bij [THEN bij_is_inj, THEN left_inverse])
+done
+
+(* ********************************************************************** *)
+(* Every element of the family is less than or equipollent to n-k (m)     *)
+(* ********************************************************************** *)
+
+lemma (in AC16) in_MM_eqpoll_n: "w \<in> MM ==> w \<approx> succ(k #+ m)"
+apply (unfold MM_def)
+apply (fast dest: includes [THEN subsetD])
+done
+
+lemma (in AC16) in_LL_eqpoll_n: "w \<in> LL ==> succ(k) \<lesssim> w"
+by (unfold LL_def MM_def, fast)
+
+lemma (in AC16) in_LL: "w \<in> LL ==> w \<subseteq> (THE x. x \<in> MM \<and> w \<subseteq> x)"
+by (erule subset_trans [OF in_LL_eq_Int [THEN equalityD1] Int_lower1])
+
+lemma (in AC16) all_in_lepoll_m: 
+      "well_ord(LL,S) ==>       
+       \<forall>b<ordertype(LL,S). GG ` (converse(ordermap(LL,S)) ` b) \<lesssim> m"
+apply (unfold GG_def)
+apply (rule oallI)
+apply (simp add: ltD ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN apply_type])
+apply (insert includes)
+apply (rule eqpoll_sum_imp_Diff_lepoll)
+apply (blast del: subsetI
+	     dest!: ltD 
+	     intro!: eqpoll_sum_imp_Diff_lepoll in_LL_eqpoll_n
+	     intro: in_LL   unique_superset1 [THEN in_MM_eqpoll_n] 
+                    ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, 
+                                  THEN apply_type])+
+done
+
+lemma (in AC16) conclusion:
+     "\<exists>a f. Ord(a) & domain(f) = a & (\<Union>b<a. f ` b) = x & (\<forall>b<a. f ` b \<lesssim> m)"
+apply (rule well_ord_LL [THEN exE])
+apply (rename_tac S)
+apply (rule_tac x = "ordertype (LL,S)" in exI)
+apply (rule_tac x = "\<lambda>b \<in> ordertype(LL,S). 
+                      GG ` (converse (ordermap (LL,S)) ` b)"  in exI)
+apply simp
+apply (blast intro: lam_funtype [THEN domain_of_fun] 
+                    Ord_ordertype  OUN_eq_x  all_in_lepoll_m [THEN ospec])
+done
+
+
+(* ********************************************************************** *)
+(* The main theorem AC16(n, k) ==> WO4(n-k)                               *)
+(* ********************************************************************** *)
+
+theorem AC16_WO4: 
+     "[| AC16(k #+ m, k); 0 < k; 0 < m; k \<in> nat; m \<in> nat |] ==> WO4(m)"
+apply (unfold AC16_def WO4_def)
+apply (rule allI)
+apply (case_tac "Finite (A)")
+apply (rule lemma1, assumption+)
+apply (cut_tac lemma2)
+apply (elim exE conjE)
+apply (erule_tac x = "A Un y" in allE)
+apply (frule infinite_Un, drule mp, assumption)
+apply (erule zero_lt_natE, assumption, clarify)
+apply (blast intro: AC16.conclusion) 
+done
 
 end
--- a/src/ZF/AC/AC16_lemmas.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,275 +0,0 @@
-(*  Title:      ZF/AC/AC16_lemmas.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Lemmas used in the proofs concerning AC16
-*)
-
-Goal "a\\<notin>A ==> cons(a,A)-{a}=A";
-by (Fast_tac 1);
-qed "cons_Diff_eq";
-
-Goalw [lepoll_def] "1 lepoll X <-> (\\<exists>x. x \\<in> X)";
-by (rtac iffI 1);
-by (fast_tac (claset() addIs [inj_is_fun RS apply_type]) 1);
-by (etac exE 1);
-by (res_inst_tac [("x","\\<lambda>a \\<in> 1. x")] exI 1);
-by (fast_tac (claset() addSIs [lam_injective]) 1);
-qed "nat_1_lepoll_iff";
-
-Goal "X eqpoll 1 <-> (\\<exists>x. X={x})";
-by (rtac iffI 1);
-by (etac eqpollE 1);
-by (dresolve_tac [nat_1_lepoll_iff RS iffD1] 1);
-by (fast_tac (claset() addSIs [lepoll_1_is_sing]) 1);
-by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1);
-qed "eqpoll_1_iff_singleton";
-
-Goalw [succ_def] "[| x eqpoll n; y\\<notin>x |] ==> cons(y,x) eqpoll succ(n)";
-by (fast_tac (claset() addSEs [cons_eqpoll_cong, mem_irrefl]) 1);
-qed "cons_eqpoll_succ";
-
-Goal "{Y \\<in> Pow(X). Y eqpoll 1} = {{x}. x \\<in> X}";
-by (rtac equalityI 1);
-by (rtac subsetI 1);
-by (etac CollectE 1);
-by (dresolve_tac [eqpoll_1_iff_singleton RS iffD1] 1);
-by (fast_tac (claset() addSIs [RepFunI]) 1);
-by (rtac subsetI 1);
-by (etac RepFunE 1);
-by (rtac CollectI 1);
-by (Fast_tac 1);
-by (fast_tac (claset() addSIs [singleton_eqpoll_1]) 1);
-qed "subsets_eqpoll_1_eq";
-
-Goalw [eqpoll_def, bij_def] "X eqpoll {{x}. x \\<in> X}";
-by (res_inst_tac [("x","\\<lambda>x \\<in> X. {x}")] exI 1);
-by (rtac IntI 1);
-by (rewrite_goals_tac [inj_def, surj_def]);
-by (Asm_full_simp_tac 1);
-by (fast_tac (claset() addSIs [lam_type, RepFunI] 
-                addIs [singleton_eq_iff RS iffD1]) 1);
-by (Asm_full_simp_tac 1);
-by (fast_tac (claset() addSIs [lam_type]) 1);
-qed "eqpoll_RepFun_sing";
-
-Goal "{Y \\<in> Pow(X). Y eqpoll 1} eqpoll X";
-by (resolve_tac [subsets_eqpoll_1_eq RS ssubst] 1);
-by (resolve_tac [eqpoll_RepFun_sing RS eqpoll_sym] 1);
-qed "subsets_eqpoll_1_eqpoll";
-
-Goal "[| InfCard(x); y \\<in> Pow(x); y eqpoll succ(z) |]  \
-\               ==> (LEAST i. i \\<in> y) \\<in> y";
-by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS 
-                succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (fast_tac (claset() addIs [LeastI]
-        addSDs [InfCard_is_Card RS Card_is_Ord, PowD RS subsetD]
-        addEs [Ord_in_Ord]) 1);
-qed "InfCard_Least_in";
-
-Goalw [lepoll_def] "[| InfCard(x); n \\<in> nat |] ==>  \
-\       {y \\<in> Pow(x). y eqpoll succ(succ(n))} lepoll  \
-\       x*{y \\<in> Pow(x). y eqpoll succ(n)}";
-by (res_inst_tac [("x","\\<lambda>y \\<in> {y \\<in> Pow(x). y eqpoll succ(succ(n))}. \
-\               <LEAST i. i \\<in> y, y-{LEAST i. i \\<in> y}>")] exI 1);
-by (res_inst_tac [("d","%z. cons(fst(z),snd(z))")] lam_injective 1);
-by (rtac SigmaI 1);
-by (etac CollectE 1);
-by (Asm_full_simp_tac 3);
-by (rtac equalityI 3);
-by (Fast_tac 4);
-by (rtac subsetI 3);
-by (etac consE 3);
-by (Fast_tac 4);
-by (rtac CollectI 2);
-by (Fast_tac 2);
-by (resolve_tac [PowD RS subsetD] 1 THEN (assume_tac 1));
-by (REPEAT (fast_tac (claset() addSIs [Diff_sing_eqpoll]
-                addIs [InfCard_Least_in]) 1));
-qed "subsets_lepoll_lemma1";
-
-val prems = goal thy "(!!y. y \\<in> z ==> Ord(y)) ==> z \\<subseteq> succ(Union(z))";
-by (rtac subsetI 1);
-by (res_inst_tac [("Q","\\<forall>y \\<in> z. y \\<subseteq> x")] (excluded_middle RS disjE) 1);
-by (Fast_tac 2);
-by (etac swap 1);
-by (rtac ballI 1);
-by (rtac Ord_linear_le 1);
-by (dtac le_imp_subset 3 THEN (assume_tac 3));
-by (fast_tac (claset() addDs prems) 1);
-by (fast_tac (claset() addDs prems) 1);
-by (fast_tac (claset() addSEs [leE,ltE]) 1);
-qed "set_of_Ord_succ_Union";
-
-Goal "j \\<subseteq> i ==> i \\<notin> j";
-by (fast_tac (claset() addSEs [mem_irrefl]) 1);
-qed "subset_not_mem";
-
-val prems = goal thy "(!!y. y \\<in> z ==> Ord(y)) ==> succ(Union(z)) \\<notin> z";
-by (resolve_tac [set_of_Ord_succ_Union RS subset_not_mem] 1);
-by (eresolve_tac prems 1);
-qed "succ_Union_not_mem";
-
-Goal "Union(cons(succ(Union(z)),z)) = succ(Union(z))";
-by (Fast_tac 1);
-qed "Union_cons_eq_succ_Union";
-
-Goal "[| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
-by (fast_tac (claset() addSDs [le_imp_subset] addEs [Ord_linear_le]) 1);
-qed "Un_Ord_disj";
-
-Goal "x \\<in> X ==> Union(X) = x Un Union(X-{x})";
-by (Fast_tac 1);
-qed "Union_eq_Un";
-
-Goal "n \\<in> nat ==>  \
-\       \\<forall>z. (\\<forall>y \\<in> z. Ord(y)) & z eqpoll n & z\\<noteq>0 --> Union(z) \\<in> z";
-by (induct_tac "n" 1);
-by (fast_tac (claset() addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
-by (REPEAT (resolve_tac [allI, impI] 1));
-by (etac natE 1);
-by (fast_tac (claset() addSDs [eqpoll_1_iff_singleton RS iffD1]
-        addSIs [Union_singleton]) 1);
-by (hyp_subst_tac 1);
-by (REPEAT (eresolve_tac [conjE, not_emptyE] 1));
-by (eres_inst_tac [("x","z-{xb}")] allE 1);
-by (etac impE 1);
-by (fast_tac (claset() addSEs [Diff_sing_eqpoll,
-                Diff_sing_eqpoll RS eqpoll_succ_imp_not_empty]) 1);
-by (resolve_tac [Union_eq_Un RSN (2, subst_elem)] 1 THEN (assume_tac 2));
-by (ftac bspec 1 THEN (assume_tac 1));
-by (dresolve_tac [Diff_subset RS subsetD RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (dtac Un_Ord_disj 1 THEN (assume_tac 1));
-by (etac DiffE 1);
-by (etac disjE 1);
-by (etac subst_elem 1 THEN (assume_tac 1));
-by (rtac subst_elem 1 THEN (TRYALL assume_tac));
-qed "Union_in_lemma";
-
-Goal "[| \\<forall>x \\<in> z. Ord(x); z eqpoll n; z\\<noteq>0; n \\<in> nat |]  \
-\               ==> Union(z) \\<in> z";
-by (dtac Union_in_lemma 1);
-by (fast_tac FOL_cs 1);
-qed "Union_in";
-
-Goal "[| InfCard(x); z \\<in> Pow(x); z eqpoll n; n \\<in> nat |]  \
-\               ==> succ(Union(z)) \\<in> x";
-by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3));
-by (etac InfCard_is_Limit 1);
-by (excluded_middle_tac "z=0" 1);
-by (fast_tac (claset() addSIs [InfCard_is_Limit RS Limit_has_0]
-                      addss (simpset())) 2);
-by (resolve_tac
-        [PowD RS subsetD RS (InfCard_is_Card RS Card_is_Ord RSN (2, ltI))] 1
-        THEN (TRYALL assume_tac));
-by (fast_tac (claset() addSIs [Union_in]
-                      addSEs [PowD RS subsetD RSN 
-		 (2, InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1);
-qed "succ_Union_in_x";
-
-Goalw [lepoll_def] "[| InfCard(x); n \\<in> nat |] ==>  \
-\       {y \\<in> Pow(x). y eqpoll succ(n)} lepoll  \
-\       {y \\<in> Pow(x). y eqpoll succ(succ(n))}";
-by (res_inst_tac [("x","\\<lambda>z \\<in> {y \\<in> Pow(x). y eqpoll succ(n)}.  \
-\       cons(succ(Union(z)), z)")] exI 1);
-by (res_inst_tac [("d","%z. z-{Union(z)}")] lam_injective 1);
-by (resolve_tac [Union_cons_eq_succ_Union RS ssubst] 2);
-by (rtac cons_Diff_eq 2);
-by (fast_tac (claset() addSDs [InfCard_is_Card RS Card_is_Ord]
-        addEs [Ord_in_Ord] addSIs [succ_Union_not_mem]) 2);
-by (rtac CollectI 1);
-by (fast_tac (claset() addSEs [cons_eqpoll_succ] 
-                    addSIs [succ_Union_not_mem] 
-                    addSDs [InfCard_is_Card RS Card_is_Ord] 
-                    addEs  [Ord_in_Ord]) 2);
-by (fast_tac (claset() addSIs [succ_Union_in_x]) 1);
-qed "succ_lepoll_succ_succ";
-
-Goal "[| InfCard(X); n \\<in> nat |]  \
-\       ==> {Y \\<in> Pow(X). Y eqpoll succ(n)} eqpoll X";
-by (induct_tac "n" 1);
-by (rtac subsets_eqpoll_1_eqpoll 1);
-by (rtac eqpollI 1);
-by (resolve_tac [subsets_lepoll_lemma1 RS lepoll_trans] 1 
-        THEN (REPEAT (assume_tac 1)));
-by (resolve_tac [InfCard_is_Card RS Card_is_Ord RS well_ord_Memrel RS
-                well_ord_InfCard_square_eq RS eqpoll_imp_lepoll
-                RSN (2, lepoll_trans)] 1 THEN (assume_tac 2));
-by (resolve_tac [InfCard_is_Card RS Card_cardinal_eq RS ssubst] 2 
-        THEN (REPEAT (assume_tac 2)));
-by (eresolve_tac [eqpoll_refl RS prod_eqpoll_cong RS eqpoll_imp_lepoll] 1);
-by (fast_tac (claset() addEs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans]
-        addSIs [succ_lepoll_succ_succ]) 1);
-qed "subsets_eqpoll_X";
-
-Goalw [surj_def] "[| f \\<in> surj(A,B); y \\<subseteq> B |]  \
-\       ==> f``(converse(f)``y) = y";
-by (fast_tac (claset() addDs [apply_equality2]
-	              addEs [apply_iff RS iffD2]) 1);
-qed "image_vimage_eq";
-
-Goal "[| f \\<in> inj(A,B); y \\<subseteq> A |] ==> converse(f)``(f``y) = y";
-by (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair]
-                addDs [inj_equality]) 1);
-qed "vimage_image_eq";
-
-Goalw [eqpoll_def] "A eqpoll B  \
-\       ==> {Y \\<in> Pow(A). Y eqpoll n} eqpoll {Y \\<in> Pow(B). Y eqpoll n}";
-by (etac exE 1);
-by (res_inst_tac [("x","\\<lambda>X \\<in> {Y \\<in> Pow(A). \\<exists>f. f \\<in> bij(Y, n)}. f``X")] exI 1);
-by (res_inst_tac [("d","%Z. converse(f)``Z")] lam_bijective 1);
-by (fast_tac (claset()
-        addSIs [bij_is_inj RS restrict_bij RS bij_converse_bij RS comp_bij] 
-        addSEs [bij_is_fun RS fun_is_rel RS image_subset RS PowI]) 1);
-by (fast_tac (claset() addSIs [bij_converse_bij RS bij_is_inj RS restrict_bij
-                        RS bij_converse_bij RS comp_bij] 
-                    addSEs [bij_converse_bij RS bij_is_fun RS fun_is_rel
-                        RS image_subset RS PowI]) 1);
-by (fast_tac (claset() addSEs [bij_is_inj RS vimage_image_eq]) 1);
-by (fast_tac (claset() addSEs [bij_is_surj RS image_vimage_eq]) 1);
-qed "subsets_eqpoll";
-
-Goalw [WO2_def] "WO2 ==> \\<exists>a. Card(a) & X eqpoll a";
-by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
-by (fast_tac (claset() addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS
-                (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym]
-                addSIs [Card_cardinal]) 1);
-qed "WO2_imp_ex_Card";
-
-Goal "[| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)";
-by (fast_tac (empty_cs addEs [notE, lepoll_Finite] addSIs [notI]) 1); 
-qed "lepoll_infinite";
-
-Goalw [InfCard_def] "[| ~Finite(X); Card(X) |] ==> InfCard(X)";
-by (fast_tac (claset() addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1);
-qed "infinite_Card_is_InfCard";
-
-Goal "[| WO2; n \\<in> nat; ~Finite(X) |]  \
-\       ==> {Y \\<in> Pow(X). Y eqpoll succ(n)} eqpoll X";
-by (dtac WO2_imp_ex_Card 1);
-by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
-by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1));
-by (dtac infinite_Card_is_InfCard 1 THEN (assume_tac 1));
-by (resolve_tac [eqpoll_trans RS eqpoll_trans] 1);
-by (etac subsets_eqpoll 1);
-by (etac subsets_eqpoll_X 1 THEN (assume_tac 1));
-by (etac eqpoll_sym 1);
-qed "WO2_infinite_subsets_eqpoll_X";
-
-Goal "well_ord(X,R) ==> \\<exists>a. Card(a) & X eqpoll a";
-by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym]
-                addSIs [Card_cardinal]) 1);
-qed "well_ord_imp_ex_Card";
-
-Goal "[| well_ord(X,R); n \\<in> nat; ~Finite(X) |]  \
-\               ==> {Y \\<in> Pow(X). Y eqpoll succ(n)} eqpoll X";
-by (dtac well_ord_imp_ex_Card 1);
-by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
-by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1));
-by (dtac infinite_Card_is_InfCard 1 THEN (assume_tac 1));
-by (resolve_tac [eqpoll_trans RS eqpoll_trans] 1);
-by (etac subsets_eqpoll 1);
-by (etac subsets_eqpoll_X 1 THEN (assume_tac 1));
-by (etac eqpoll_sym 1);
-qed "well_ord_infinite_subsets_eqpoll_X";
-
--- a/src/ZF/AC/AC16_lemmas.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC16_lemmas.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,244 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/AC16_lemmas.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+Lemmas used in the proofs concerning AC16
+*)
+
+theory AC16_lemmas = AC_Equiv + Hartog + Cardinal_aux:
+
+lemma cons_Diff_eq: "a\<notin>A ==> cons(a,A)-{a}=A"
+by fast
+
+lemma nat_1_lepoll_iff: "1\<lesssim>X <-> (\<exists>x. x \<in> X)"
+apply (unfold lepoll_def)
+apply (rule iffI)
+apply (fast intro: inj_is_fun [THEN apply_type])
+apply (erule exE)
+apply (rule_tac x = "\<lambda>a \<in> 1. x" in exI)
+apply (fast intro!: lam_injective)
+done
+
+lemma eqpoll_1_iff_singleton: "X\<approx>1 <-> (\<exists>x. X={x})"
+apply (rule iffI)
+apply (erule eqpollE)
+apply (drule nat_1_lepoll_iff [THEN iffD1])
+apply (fast intro!: lepoll_1_is_sing)
+apply (fast intro!: singleton_eqpoll_1)
+done
+
+lemma cons_eqpoll_succ: "[| x\<approx>n; y\<notin>x |] ==> cons(y,x)\<approx>succ(n)"
+apply (unfold succ_def)
+apply (fast elim!: cons_eqpoll_cong mem_irrefl)
+done
+
+lemma subsets_eqpoll_1_eq: "{Y \<in> Pow(X). Y\<approx>1} = {{x}. x \<in> X}"
+apply (rule equalityI)
+apply (rule subsetI)
+apply (erule CollectE)
+apply (drule eqpoll_1_iff_singleton [THEN iffD1])
+apply (fast intro!: RepFunI)
+apply (rule subsetI)
+apply (erule RepFunE)
+apply (rule CollectI)
+apply fast
+apply (fast intro!: singleton_eqpoll_1)
+done
+
+lemma eqpoll_RepFun_sing: "X\<approx>{{x}. x \<in> X}"
+apply (unfold eqpoll_def bij_def)
+apply (rule_tac x = "\<lambda>x \<in> X. {x}" in exI)
+apply (rule IntI)
+apply (unfold inj_def surj_def)
+apply simp
+apply (fast intro!: lam_type RepFunI intro: singleton_eq_iff [THEN iffD1])
+apply simp
+apply (fast intro!: lam_type)
+done
+
+lemma subsets_eqpoll_1_eqpoll: "{Y \<in> Pow(X). Y\<approx>1}\<approx>X"
+apply (rule subsets_eqpoll_1_eq [THEN ssubst])
+apply (rule eqpoll_RepFun_sing [THEN eqpoll_sym])
+done
+
+lemma InfCard_Least_in:
+     "[| InfCard(x); y \<subseteq> x; y \<approx> succ(z) |] ==> (LEAST i. i \<in> y) \<in> y"
+apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, 
+                         THEN succ_lepoll_imp_not_empty, THEN not_emptyE])
+apply (fast intro: LeastI 
+            dest!: InfCard_is_Card [THEN Card_is_Ord] 
+            elim: Ord_in_Ord)
+done
+
+lemma subsets_lepoll_lemma1:
+     "[| InfCard(x); n \<in> nat |] 
+      ==> {y \<in> Pow(x). y\<approx>succ(succ(n))} \<lesssim> x*{y \<in> Pow(x). y\<approx>succ(n)}"
+apply (unfold lepoll_def)
+apply (rule_tac x = "\<lambda>y \<in> {y \<in> Pow(x) . y\<approx>succ (succ (n))}. 
+                      <LEAST i. i \<in> y, y-{LEAST i. i \<in> y}>" in exI)
+apply (rule_tac d = "%z. cons (fst(z), snd(z))" in lam_injective);
+ apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in);
+apply (simp, blast intro: InfCard_Least_in);
+done
+
+lemma set_of_Ord_succ_Union: "(\<forall>y \<in> z. Ord(y)) ==> z \<subseteq> succ(Union(z))"
+apply (rule subsetI)
+apply (case_tac "\<forall>y \<in> z. y \<subseteq> x", blast );
+apply (simp, erule bexE); 
+apply (rule_tac i=xa and j=x in Ord_linear_le)
+apply (blast dest: le_imp_subset elim: leE ltE)+
+done
+
+lemma subset_not_mem: "j \<subseteq> i ==> i \<notin> j"
+by (fast elim!: mem_irrefl)
+
+lemma succ_Union_not_mem:
+     "(!!y. y \<in> z ==> Ord(y)) ==> succ(Union(z)) \<notin> z"
+apply (rule set_of_Ord_succ_Union [THEN subset_not_mem]);
+apply blast
+done
+
+lemma Union_cons_eq_succ_Union:
+     "Union(cons(succ(Union(z)),z)) = succ(Union(z))"
+by fast
+
+lemma Un_Ord_disj: "[| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j"
+by (fast dest!: le_imp_subset elim: Ord_linear_le)
+
+lemma Union_eq_Un: "x \<in> X ==> Union(X) = x Un Union(X-{x})"
+by fast
 
-AC16_lemmas = AC_Equiv + Hartog
+lemma Union_in_lemma [rule_format]:
+     "n \<in> nat ==> \<forall>z. (\<forall>y \<in> z. Ord(y)) & z\<approx>n & z\<noteq>0 --> Union(z) \<in> z"
+apply (induct_tac "n")
+apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
+apply (intro allI impI)
+apply (erule natE)
+apply (fast dest!: eqpoll_1_iff_singleton [THEN iffD1]
+            intro!: Union_singleton)
+apply (clarify ); 
+apply (elim not_emptyE)
+apply (erule_tac x = "z-{xb}" in allE)
+apply (erule impE)
+apply (fast elim!: Diff_sing_eqpoll
+                   Diff_sing_eqpoll [THEN eqpoll_succ_imp_not_empty])
+apply (subgoal_tac "xb \<union> \<Union>(z - {xb}) \<in> z");
+apply (simp add: Union_eq_Un [symmetric]);
+apply (frule bspec, assumption)
+apply (drule bspec); 
+apply (erule Diff_subset [THEN subsetD]);
+apply (drule Un_Ord_disj, assumption)
+apply (auto ); 
+done
+
+lemma Union_in: "[| \<forall>x \<in> z. Ord(x); z\<approx>n; z\<noteq>0; n \<in> nat |] ==> Union(z) \<in> z"
+apply (blast intro: Union_in_lemma); 
+done
+
+lemma succ_Union_in_x:
+     "[| InfCard(x); z \<in> Pow(x); z\<approx>n; n \<in> nat |] ==> succ(Union(z)) \<in> x"
+apply (rule Limit_has_succ [THEN ltE]);
+prefer 3 apply assumption
+apply (erule InfCard_is_Limit)
+apply (case_tac "z=0");
+apply (simp, fast intro!: InfCard_is_Limit [THEN Limit_has_0]);
+apply (rule ltI [OF PowD [THEN subsetD] InfCard_is_Card [THEN Card_is_Ord]]);
+apply assumption; 
+apply (blast intro: Union_in
+                    InfCard_is_Card [THEN Card_is_Ord, THEN Ord_in_Ord])+
+done
+
+lemma succ_lepoll_succ_succ:
+     "[| InfCard(x); n \<in> nat |] 
+      ==> {y \<in> Pow(x). y\<approx>succ(n)} \<lesssim> {y \<in> Pow(x). y\<approx>succ(succ(n))}"
+apply (unfold lepoll_def);
+apply (rule_tac x = "\<lambda>z \<in> {y\<in>Pow(x). y\<approx>succ(n)}. cons(succ(Union(z)), z)" 
+       in exI)
+apply (rule_tac d = "%z. z-{Union (z) }" in lam_injective)
+apply (blast intro!: succ_Union_in_x succ_Union_not_mem
+             intro: cons_eqpoll_succ Ord_in_Ord
+             dest!: InfCard_is_Card [THEN Card_is_Ord])
+apply (simp only: Union_cons_eq_succ_Union); 
+apply (rule cons_Diff_eq);
+apply (fast dest!: InfCard_is_Card [THEN Card_is_Ord]
+            elim: Ord_in_Ord 
+            intro!: succ_Union_not_mem);
+done
+
+lemma subsets_eqpoll_X:
+     "[| InfCard(X); n \<in> nat |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)} \<approx> X"
+apply (induct_tac "n")
+apply (rule subsets_eqpoll_1_eqpoll)
+apply (rule eqpollI)
+apply (rule subsets_lepoll_lemma1 [THEN lepoll_trans], assumption+);
+apply (rule eqpoll_trans [THEN eqpoll_imp_lepoll]); 
+ apply (erule eqpoll_refl [THEN prod_eqpoll_cong]);
+apply (erule InfCard_square_eqpoll)
+apply (fast elim: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] 
+            intro!: succ_lepoll_succ_succ)
+done
+
+lemma image_vimage_eq:
+     "[| f \<in> surj(A,B); y \<subseteq> B |] ==> f``(converse(f)``y) = y"
+apply (unfold surj_def)
+apply (fast dest: apply_equality2 elim: apply_iff [THEN iffD2])
+done
+
+lemma vimage_image_eq: "[| f \<in> inj(A,B); y \<subseteq> A |] ==> converse(f)``(f``y) = y"
+apply (fast elim!: inj_is_fun [THEN apply_Pair] dest: inj_equality)
+done
+
+lemma subsets_eqpoll:
+     "A\<approx>B ==> {Y \<in> Pow(A). Y\<approx>n}\<approx>{Y \<in> Pow(B). Y\<approx>n}"
+apply (unfold eqpoll_def)
+apply (erule exE)
+apply (rule_tac x = "\<lambda>X \<in> {Y \<in> Pow (A) . \<exists>f. f \<in> bij (Y, n) }. f``X" in exI)
+apply (rule_tac d = "%Z. converse (f) ``Z" in lam_bijective)
+apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij, 
+                                THEN comp_bij] 
+            elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset])
+apply (blast intro!:  bij_is_inj [THEN restrict_bij] 
+                      comp_bij bij_converse_bij
+                      bij_is_fun [THEN fun_is_rel, THEN image_subset])
+apply (fast elim!: bij_is_inj [THEN vimage_image_eq])
+apply (fast elim!: bij_is_surj [THEN image_vimage_eq])
+done
+
+lemma WO2_imp_ex_Card: "WO2 ==> \<exists>a. Card(a) & X\<approx>a"
+apply (unfold WO2_def)
+apply (drule spec [of _ X])
+apply (blast intro: Card_cardinal eqpoll_trans
+          well_ord_Memrel [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym])
+done
+
+lemma lepoll_infinite: "[| X\<lesssim>Y; ~Finite(X) |] ==> ~Finite(Y)"
+by (blast intro: lepoll_Finite)
+
+lemma infinite_Card_is_InfCard: "[| ~Finite(X); Card(X) |] ==> InfCard(X)"
+apply (unfold InfCard_def)
+apply (fast elim!: Card_is_Ord [THEN nat_le_infinite_Ord])
+done
+
+lemma WO2_infinite_subsets_eqpoll_X: "[| WO2; n \<in> nat; ~Finite(X) |]   
+        ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
+apply (drule WO2_imp_ex_Card)
+apply (elim allE exE conjE);
+apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)
+apply (drule infinite_Card_is_InfCard, assumption)
+apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans); 
+done
+
+lemma well_ord_imp_ex_Card: "well_ord(X,R) ==> \<exists>a. Card(a) & X\<approx>a"
+by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym] 
+         intro!: Card_cardinal)
+
+lemma well_ord_infinite_subsets_eqpoll_X:
+     "[| well_ord(X,R); n \<in> nat; ~Finite(X) |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
+apply (drule well_ord_imp_ex_Card)
+apply (elim allE exE conjE)
+apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)
+apply (drule infinite_Card_is_InfCard, assumption)
+apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans); 
+done
+
+end
--- a/src/ZF/AC/AC17_AC1.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,310 +0,0 @@
-(*  Title:      ZF/AC/AC1_AC17.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The equivalence of AC0, AC1 and AC17
-
-Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent
-to AC0 and AC1.
-*)
-
-
-(** AC0 is equivalent to AC1.  
-    AC0 comes from Suppes, AC1 from Rubin & Rubin **)
-
-Goal "[| f:(\\<Pi>X \\<in> A. X); D \\<subseteq> A |] ==> \\<exists>g. g:(\\<Pi>X \\<in> D. X)";
-by (fast_tac (claset() addSIs [restrict_type, apply_type]) 1);
-val lemma1 = result();
-
-Goalw AC_defs "AC0 ==> AC1"; 
-by (blast_tac (claset() addIs [lemma1]) 1); 
-qed "AC0_AC1";
-
-Goalw AC_defs "AC1 ==> AC0";
-by (Blast_tac 1); 
-qed "AC1_AC0";
-
-
-(**** The proof of AC1 ==> AC17 ****)
-
-Goal "f \\<in> (\\<Pi>X \\<in> Pow(A) - {0}. X) ==> f \\<in> (Pow(A) - {0} -> A)";
-by (rtac Pi_type 1 THEN (assume_tac 1));
-by (dtac apply_type 1 THEN (assume_tac 1));
-by (Fast_tac 1);
-val lemma1 = result();
-
-Goalw AC_defs "AC1 ==> AC17";
-by (rtac allI 1);
-by (rtac ballI 1);
-by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
-by (etac impE 1);
-by (Fast_tac 1);
-by (etac exE 1);
-by (rtac bexI 1);
-by (etac lemma1 2);
-by (rtac apply_type 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSDs [lemma1] addSEs [apply_type]) 1);
-qed "AC1_AC17";
-
-
-(**** The proof of AC17 ==> AC1 ****)
-
-(* *********************************************************************** *)
-(* more properties of HH                                                   *)
-(* *********************************************************************** *)
-
-Goal "[| x - (\\<Union>j \\<in> LEAST i. HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X}, x, i) = {x}. \
-\       HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X}, x, j)) = 0;  \
-\       f \\<in> Pow(x)-{0} -> x |]  \
-\       ==> \\<exists>r. well_ord(x,r)";
-by (rtac exI 1);
-by (eresolve_tac [[bij_Least_HH_x RS bij_converse_bij RS bij_is_inj,
-                Ord_Least RS well_ord_Memrel] MRS well_ord_rvimage] 1);
-by (assume_tac 1);
-qed "UN_eq_imp_well_ord";
-
-(* *********************************************************************** *)
-(* theorems closer to the proof                                            *)
-(* *********************************************************************** *)
-
-Goalw AC_defs "~AC1 ==>  \
-\               \\<exists>A. \\<forall>f \\<in> Pow(A)-{0} -> A. \\<exists>u \\<in> Pow(A)-{0}. f`u \\<notin> u";
-by (etac swap 1);
-by (rtac allI 1);
-by (etac swap 1);
-by (res_inst_tac [("x","Union(A)")] exI 1);
-by (rtac ballI 1);
-by (etac swap 1);
-by (rtac impI 1);
-by (fast_tac (claset() addSIs [restrict_type]) 1);
-qed "not_AC1_imp_ex";
-
-Goal "[| \\<forall>f \\<in> Pow(x) - {0} -> x. \\<exists>u \\<in> Pow(x) - {0}. f`u\\<notin>u;  \
-\       \\<exists>f \\<in> Pow(x)-{0}->x. \
-\       x - (\\<Union>a \\<in> (LEAST i. HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X},x,i)={x}).  \
-\       HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X},x,a)) = 0 |] \
-\       ==> P";
-by (etac bexE 1);
-by (eresolve_tac [UN_eq_imp_well_ord RS exE] 1 THEN (assume_tac 1));
-by (eresolve_tac [ex_choice_fun_Pow RS exE] 1);
-by (etac ballE 1);
-by (fast_tac (FOL_cs addEs [bexE, notE, apply_type]) 1);
-by (etac notE 1);
-by (rtac Pi_type 1 THEN (assume_tac 1));
-by (resolve_tac [apply_type RSN (2, subsetD)] 1 THEN TRYALL assume_tac);
-by (Fast_tac 1);
-val lemma1 = result();
-
-Goal "~ (\\<exists>f \\<in> Pow(x)-{0}->x. x - F(f) = 0)  \
-\       ==> (\\<lambda>f \\<in> Pow(x)-{0}->x. x - F(f))  \
-\               \\<in> (Pow(x) -{0} -> x) -> Pow(x) - {0}";
-by (fast_tac (claset() addSIs [lam_type] addSDs [Diff_eq_0_iff RS iffD1]) 1);
-val lemma2 = result();
-
-Goal "[| f`Z \\<in> Z; Z \\<in> Pow(x)-{0} |] ==>  \
-\       (\\<lambda>X \\<in> Pow(x)-{0}. {f`X})`Z \\<in> Pow(Z)-{0}";
-by Auto_tac;
-val lemma3 = result();
-
-Goal "\\<exists>f \\<in> F. f`((\\<lambda>f \\<in> F. Q(f))`f) \\<in> (\\<lambda>f \\<in> F. Q(f))`f  \
-\       ==> \\<exists>f \\<in> F. f`Q(f) \\<in> Q(f)";
-by (Asm_full_simp_tac 1);
-val lemma4 = result();
-
-Goalw [AC17_def] "AC17 ==> AC1";
-by (rtac classical 1);
-by (eresolve_tac [not_AC1_imp_ex RS exE] 1);
-by (excluded_middle_tac
-        "\\<exists>f \\<in> Pow(x)-{0}->x. \
-\       x - (\\<Union>a \\<in> (LEAST i. HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X},x,i)={x}).  \
-\       HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X},x,a)) = 0" 1);
-by (etac lemma1 2 THEN (assume_tac 2));
-by (dtac lemma2 1);
-by (etac allE 1);
-by (dtac bspec 1 THEN (assume_tac 1));
-by (dtac lemma4 1);
-by (etac bexE 1);
-by (dtac apply_type 1 THEN (assume_tac 1));
-by (dresolve_tac [beta RS sym RSN (2, subst_elem)] 1);
-by (assume_tac 1);
-by (dtac lemma3 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem),
-                f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1);
-qed "AC17_AC1";
-
-
-(* **********************************************************************
-    AC1 ==> AC2 ==> AC1
-    AC1 ==> AC4 ==> AC3 ==> AC1
-    AC4 ==> AC5 ==> AC4
-    AC1 <-> AC6
-************************************************************************* *)
-
-(* ********************************************************************** *)
-(* AC1 ==> AC2                                                            *)
-(* ********************************************************************** *)
-
-Goal "[| f:(\\<Pi>X \\<in> A. X);  B \\<in> A;  0\\<notin>A |] ==> {f`B} \\<subseteq> B Int {f`C. C \\<in> A}";
-by (fast_tac (claset() addSEs [apply_type]) 1);
-val lemma1 = result();
-
-Goalw [pairwise_disjoint_def]
-        "[| pairwise_disjoint(A); B \\<in> A; C \\<in> A; D \\<in> B; D \\<in> C |] ==> f`B = f`C";
-by (Fast_tac 1);
-val lemma2 = result();
-
-Goalw AC_defs "AC1 ==> AC2"; 
-by (rtac allI 1);
-by (rtac impI 1);
-by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
-by (REPEAT (resolve_tac [exI,ballI,equalityI] 1));
-by (rtac lemma1 2 THEN (REPEAT (assume_tac 2)));
-by (fast_tac (claset() addSEs [lemma2] addEs [apply_type]) 1);
-qed "AC1_AC2";
-
-
-(* ********************************************************************** *)
-(* AC2 ==> AC1                                                            *)
-(* ********************************************************************** *)
-
-Goal "0\\<notin>A ==> 0 \\<notin> {B*{B}. B \\<in> A}";
-by (fast_tac (claset() addSDs [sym RS (Sigma_empty_iff RS iffD1)]) 1);
-val lemma1 = result();
-
-Goal "[| X*{X} Int C = {y}; X \\<in> A |]  \
-\               ==> (THE y. X*{X} Int C = {y}): X*A";
-by (rtac subst_elem 1);
-by (fast_tac (claset() addSIs [the_equality]
-                addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2);
-by (blast_tac (claset() addSEs [equalityE]) 1);
-val lemma2 = result();
-
-Goal "\\<forall>D \\<in> {E*{E}. E \\<in> A}. \\<exists>y. D Int C = {y}  \
-\     ==> (\\<lambda>x \\<in> A. fst(THE z. (x*{x} Int C = {z}))) \\<in> (\\<Pi>X \\<in> A. X)";
-by (fast_tac (claset() addSEs [lemma2] 
-                       addSIs [lam_type, RepFunI, fst_type]) 1);
-val lemma3 = result();
-
-Goalw (AC_defs@AC_aux_defs) "AC2 ==> AC1";
-by (REPEAT (resolve_tac [allI, impI] 1));
-by (REPEAT (eresolve_tac [allE, impE] 1));
-by (fast_tac (claset() addSEs [lemma3]) 2);
-by (fast_tac (claset() addSIs [lemma1, equals0I]) 1);
-qed "AC2_AC1";
-
-
-(* ********************************************************************** *)
-(* AC1 ==> AC4                                                            *)
-(* ********************************************************************** *)
-
-Goal "0 \\<notin> {R``{x}. x \\<in> domain(R)}";
-by (Blast_tac 1);
-val lemma = result();
-
-Goalw AC_defs "AC1 ==> AC4";
-by (REPEAT (resolve_tac [allI, impI] 1));
-by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
-by (best_tac (claset() addSIs [lam_type] addSEs [apply_type]) 1);
-qed "AC1_AC4";
-
-
-(* ********************************************************************** *)
-(* AC4 ==> AC3                                                            *)
-(* ********************************************************************** *)
-
-Goal "f \\<in> A->B ==> (\\<Union>z \\<in> A. {z}*f`z) \\<subseteq> A*Union(B)";
-by (fast_tac (claset() addSDs [apply_type]) 1);
-val lemma1 = result();
-
-Goal "domain(\\<Union>z \\<in> A. {z}*f(z)) = {a \\<in> A. f(a)\\<noteq>0}";
-by (Blast_tac 1);
-val lemma2 = result();
-
-Goal "x \\<in> A ==> (\\<Union>z \\<in> A. {z}*f(z))``{x} = f(x)";
-by (Fast_tac 1);
-val lemma3 = result();
-
-Goalw AC_defs "AC4 ==> AC3";
-by (REPEAT (resolve_tac [allI,ballI] 1));
-by (REPEAT (eresolve_tac [allE,impE] 1));
-by (etac lemma1 1);
-by (asm_full_simp_tac (simpset() addsimps [lemma2, lemma3]
-                                 addcongs [Pi_cong]) 1);
-qed "AC4_AC3";
-
-(* ********************************************************************** *)
-(* AC3 ==> AC1                                                            *)
-(* ********************************************************************** *)
-
-Goal "b\\<notin>A ==> (\\<Pi>x \\<in> {a \\<in> A. id(A)`a\\<noteq>b}. id(A)`x) = (\\<Pi>x \\<in> A. x)";
-by (asm_full_simp_tac (simpset() addsimps [id_def] addcongs [Pi_cong]) 1);
-by (res_inst_tac [("b","A")] subst_context 1);
-by (Fast_tac 1);
-val lemma = result();
-
-Goalw AC_defs "AC3 ==> AC1";
-by (fast_tac (claset() addSIs [id_type] addEs [lemma RS subst]) 1);
-qed "AC3_AC1";
-
-(* ********************************************************************** *)
-(* AC4 ==> AC5                                                            *)
-(* ********************************************************************** *)
-
-Goalw (range_def::AC_defs) "AC4 ==> AC5";
-by (REPEAT (resolve_tac [allI,ballI] 1));
-by (REPEAT (eresolve_tac [allE,impE] 1));
-by (eresolve_tac [fun_is_rel RS converse_type] 1);
-by (etac exE 1);
-by (rtac bexI 1);
-by (rtac Pi_type 2 THEN (assume_tac 2));
-by (fast_tac (claset() addSDs [apply_type]
-        addSEs [fun_is_rel RS converse_type RS subsetD RS SigmaD2]) 2);
-by (rtac ballI 1);
-by (rtac apply_equality 1 THEN (assume_tac 2));
-by (etac domainE 1);
-by (ftac range_type 1 THEN (assume_tac 1));
-by (fast_tac (claset() addDs [apply_equality]) 1);
-qed "AC4_AC5";
-
-
-(* ********************************************************************** *)
-(* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
-(* ********************************************************************** *)
-
-Goal "R \\<subseteq> A*B ==> (\\<lambda>x \\<in> R. fst(x)) \\<in> R -> A";
-by (fast_tac (claset() addSIs [lam_type, fst_type]) 1);
-val lemma1 = result();
-
-Goalw [range_def] "R \\<subseteq> A*B ==> range(\\<lambda>x \\<in> R. fst(x)) = domain(R)";
-by (force_tac (claset() addIs [lamI RS subst_elem] addSEs [lamE], 
-	       simpset()) 1);
-val lemma2 = result();
-
-Goal "[| \\<exists>f \\<in> A->C. P(f,domain(f)); A=B |] ==>  \\<exists>f \\<in> B->C. P(f,B)";
-by (etac bexE 1);
-by (ftac domain_of_fun 1);
-by (Fast_tac 1);
-val lemma3 = result();
-
-Goal "[| R \\<subseteq> A*B; g \\<in> C->R; \\<forall>x \\<in> C. (\\<lambda>z \\<in> R. fst(z))` (g`x) = x |] \
-\               ==> (\\<lambda>x \\<in> C. snd(g`x)): (\\<Pi>x \\<in> C. R``{x})";
-by (rtac lam_type 1);
-by (force_tac (claset() addDs [apply_type], simpset()) 1);
-val lemma4 = result();
-
-Goalw AC_defs "AC5 ==> AC4";
-by (Clarify_tac 1);
-by (REPEAT (eresolve_tac [allE,ballE] 1));
-by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
-by (dresolve_tac [lemma2 RSN (2, lemma3)] 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [lemma4]) 1);
-qed "AC5_AC4";
-
-
-(* ********************************************************************** *)
-(* AC1 <-> AC6                                                            *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC1 <-> AC6";
-by (Blast_tac 1);
-qed "AC1_iff_AC6";
--- a/src/ZF/AC/AC17_AC1.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC17_AC1.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,300 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/AC1_AC17.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+The equivalence of AC0, AC1 and AC17
+
+Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent
+to AC0 and AC1.
+*)
+
+theory AC17_AC1 = HH:
+
+
+(** AC0 is equivalent to AC1.  
+    AC0 comes from Suppes, AC1 from Rubin & Rubin **)
+
+lemma AC0_AC1_lemma: "[| f:(\<Pi>X \<in> A. X); D \<subseteq> A |] ==> \<exists>g. g:(\<Pi>X \<in> D. X)"
+by (fast intro!: restrict_type apply_type)
+
+lemma AC0_AC1: "AC0 ==> AC1"
+apply (unfold AC0_def AC1_def)
+apply (blast intro: AC0_AC1_lemma)
+done
+
+lemma AC1_AC0: "AC1 ==> AC0"
+by (unfold AC0_def AC1_def, blast)
+
+
+(**** The proof of AC1 ==> AC17 ****)
+
+lemma AC1_AC17_lemma: "f \<in> (\<Pi>X \<in> Pow(A) - {0}. X) ==> f \<in> (Pow(A) - {0} -> A)"
+apply (rule Pi_type, assumption)
+apply (drule apply_type, assumption, fast)
+done
+
+lemma AC1_AC17: "AC1 ==> AC17"
+apply (unfold AC1_def AC17_def)
+apply (rule allI)
+apply (rule ballI)
+apply (erule_tac x = "Pow (A) -{0}" in allE)
+apply (erule impE, fast)
+apply (erule exE)
+apply (rule bexI)
+apply (erule_tac [2] AC1_AC17_lemma)
+apply (rule apply_type, assumption)
+apply (fast dest!: AC1_AC17_lemma elim!: apply_type)
+done
+
+
+(**** The proof of AC17 ==> AC1 ****)
+
+(* *********************************************************************** *)
+(* more properties of HH                                                   *)
+(* *********************************************************************** *)
+
+lemma UN_eq_imp_well_ord:
+     "[| x - (\<Union>j \<in> LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x}.  
+        HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, j)) = 0;   
+        f \<in> Pow(x)-{0} -> x |]   
+        ==> \<exists>r. well_ord(x,r)"
+apply (rule exI)
+apply (erule well_ord_rvimage 
+        [OF bij_Least_HH_x [THEN bij_converse_bij, THEN bij_is_inj] 
+            Ord_Least [THEN well_ord_Memrel]], assumption)
+done
+
+(* *********************************************************************** *)
+(* theorems closer to the proof                                            *)
+(* *********************************************************************** *)
+
+lemma not_AC1_imp_ex:
+     "~AC1 ==> \<exists>A. \<forall>f \<in> Pow(A)-{0} -> A. \<exists>u \<in> Pow(A)-{0}. f`u \<notin> u"
+apply (unfold AC1_def)
+apply (erule swap)
+apply (rule allI)
+apply (erule swap)
+apply (rule_tac x = "Union (A)" in exI)
+apply (blast intro: restrict_type)
+done
+
+lemma lemma1:
+     "[| \<forall>f \<in> Pow(x) - {0} -> x. \<exists>u \<in> Pow(x) - {0}. f`u\<notin>u;   
+         \<exists>f \<in> Pow(x)-{0}->x.  
+            x - (\<Union>a \<in> (LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,i)={x}).   
+            HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,a)) = 0 |]  
+        ==> P"
+apply (erule bexE)
+apply (erule UN_eq_imp_well_ord [THEN exE], assumption)
+apply (erule ex_choice_fun_Pow [THEN exE])
+apply (erule ballE) 
+apply (fast intro: apply_type del: DiffE)
+apply (erule notE)
+apply (rule Pi_type, assumption)
+apply (blast dest: apply_type) 
+done
+
+lemma lemma2:
+      "~ (\<exists>f \<in> Pow(x)-{0}->x. x - F(f) = 0)   
+       ==> (\<lambda>f \<in> Pow(x)-{0}->x . x - F(f))   
+           \<in> (Pow(x) -{0} -> x) -> Pow(x) - {0}"
+by (fast intro!: lam_type dest!: Diff_eq_0_iff [THEN iffD1])
+
+lemma lemma3:
+     "[| f`Z \<in> Z; Z \<in> Pow(x)-{0} |] 
+      ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X})`Z \<in> Pow(Z)-{0}"
+by auto
+
+lemma lemma4:
+     "\<exists>f \<in> F. f`((\<lambda>f \<in> F. Q(f))`f) \<in> (\<lambda>f \<in> F. Q(f))`f   
+      ==> \<exists>f \<in> F. f`Q(f) \<in> Q(f)"
+by simp
+
+lemma AC17_AC1: "AC17 ==> AC1"
+apply (unfold AC17_def)
+apply (rule classical)
+apply (erule not_AC1_imp_ex [THEN exE])
+apply (case_tac 
+       "\<exists>f \<in> Pow(x)-{0} -> x. 
+        x - (\<Union>a \<in> (LEAST i. HH (\<lambda>X \<in> Pow (x) -{0}. {f`X},x,i) ={x}) . HH (\<lambda>X \<in> Pow (x) -{0}. {f`X},x,a)) = 0")
+apply (erule lemma1, assumption)
+apply (drule lemma2)
+apply (erule allE)
+apply (drule bspec, assumption)
+apply (drule lemma4)
+apply (erule bexE)
+apply (drule apply_type, assumption)
+apply (simp add: HH_Least_eq_x del: Diff_iff ) 
+apply (drule lemma3, assumption) 
+apply (fast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]]
+                   f_subset_imp_HH_subset elim!: mem_irrefl)
+done
+
+
+(* **********************************************************************
+    AC1 ==> AC2 ==> AC1
+    AC1 ==> AC4 ==> AC3 ==> AC1
+    AC4 ==> AC5 ==> AC4
+    AC1 <-> AC6
+************************************************************************* *)
+
+(* ********************************************************************** *)
+(* AC1 ==> AC2                                                            *)
+(* ********************************************************************** *)
+
+lemma lemma1:
+     "[| f:(\<Pi>X \<in> A. X);  B \<in> A;  0\<notin>A |] ==> {f`B} \<subseteq> B Int {f`C. C \<in> A}"
+by (fast elim!: apply_type)
 
-AC17_AC1 = HH
+lemma lemma2: 
+        "[| pairwise_disjoint(A); B \<in> A; C \<in> A; D \<in> B; D \<in> C |] ==> f`B = f`C"
+by (unfold pairwise_disjoint_def, fast)
+
+lemma AC1_AC2: "AC1 ==> AC2"
+apply (unfold AC1_def AC2_def)
+apply (rule allI)
+apply (rule impI)  
+apply (elim asm_rl conjE allE exE impE, assumption)
+apply (intro exI ballI equalityI)
+prefer 2 apply (rule lemma1, assumption+)
+apply (fast elim!: lemma2 elim: apply_type)
+done
+
+
+(* ********************************************************************** *)
+(* AC2 ==> AC1                                                            *)
+(* ********************************************************************** *)
+
+lemma lemma1: "0\<notin>A ==> 0 \<notin> {B*{B}. B \<in> A}"
+by (fast dest!: sym [THEN Sigma_empty_iff [THEN iffD1]])
+
+lemma lemma2: "[| X*{X} Int C = {y}; X \<in> A |]   
+               ==> (THE y. X*{X} Int C = {y}): X*A"
+apply (rule subst_elem [of y])
+apply (blast elim!: equalityE)
+apply (auto simp add: singleton_eq_iff) 
+done
+
+lemma lemma3:
+     "\<forall>D \<in> {E*{E}. E \<in> A}. \<exists>y. D Int C = {y}   
+      ==> (\<lambda>x \<in> A. fst(THE z. (x*{x} Int C = {z}))) \<in> (\<Pi>X \<in> A. X)"
+apply (rule lam_type)
+apply (drule bspec, blast)
+apply (blast intro: lemma2 fst_type)
+done
+
+lemma AC2_AC1: "AC2 ==> AC1"
+apply (unfold AC1_def AC2_def pairwise_disjoint_def)
+apply (intro allI impI)
+apply (elim allE impE)
+prefer 2 apply (fast elim!: lemma3) 
+apply (blast intro!: lemma1)
+done
+
+
+(* ********************************************************************** *)
+(* AC1 ==> AC4                                                            *)
+(* ********************************************************************** *)
+
+lemma empty_notin_images: "0 \<notin> {R``{x}. x \<in> domain(R)}"
+by blast
+
+lemma AC1_AC4: "AC1 ==> AC4"
+apply (unfold AC1_def AC4_def)
+apply (intro allI impI)
+apply (drule spec, drule mp [OF _ empty_notin_images]) 
+apply (best intro!: lam_type elim!: apply_type)
+done
+
+
+(* ********************************************************************** *)
+(* AC4 ==> AC3                                                            *)
+(* ********************************************************************** *)
+
+lemma lemma1: "f \<in> A->B ==> (\<Union>z \<in> A. {z}*f`z) \<subseteq> A*Union(B)"
+by (fast dest!: apply_type)
+
+lemma lemma2: "domain(\<Union>z \<in> A. {z}*f(z)) = {a \<in> A. f(a)\<noteq>0}"
+by blast
+
+lemma lemma3: "x \<in> A ==> (\<Union>z \<in> A. {z}*f(z))``{x} = f(x)"
+by fast
+
+lemma AC4_AC3: "AC4 ==> AC3"
+apply (unfold AC3_def AC4_def)
+apply (intro allI ballI)
+apply (elim allE impE)
+apply (erule lemma1)
+apply (simp add: lemma2 lemma3 cong add: Pi_cong)
+done
+
+(* ********************************************************************** *)
+(* AC3 ==> AC1                                                            *)
+(* ********************************************************************** *)
+
+lemma AC3_AC1_lemma:
+     "b\<notin>A ==> (\<Pi>x \<in> {a \<in> A. id(A)`a\<noteq>b}. id(A)`x) = (\<Pi>x \<in> A. x)"
+apply (simp add: id_def cong add: Pi_cong)
+apply (rule_tac b = "A" in subst_context, fast)
+done
+
+lemma AC3_AC1: "AC3 ==> AC1"
+apply (unfold AC1_def AC3_def)
+apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst])
+done
+
+(* ********************************************************************** *)
+(* AC4 ==> AC5                                                            *)
+(* ********************************************************************** *)
+
+lemma AC4_AC5: "AC4 ==> AC5"
+apply (unfold range_def AC4_def AC5_def)
+apply (intro allI ballI)
+apply (elim allE impE)
+apply (erule fun_is_rel [THEN converse_type])
+apply (erule exE)
+apply (rename_tac g)
+apply (rule_tac x=g in bexI)
+apply (blast dest: apply_equality range_type) 
+apply (blast intro: Pi_type dest: apply_type fun_is_rel)
+done
+
+
+(* ********************************************************************** *)
+(* AC5 ==> AC4, Rubin & Rubin, p. 11                                      *)
+(* ********************************************************************** *)
+
+lemma lemma1: "R \<subseteq> A*B ==> (\<lambda>x \<in> R. fst(x)) \<in> R -> A"
+by (fast intro!: lam_type fst_type)
+
+lemma lemma2: "R \<subseteq> A*B ==> range(\<lambda>x \<in> R. fst(x)) = domain(R)"
+by (unfold lam_def, force)
+
+lemma lemma3: "[| \<exists>f \<in> A->C. P(f,domain(f)); A=B |] ==>  \<exists>f \<in> B->C. P(f,B)"
+apply (erule bexE)
+apply (frule domain_of_fun, fast)
+done
+
+lemma lemma4: "[| R \<subseteq> A*B; g \<in> C->R; \<forall>x \<in> C. (\<lambda>z \<in> R. fst(z))` (g`x) = x |]  
+                ==> (\<lambda>x \<in> C. snd(g`x)): (\<Pi>x \<in> C. R``{x})"
+apply (rule lam_type)
+apply (force dest: apply_type)
+done
+
+lemma AC5_AC4: "AC5 ==> AC4"
+apply (unfold AC4_def AC5_def, clarify)
+apply (elim allE ballE)
+apply (drule lemma3 [OF _ lemma2], assumption)
+apply (fast elim!: lemma4)
+apply (blast intro: lemma1) 
+done
+
+
+(* ********************************************************************** *)
+(* AC1 <-> AC6                                                            *)
+(* ********************************************************************** *)
+
+lemma AC1_iff_AC6: "AC1 <-> AC6"
+by (unfold AC1_def AC6_def, blast)
+
+end
--- a/src/ZF/AC/AC18_AC19.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(*  Title:      ZF/AC/AC18_AC19.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The proof of AC1 ==> AC18 ==> AC19 ==> AC1
-*)
-
-(* ********************************************************************** *)
-(* AC1 ==> AC18                                                           *)
-(* ********************************************************************** *)
-
-Goal "[| f \\<in> (\\<Pi>b \\<in> {P(a). a \\<in> A}. b);  \\<forall>a \\<in> A. P(a)<=Q(a) |]  \
-\     ==> (\\<lambda>a \\<in> A. f`P(a)) \\<in> (\\<Pi>a \\<in> A. Q(a))";
-by (rtac lam_type 1);
-by (dtac apply_type 1);
-by Auto_tac;  
-qed "PROD_subsets";
-
-Goal "[| \\<forall>A. 0 \\<notin> A --> (\\<exists>f. f \\<in> (\\<Pi>X \\<in> A. X)); A \\<noteq> 0 |] ==>  \
-\  (\\<Inter>a \\<in> A. \\<Union>b \\<in> B(a). X(a, b)) \\<subseteq> (\\<Union>f \\<in> \\<Pi>a \\<in> A. B(a). \\<Inter>a \\<in> A. X(a, f`a))";
-by (rtac subsetI 1);
-by (eres_inst_tac [("x","{{b \\<in> B(a). x \\<in> X(a,b)}. a \\<in> A}")] allE 1);
-by (etac impE 1);
-by (Fast_tac 1);
-by (etac exE 1);
-by (rtac UN_I 1);
-by (fast_tac (claset() addSEs [PROD_subsets]) 1);
-by (Simp_tac 1);
-by (fast_tac (claset() addSEs [not_emptyE] 
-                       addDs [RepFunI RSN (2, apply_type)]) 1);
-qed "lemma_AC18";
-
-val [prem] = goalw thy (AC18_def::AC_defs) "AC1 ==> AC18";
-by (resolve_tac [prem RS revcut_rl] 1);
-by (fast_tac (claset() addSEs [lemma_AC18, not_emptyE, apply_type]
-                addSIs [equalityI, INT_I, UN_I]) 1);
-qed "AC1_AC18";
-
-(* ********************************************************************** *)
-(* AC18 ==> AC19                                                          *)
-(* ********************************************************************** *)
-
-val [prem] = goalw thy [AC18_def, AC19_def] "AC18 ==> AC19";
-by (rtac allI 1);
-by (res_inst_tac [("B1","%x. x")] (forall_elim_vars 0 prem RS revcut_rl) 1);
-by (Fast_tac 1);
-qed "AC18_AC19";
-
-(* ********************************************************************** *)
-(* AC19 ==> AC1                                                           *)
-(* ********************************************************************** *)
-
-Goalw [u_def]
-        "[| A \\<noteq> 0; 0 \\<notin> A |] ==> {u_(a). a \\<in> A} \\<noteq> 0 & 0 \\<notin> {u_(a). a \\<in> A}";
-by (fast_tac (claset() addSIs [not_emptyI]
-                addSEs [not_emptyE]
-                addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1);
-qed "RepRep_conj";
-
-Goal "[|c \\<in> a; x = c Un {0}; x \\<notin> a |] ==> x - {0} \\<in> a";
-by (hyp_subst_tac 1);
-by (rtac subst_elem 1 THEN (assume_tac 1));
-by (rtac equalityI 1);
-by (Fast_tac 1);
-by (rtac subsetI 1);
-by (excluded_middle_tac "x=0" 1);
-by (Fast_tac 1);
-by (fast_tac (claset() addEs [notE, subst_elem])  1);
-val lemma1_1 = result();
-
-Goalw [u_def]
-        "[| f`(u_(a)) \\<notin> a; f \\<in> (\\<Pi>B \\<in> {u_(a). a \\<in> A}. B); a \\<in> A |]  \
-\               ==> f`(u_(a))-{0} \\<in> a";
-by (fast_tac (claset() addSEs [lemma1_1] addSDs [apply_type]) 1);
-val lemma1_2 = result();
-
-Goal  "\\<exists>f. f \\<in> (\\<Pi>B \\<in> {u_(a). a \\<in> A}. B) ==> \\<exists>f. f \\<in> (\\<Pi>B \\<in> A. B)";
-by (etac exE 1);
-by (res_inst_tac
-        [("x","\\<lambda>a \\<in> A. if(f`(u_(a)) \\<in> a, f`(u_(a)), f`(u_(a))-{0})")] exI 1);
-by (rtac lam_type 1);
-by (split_tac [split_if] 1);
-by (rtac conjI 1);
-by (Fast_tac 1);
-by (fast_tac (claset() addSEs [lemma1_2]) 1);
-val lemma1 = result();
-
-Goalw [u_def] "a\\<noteq>0 ==> 0 \\<in> (\\<Union>b \\<in> u_(a). b)";
-by (fast_tac (claset() addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1);
-val lemma2_1 = result();
-
-Goal "[| A\\<noteq>0; 0\\<notin>A |] ==> (\\<Inter>x \\<in> {u_(a). a \\<in> A}. \\<Union>b \\<in> x. b) \\<noteq> 0";
-by (etac not_emptyE 1);
-by (res_inst_tac [("a","0")] not_emptyI 1);
-by (fast_tac (claset() addSIs [lemma2_1]) 1);
-val lemma2 = result();
-
-Goal "(\\<Union>f \\<in> F. P(f)) \\<noteq> 0 ==> F \\<noteq> 0";
-by (fast_tac (claset() addSEs [not_emptyE]) 1);
-val lemma3 = result();
-
-Goalw AC_defs "AC19 ==> AC1";
-by (Clarify_tac 1);
-by (case_tac "A=0" 1);
-by (Force_tac 1);
-by (eres_inst_tac [("x","{u_(a). a \\<in> A}")] allE 1);
-by (etac impE 1);
-by (etac RepRep_conj 1 THEN (assume_tac 1));
-by (rtac lemma1 1);
-by (dtac lemma2 1 THEN (assume_tac 1));
-by (dres_inst_tac [("P","%x. x\\<noteq>0")] subst 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [lemma3 RS not_emptyE]) 1);
-qed "AC19_AC1";
-
-
--- a/src/ZF/AC/AC18_AC19.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC18_AC19.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -2,17 +2,106 @@
     ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-Additional definition used in the proof AC19 ==> AC1 which is a part
-of the chain AC1 ==> AC18 ==> AC19 ==> AC1
+The proof of AC1 ==> AC18 ==> AC19 ==> AC1
 *)
 
-AC18_AC19 = AC_Equiv +
+theory AC18_AC19 = AC_Equiv:
+
+constdefs
+  uu    :: "i => i"
+    "uu(a) == {c Un {0}. c \<in> a}"
+
+
+(* ********************************************************************** *)
+(* AC1 ==> AC18                                                           *)
+(* ********************************************************************** *)
+
+lemma PROD_subsets:
+     "[| f \<in> (\<Pi>b \<in> {P(a). a \<in> A}. b);  \<forall>a \<in> A. P(a)<=Q(a) |]   
+      ==> (\<lambda>a \<in> A. f`P(a)) \<in> (\<Pi>a \<in> A. Q(a))"
+by (rule lam_type, drule apply_type, auto)
+
+lemma lemma_AC18:
+     "[| \<forall>A. 0 \<notin> A --> (\<exists>f. f \<in> (\<Pi>X \<in> A. X)); A \<noteq> 0 |] 
+      ==> (\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a, b)) \<subseteq> 
+          (\<Union>f \<in> \<Pi>a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a))"
+apply (rule subsetI)
+apply (erule_tac x = "{{b \<in> B (a) . x \<in> X (a,b) }. a \<in> A}" in allE)
+apply (erule impE, fast)
+apply (erule exE)
+apply (rule UN_I)
+ apply (fast elim!: PROD_subsets)
+apply (simp, fast elim!: not_emptyE dest: apply_type [OF _ RepFunI])
+done
+
+lemma AC1_AC18: "AC1 ==> AC18"
+apply (unfold AC1_def AC18_def)
+apply (fast elim!: lemma_AC18 apply_type intro!: equalityI INT_I UN_I)
+done
+
+(* ********************************************************************** *)
+(* AC18 ==> AC19                                                          *)
+(* ********************************************************************** *)
+
+text{*Hard to express because of the need for meta-quantifiers in AC18*}
+lemma "AC18 ==> AC19"
+proof -
+  assume ac18 [unfolded AC18_def, norm_hhf]: AC18
+  show AC19
+    apply (unfold AC18_def AC19_def)
+    apply (intro allI impI) 
+    apply (rule ac18 [of _ "%x. x", THEN mp], blast) 
+    done
+qed
 
-consts
-  u_    :: i => i
-  
-defs
+(* ********************************************************************** *)
+(* AC19 ==> AC1                                                           *)
+(* ********************************************************************** *)
+
+lemma RepRep_conj: 
+        "[| A \<noteq> 0; 0 \<notin> A |] ==> {uu(a). a \<in> A} \<noteq> 0 & 0 \<notin> {uu(a). a \<in> A}"
+apply (unfold uu_def, auto) 
+apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]])
+done
+
+lemma lemma1_1: "[|c \<in> a; x = c Un {0}; x \<notin> a |] ==> x - {0} \<in> a"
+apply clarify 
+apply (rule subst_elem , (assumption))
+apply (fast elim: notE subst_elem)
+done
+
+lemma lemma1_2: 
+        "[| f`(uu(a)) \<notin> a; f \<in> (\<Pi>B \<in> {uu(a). a \<in> A}. B); a \<in> A |]   
+                ==> f`(uu(a))-{0} \<in> a"
+apply (unfold uu_def, fast elim!: lemma1_1 dest!: apply_type)
+done
 
-  u_def "u_(a) == {c Un {0}. c \\<in> a}"
+lemma lemma1: "\<exists>f. f \<in> (\<Pi>B \<in> {uu(a). a \<in> A}. B) ==> \<exists>f. f \<in> (\<Pi>B \<in> A. B)"
+apply (erule exE)
+apply (rule_tac x = "\<lambda>a\<in>A. if (f` (uu(a)) \<in> a, f` (uu(a)), f` (uu(a))-{0})" 
+       in exI)
+apply (rule lam_type) 
+apply (simp add: lemma1_2)
+done
+
+lemma lemma2_1: "a\<noteq>0 ==> 0 \<in> (\<Union>b \<in> uu(a). b)"
+by (unfold uu_def, auto)
+
+lemma lemma2: "[| A\<noteq>0; 0\<notin>A |] ==> (\<Inter>x \<in> {uu(a). a \<in> A}. \<Union>b \<in> x. b) \<noteq> 0"
+apply (erule not_emptyE) 
+apply (rule_tac a = "0" in not_emptyI)
+apply (fast intro!: lemma2_1)
+done
+
+lemma AC19_AC1: "AC19 ==> AC1"
+apply (unfold AC19_def AC1_def, clarify)
+apply (case_tac "A=0", force)
+apply (erule_tac x = "{uu (a) . a \<in> A}" in allE)
+apply (erule impE)
+apply (erule RepRep_conj , (assumption))
+apply (rule lemma1)
+apply (drule lemma2 , (assumption))
+apply auto 
+done
 
 end
--- a/src/ZF/AC/AC1_WO2.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-(*  Title:      ZF/AC/AC1_WO2.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The proof of AC1 ==> WO2
-*)
-
-(*Establishing the existence of a bijection -- hence the need for uresult*)
-val [prem] = goal thy "f \\<in> (\\<Pi>X \\<in> Pow(x) - {0}. X) ==>  \
-\       ?g(f) \\<in> bij(x, LEAST i. HH(\\<lambda>X \\<in> Pow(x)-{0}. {f`X}, x, i) = {x})";
-by (resolve_tac [bij_Least_HH_x RS bij_converse_bij] 1);
-by (rtac f_subsets_imp_UN_HH_eq_x 1);
-by (resolve_tac [lam_type RS apply_type] 1 THEN (assume_tac 2));
-by (fast_tac (claset() addSDs [prem RS apply_type]) 1);
-by (fast_tac (claset() addSIs [prem RS Pi_weaken_type]) 1);
-val lemma1 = uresult() |> standard;
-
-Goalw [AC1_def, WO2_def, eqpoll_def] "AC1 ==> WO2";
-by (rtac allI 1);
-by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
-by (fast_tac (claset() addSDs [lemma1] addSIs [Ord_Least]) 1);
-qed "AC1_WO2";
--- a/src/ZF/AC/AC1_WO2.thy	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-(*Dummy theory to document dependencies *)
-
-AC1_WO2 = HH
--- a/src/ZF/AC/AC7_AC9.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,188 +0,0 @@
-(*  Title:      ZF/AC/AC7-AC9.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous
-instances of AC.
-*)
-
-(* ********************************************************************** *)
-(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1                  *)
-(*  - Sigma_fun_space_not0                                                *)
-(*  - Sigma_fun_space_eqpoll                                              *)
-(* ********************************************************************** *)
-
-Goal "[| 0\\<notin>A; B \\<in> A |] ==> (nat->Union(A)) * B \\<noteq> 0";
-by (blast_tac (claset() addSDs [Sigma_empty_iff RS iffD1, 
-				Union_empty_iff RS iffD1]) 1);
-qed "Sigma_fun_space_not0";
-
-Goalw [inj_def]
-        "C \\<in> A ==> (\\<lambda>g \\<in> (nat->Union(A))*C.  \
-\               (\\<lambda>n \\<in> nat. if(n=0, snd(g), fst(g)`(n #- 1))))  \
-\               \\<in> inj((nat->Union(A))*C, (nat->Union(A)) ) ";
-by (rtac CollectI 1);
-by (fast_tac (claset() addSIs [lam_type,RepFunI,if_type,snd_type,apply_type,
-                                fst_type,diff_type,nat_succI,nat_0I]) 1);
-by (REPEAT (resolve_tac [ballI, impI] 1));
-by (Asm_full_simp_tac 1);
-by (REPEAT (etac SigmaE 1));
-by (REPEAT (hyp_subst_tac 1));
-by (Asm_full_simp_tac 1);
-by (rtac conjI 1);
-by (dresolve_tac [nat_0I RSN (2, lam_eqE)] 2);
-by (Asm_full_simp_tac 2);
-by (rtac fun_extension 1 THEN  REPEAT (assume_tac 1));
-by (dresolve_tac [nat_succI RSN (2, lam_eqE)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [succ_not_0 RS if_not_P]) 1);
-val lemma = result();
-
-Goal "[| C \\<in> A; 0\\<notin>A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
-by (rtac eqpollI 1);
-by (fast_tac (claset() addSEs [prod_lepoll_self, not_sym RS not_emptyE,
-                subst_elem] addEs [swap]) 2);
-by (rewtac lepoll_def);
-by (fast_tac (claset() addSIs [lemma]) 1);
-qed "Sigma_fun_space_eqpoll";
-
-
-(* ********************************************************************** *)
-(* AC6 ==> AC7                                                            *)
-(* ********************************************************************** *)
-
-Goalw AC_defs "AC6 ==> AC7";
-by (Blast_tac 1);
-qed "AC6_AC7";
-
-(* ********************************************************************** *)
-(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8                          *)
-(* The case of the empty family of sets added in order to complete        *)
-(* the proof.                                                             *)
-(* ********************************************************************** *)
-
-Goal "y \\<in> (\\<Pi>B \\<in> A. Y*B) ==> (\\<lambda>B \\<in> A. snd(y`B)): (\\<Pi>B \\<in> A. B)";
-by (fast_tac (claset() addSIs [lam_type, snd_type, apply_type]) 1);
-val lemma1_1 = result();
-
-Goal "y \\<in> (\\<Pi>B \\<in> {Y*C. C \\<in> A}. B) ==> (\\<lambda>B \\<in> A. y`(Y*B)): (\\<Pi>B \\<in> A. Y*B)";
-by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
-val lemma1_2 = result();
-
-Goal "(\\<Pi>B \\<in> {(nat->Union(A))*C. C \\<in> A}. B) \\<noteq> 0 ==> (\\<Pi>B \\<in> A. B) \\<noteq> 0";
-by (fast_tac (claset() addSIs [equals0I,lemma1_1, lemma1_2]) 1);
-val lemma1 = result();
-
-Goal "0 \\<notin> A ==> 0 \\<notin> {(nat -> Union(A)) * C. C \\<in> A}";
-by (fast_tac (claset() addEs [Sigma_fun_space_not0 RS not_sym RS notE]) 1);
-val lemma2 = result();
-
-Goalw AC_defs "AC7 ==> AC6";
-by (rtac allI 1);
-by (rtac impI 1);
-by (case_tac "A=0" 1);
-by (Asm_simp_tac 1);
-by (rtac lemma1 1);
-by (etac allE 1);
-by (etac impE 1 THEN (assume_tac 2));
-by (blast_tac (claset() addSIs [lemma2] 
-                addIs [eqpoll_sym, eqpoll_trans, Sigma_fun_space_eqpoll]) 1); 
-qed "AC7_AC6";
-
-
-(* ********************************************************************** *)
-(* AC1 ==> AC8                                                            *)
-(* ********************************************************************** *)
-
-Goalw [eqpoll_def]
-        "\\<forall>B \\<in> A. \\<exists>B1 B2. B=<B1,B2> & B1 eqpoll B2  \
-\       ==> 0 \\<notin> { bij(fst(B),snd(B)). B \\<in> A }";
-by Auto_tac;
-val lemma1 = result();
-
-Goal "[| f \\<in> (\\<Pi>X \\<in> RepFun(A,p). X); D \\<in> A |] ==> (\\<lambda>x \\<in> A. f`p(x))`D \\<in> p(D)";
-by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [apply_type]) 1);
-val lemma2 = result();
-
-Goalw AC_defs "AC1 ==> AC8";
-by (Clarify_tac 1);
-by (dtac lemma1 1);
-by (fast_tac (claset() addSEs [lemma2]) 1);
-qed "AC1_AC8";
-
-
-(* ********************************************************************** *)
-(* AC8 ==> AC9                                                            *)
-(*  - this proof replaces the following two from Rubin & Rubin:           *)
-(*    AC8 ==> AC1 and AC1 ==> AC9                                         *)
-(* ********************************************************************** *)
-
-Goal "\\<forall>B1 \\<in> A. \\<forall>B2 \\<in> A. B1 eqpoll B2  \
-\     ==> \\<forall>B \\<in> A*A. \\<exists>B1 B2. B=<B1,B2> & B1 eqpoll B2";
-by (Fast_tac 1);
-val lemma1 = result();
-
-Goal "f \\<in> bij(fst(<a,b>),snd(<a,b>)) ==> f \\<in> bij(a,b)";
-by (Asm_full_simp_tac 1);
-val lemma2 = result();
-
-Goalw AC_defs "AC8 ==> AC9";
-by (rtac allI 1);
-by (rtac impI 1);
-by (etac allE 1);
-by (etac impE 1);
-by (etac lemma1 1);
-by (fast_tac (claset() addSEs [lemma2]) 1);
-qed "AC8_AC9";
-
-
-(* ********************************************************************** *)
-(* AC9 ==> AC1                                                            *)
-(* The idea of this proof comes from "Equivalents of the Axiom of Choice" *)
-(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to        *)
-(* (x * y) Un {0} when y is a set of total functions acting from nat to   *)
-(* Union(A) -- therefore we have used the set (y * nat) instead of y.     *)
-(* ********************************************************************** *)
-
-(* Rules nedded to prove lemma1 *)
-val snd_lepoll_SigmaI = prod_lepoll_self RS 
-        ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans));
-
-
-Goal "[|0 \\<notin> A; B \\<in> A|] ==> nat \\<lesssim> ((nat \\<rightarrow> Union(A)) \\<times> B) \\<times> nat";
-by (blast_tac (claset() addDs [Sigma_fun_space_not0]
-                        addIs [snd_lepoll_SigmaI]) 1);
-qed "nat_lepoll_lemma";
-
-
-Goal "[| 0\\<notin>A;  A\\<noteq>0;  \
-\        C = {((nat->Union(A))*B)*nat. B \\<in> A}  Un \
-\            {cons(0,((nat->Union(A))*B)*nat). B \\<in> A}; \
-\        B1: C;  B2: C |]  \
-\     ==> B1 eqpoll B2";
-by (blast_tac
-    (claset() delrules [eqpoll_refl]
-	      addSIs [nat_lepoll_lemma, nat_cons_eqpoll RS eqpoll_trans, 
-                      eqpoll_refl RSN (2, prod_eqpoll_cong)]
-              addIs [eqpoll_trans, eqpoll_sym, Sigma_fun_space_eqpoll]) 1);
-val lemma1 = result();
-
-Goal "\\<forall>B1 \\<in> {(F*B)*N. B \\<in> A} Un {cons(0,(F*B)*N). B \\<in> A}.  \
-\     \\<forall>B2 \\<in> {(F*B)*N. B \\<in> A} Un {cons(0,(F*B)*N). B \\<in> A}.  \
-\       f`<B1,B2> \\<in> bij(B1, B2)  \
-\   ==> (\\<lambda>B \\<in> A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) \\<in> (\\<Pi>X \\<in> A. X)";
-by (rtac lam_type 1);
-by (rtac snd_type 1);
-by (rtac fst_type 1);
-by (resolve_tac [consI1 RSN (2, apply_type)] 1);
-by (fast_tac (claset() addSIs [fun_weaken_type, bij_is_fun]) 1);
-val lemma2 = result();
-
-Goalw AC_defs "AC9 ==> AC1";
-by (rtac allI 1);
-by (rtac impI 1);
-by (etac allE 1);
-by (case_tac "A=0" 1);
-by (blast_tac (claset() addDs [lemma1,lemma2]) 2); 
-by Auto_tac;  
-qed "AC9_AC1";
--- a/src/ZF/AC/AC_Equiv.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,126 +0,0 @@
-(*  Title:      ZF/AC/AC_Equiv.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-*)
-
-val WO_defs = [WO1_def, WO2_def, WO3_def, WO4_def, WO5_def, WO6_def, WO8_def];
- 
-val AC_defs = [AC0_def, AC1_def, AC2_def, AC3_def, AC4_def, AC5_def, 
-               AC6_def, AC7_def, AC8_def, AC9_def, AC10_def, AC11_def, 
-               AC12_def, AC13_def, AC14_def, AC15_def, AC16_def, 
-               AC17_def, AC18_def, AC19_def];
- 
-val AC_aux_defs = [pairwise_disjoint_def, sets_of_size_between_def];
- 
-
-(* ********************************************************************** *)
-(*             lemmas concerning FOL and pure ZF theory                   *)
-(* ********************************************************************** *)
-
-(* used only in WO1_DC.ML *)
-(*Note simpler proof*)
-Goal "[| \\<forall>x \\<in> A. f`x=g`x; f \\<in> Df->Cf; g \\<in> Dg->Cg; A \\<subseteq> Df; A \\<subseteq> Dg |] ==> f``A=g``A";
-by (asm_simp_tac (simpset() addsimps [image_fun]) 1);
-qed "images_eq";
-
-(* used in \\<in> AC10-AC15.ML AC16WO4.ML WO6WO1.ML *)
-(*I don't know where to put this one.*)
-Goal
-     "!!m A B. [| A lepoll succ(m); B \\<subseteq> A; B\\<noteq>0 |] ==> A-B lepoll m";
-by (rtac not_emptyE 1 THEN (assume_tac 1));
-by (ftac singleton_subsetI 1);
-by (ftac subsetD 1 THEN (assume_tac 1));
-by (res_inst_tac [("A2","A")] 
-     (Diff_sing_lepoll RSN (2, subset_imp_lepoll RS lepoll_trans)) 1 
-    THEN (REPEAT (assume_tac 2)));
-by (Fast_tac 1);
-qed "Diff_lepoll";
-
-(* ********************************************************************** *)
-(*              lemmas concerning lepoll and eqpoll relations             *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(*                    Theorems concerning ordinals                        *)
-(* ********************************************************************** *)
-
-(* lemma for ordertype_Int *)
-goalw Cardinal.thy [rvimage_def] "rvimage(A,id(A),r) = r Int A*A";
-by (rtac equalityI 1);
-by Safe_tac;
-by (dres_inst_tac [("P","%a. <id(A)`xb,a>:r")] (id_conv RS subst) 1
-    THEN (assume_tac 1));
-by (dres_inst_tac [("P","%a. <a,ya>:r")] (id_conv RS subst) 1
-    THEN (REPEAT (assume_tac 1)));
-by (fast_tac (claset() addIs [id_conv RS ssubst]) 1);
-qed "rvimage_id";
-
-(* used only in Hartog.ML *)
-goal Cardinal.thy
-        "!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)";
-by (res_inst_tac [("P","%a. ordertype(A,a)=ordertype(A,r)")] 
-    (rvimage_id RS subst) 1);
-by (eresolve_tac [id_bij RS bij_ordertype_vimage] 1);
-qed "ordertype_Int";
-
-(* used only in AC16_lemmas.ML *)
-Goalw [InfCard_def]
-        "!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)";
-by (asm_simp_tac (simpset() addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
-qed "Inf_Card_is_InfCard";
-
-Goal "(THE z. {x}={z}) = x";
-by (fast_tac (claset() addSEs [singleton_eq_iff RS iffD1 RS sym]) 1);
-qed "the_element";
-
-Goal "(\\<lambda>x \\<in> A. {x}) \\<in> bij(A, {{x}. x \\<in> A})";
-by (res_inst_tac [("d","%z. THE x. z={x}")] lam_bijective 1);
-by (TRYALL (eresolve_tac [RepFunI, RepFunE]));
-by (REPEAT (asm_full_simp_tac (simpset() addsimps [the_element]) 1));
-qed "lam_sing_bij";
-
-val [major, minor] = Goalw [inj_def]
-        "[| f \\<in> inj(A, B);  !!a. a \\<in> A ==> f`a \\<in> C |] ==> f \\<in> inj(A,C)";
-by (fast_tac (claset() addSEs [minor]
-        addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1);
-qed "inj_strengthen_type";
-
-Goalw [Finite_def] "~Finite(nat)";
-by (fast_tac (claset() addSDs [eqpoll_imp_lepoll]
-                addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1);
-qed "nat_not_Finite";
-
-val le_imp_lepoll = le_imp_subset RS subset_imp_lepoll;
-
-(* ********************************************************************** *)
-(* Another elimination rule for \\<exists>!                                       *)
-(* ********************************************************************** *)
-
-Goal "[| \\<exists>! x. P(x); P(x); P(y) |] ==> x=y";
-by (etac ex1E 1);
-by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1);
-by (Fast_tac 1);
-by (Fast_tac 1);
-qed "ex1_two_eq";
-
-(* ********************************************************************** *)
-(* image of a surjection                                                  *)
-(* ********************************************************************** *)
-
-Goalw [surj_def] "f \\<in> surj(A, B) ==> f``A = B";
-by (etac CollectE 1);
-by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 
-    THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [apply_type] addIs [equalityI]) 1);
-qed "surj_image_eq";
-
-
-Goal "succ(x) lepoll y ==> y \\<noteq> 0";
-by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
-qed "succ_lepoll_imp_not_empty";
-
-Goal "x eqpoll succ(n) ==> x \\<noteq> 0";
-by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
-qed "eqpoll_succ_imp_not_empty";
-
--- a/src/ZF/AC/AC_Equiv.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/AC_Equiv.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -12,114 +12,231 @@
 but slightly changed.
 *)
 
+theory AC_Equiv = Main: (*obviously not Main_ZFC*)
 
-AC_Equiv = Main + (*obviously not Main_ZFC*)
-
-consts
+constdefs
   
 (* Well Ordering Theorems *)
-  WO1, WO2, WO3, WO5, WO6, WO7, WO8 :: o
-  WO4                               :: i => o
+  WO1 :: o
+    "WO1 == \<forall>A. \<exists>R. well_ord(A,R)"
+
+  WO2 :: o
+    "WO2 == \<forall>A. \<exists>a. Ord(a) & A\<approx>a"
 
-(* Axioms of Choice *)  
-  AC0, AC1, AC2, AC3, AC4, AC5, AC6, AC7, AC8, AC9,
-  AC11, AC12, AC14, AC15, AC17, AC19 :: o
-  AC10, AC13              :: i => o
-  AC16                    :: [i, i] => o
-  AC18                    :: prop       ("AC18")
+  WO3 :: o
+    "WO3 == \<forall>A. \<exists>a. Ord(a) & (\<exists>b. b \<subseteq> a & A\<approx>b)"
 
-(* Auxiliary definitions used in definitions *)
-  pairwise_disjoint       :: i => o
-  sets_of_size_between    :: [i, i, i] => o
+  WO4 :: "i => o"
+    "WO4(m) == \<forall>A. \<exists>a f. Ord(a) & domain(f)=a &   
+		         (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m)"
 
-defs
-
-(* Well Ordering Theorems *)
+  WO5 :: o
+    "WO5 == \<exists>m \<in> nat. 1\<le>m & WO4(m)"
 
-  WO1_def "WO1 == \\<forall>A. \\<exists>R. well_ord(A,R)"
-
-  WO2_def "WO2 == \\<forall>A. \\<exists>a. Ord(a) & A eqpoll a"
+  WO6 :: o
+    "WO6 == \<forall>A. \<exists>m \<in> nat. 1\<le>m & (\<exists>a f. Ord(a) & domain(f)=a
+		               & (\<Union>b<a. f`b) = A & (\<forall>b<a. f`b \<lesssim> m))"
 
-  WO3_def "WO3 == \\<forall>A. \\<exists>a. Ord(a) & (\\<exists>b. b \\<subseteq> a & A eqpoll b)"
-
-  WO4_def "WO4(m) == \\<forall>A. \\<exists>a f. Ord(a) & domain(f)=a &   
-                     (\\<Union>b<a. f`b) = A & (\\<forall>b<a. f`b lepoll m)"
+  WO7 :: o
+    "WO7 == \<forall>A. Finite(A) <-> (\<forall>R. well_ord(A,R) --> well_ord(A,converse(R)))"
 
-  WO5_def "WO5 == \\<exists>m \\<in> nat. 1 le m & WO4(m)"
+  WO8 :: o
+    "WO8 == \<forall>A. (\<exists>f. f \<in> (\<Pi>X \<in> A. X)) --> (\<exists>R. well_ord(A,R))"
 
-  WO6_def "WO6 == \\<forall>A. \\<exists>m \\<in> nat. 1 le m & (\\<exists>a f. Ord(a) & domain(f)=a   
-                    & (\\<Union>b<a. f`b) = A & (\\<forall>b<a. f`b lepoll m))"
 
-  WO7_def "WO7 == \\<forall>A. Finite(A) <-> (\\<forall>R. well_ord(A,R) -->   
-                    well_ord(A,converse(R)))"
+(* Auxiliary concepts needed below *)
+  pairwise_disjoint :: "i => o"
+    "pairwise_disjoint(A) == \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 Int A2 \<noteq> 0 --> A1=A2"
 
-  WO8_def "WO8 == \\<forall>A. (\\<exists>f. f \\<in> (\\<Pi>X \\<in> A. X)) --> (\\<exists>R. well_ord(A,R))"
+  sets_of_size_between :: "[i, i, i] => o"
+    "sets_of_size_between(A,m,n) == \<forall>B \<in> A. m \<lesssim> B & B \<lesssim> n"
+
 
 (* Axioms of Choice *)  
+  AC0 :: o
+    "AC0 == \<forall>A. \<exists>f. f \<in> (\<Pi>X \<in> Pow(A)-{0}. X)"
 
-  AC0_def "AC0 == \\<forall>A. \\<exists>f. f \\<in> (\\<Pi>X \\<in> Pow(A)-{0}. X)"
+  AC1 :: o
+    "AC1 == \<forall>A. 0\<notin>A --> (\<exists>f. f \<in> (\<Pi>X \<in> A. X))"
+
+  AC2 :: o
+    "AC2 == \<forall>A. 0\<notin>A & pairwise_disjoint(A)   
+		   --> (\<exists>C. \<forall>B \<in> A. \<exists>y. B Int C = {y})"
+  AC3 :: o
+    "AC3 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g. g \<in> (\<Pi>x \<in> {a \<in> A. f`a\<noteq>0}. f`x)"
 
-  AC1_def "AC1 == \\<forall>A. 0\\<notin>A --> (\\<exists>f. f \\<in> (\\<Pi>X \\<in> A. X))"
+  AC4 :: o
+    "AC4 == \<forall>R A B. (R \<subseteq> A*B --> (\<exists>f. f \<in> (\<Pi>x \<in> domain(R). R``{x})))"
+
+  AC5 :: o
+    "AC5 == \<forall>A B. \<forall>f \<in> A->B. \<exists>g \<in> range(f)->A. \<forall>x \<in> domain(g). f`(g`x) = x"
+
+  AC6 :: o
+    "AC6 == \<forall>A. 0\<notin>A --> (\<Pi>B \<in> A. B)\<noteq>0"
 
-  AC2_def "AC2 == \\<forall>A. 0\\<notin>A & pairwise_disjoint(A)   
-                    --> (\\<exists>C. \\<forall>B \\<in> A. \\<exists>y. B Int C = {y})"
+  AC7 :: o
+    "AC7 == \<forall>A. 0\<notin>A & (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) --> (\<Pi>B \<in> A. B) \<noteq> 0"
 
-  AC3_def "AC3 == \\<forall>A B. \\<forall>f \\<in> A->B. \\<exists>g. g \\<in> (\\<Pi>x \\<in> {a \\<in> A. f`a\\<noteq>0}. f`x)"
+  AC8 :: o
+    "AC8 == \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1\<approx>B2)   
+		   --> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
+
+  AC9 :: o
+    "AC9 == \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) -->   
+		   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
 
-  AC4_def "AC4 == \\<forall>R A B. (R \\<subseteq> A*B --> (\\<exists>f. f \\<in> (\\<Pi>x \\<in> domain(R). R``{x})))"
+  AC10 :: "i => o"
+    "AC10(n) ==  \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->   
+		   (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
+		   sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))"
 
-  AC5_def "AC5 == \\<forall>A B. \\<forall>f \\<in> A->B. \\<exists>g \\<in> range(f)->A.   
-                    \\<forall>x \\<in> domain(g). f`(g`x) = x"
+  AC11 :: o
+    "AC11 == \<exists>n \<in> nat. 1\<le>n & AC10(n)"
+
+  AC12 :: o
+    "AC12 == \<forall>A. (\<forall>B \<in> A. ~Finite(B)) -->
+  	         (\<exists>n \<in> nat. 1\<le>n & (\<exists>f. \<forall>B \<in> A. (pairwise_disjoint(f`B) &   
+	              sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))"
 
-  AC6_def "AC6 == \\<forall>A. 0\\<notin>A --> (\\<Pi>B \\<in> A. B)\\<noteq>0"
+  AC13 :: "i => o"
+    "AC13(m) == \<forall>A. 0\<notin>A --> (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m)"
+
+  AC14 :: o
+    "AC14 == \<exists>m \<in> nat. 1\<le>m & AC13(m)"
+
+  AC15 :: o
+    "AC15 == \<forall>A. 0\<notin>A --> 
+                 (\<exists>m \<in> nat. 1\<le>m & (\<exists>f. \<forall>B \<in> A. f`B\<noteq>0 & f`B \<subseteq> B & f`B \<lesssim> m))"
 
-  AC7_def "AC7 == \\<forall>A. 0\\<notin>A & (\\<forall>B1 \\<in> A. \\<forall>B2 \\<in> A. B1 eqpoll B2)   
-                    --> (\\<Pi>B \\<in> A. B)\\<noteq>0"
+  AC16 :: "[i, i] => o"
+    "AC16(n, k)  == 
+       \<forall>A. ~Finite(A) -->   
+	   (\<exists>T. T \<subseteq> {X \<in> Pow(A). X\<approx>succ(n)} &   
+	   (\<forall>X \<in> {X \<in> Pow(A). X\<approx>succ(k)}. \<exists>! Y. Y \<in> T & X \<subseteq> Y))"
 
-  AC8_def "AC8 == \\<forall>A. (\\<forall>B \\<in> A. \\<exists>B1 B2. B=<B1,B2> & B1 eqpoll B2)   
-                    --> (\\<exists>f. \\<forall>B \\<in> A. f`B \\<in> bij(fst(B),snd(B)))"
+  AC17 :: o
+    "AC17 == \<forall>A. \<forall>g \<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
+		   \<exists>f \<in> Pow(A)-{0} -> A. f`(g`f) \<in> g`f"
 
-  AC9_def "AC9 == \\<forall>A. (\\<forall>B1 \\<in> A. \\<forall>B2 \\<in> A. B1 eqpoll B2) -->   
-                    (\\<exists>f. \\<forall>B1 \\<in> A. \\<forall>B2 \\<in> A. f`<B1,B2> \\<in> bij(B1,B2))"
+  AC18 :: "prop" ("AC18")
+    "AC18 == (!!X A B. A\<noteq>0 & (\<forall>a \<in> A. B(a) \<noteq> 0) -->   
+		((\<Inter>a \<in> A. \<Union>b \<in> B(a). X(a,b)) =   
+		(\<Union>f \<in> \<Pi>a \<in> A. B(a). \<Inter>a \<in> A. X(a, f`a))))"
+  --"AC18 can be expressed only using meta-level quantification"
+
+  AC19 :: o
+    "AC19 == \<forall>A. A\<noteq>0 & 0\<notin>A --> ((\<Inter>a \<in> A. \<Union>b \<in> a. b) =   
+		   (\<Union>f \<in> (\<Pi>B \<in> A. B). \<Inter>a \<in> A. f`a))"
+
+
 
-  AC10_def "AC10(n) ==  \\<forall>A. (\\<forall>B \\<in> A. ~Finite(B)) -->   
-                    (\\<exists>f. \\<forall>B \\<in> A. (pairwise_disjoint(f`B) &   
-                    sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B))"
+(* ********************************************************************** *)
+(*                    Theorems concerning ordinals                        *)
+(* ********************************************************************** *)
 
-  AC11_def "AC11 == \\<exists>n \\<in> nat. 1 le n & AC10(n)"
+(* lemma for ordertype_Int *)
+lemma rvimage_id: "rvimage(A,id(A),r) = r Int A*A"
+apply (unfold rvimage_def)
+apply (rule equalityI, safe)
+apply (drule_tac P = "%a. <id (A) `xb,a>:r" in id_conv [THEN subst],
+       (assumption))
+apply (drule_tac P = "%a. <a,ya>:r" in id_conv [THEN subst], (assumption+))
+apply (fast intro: id_conv [THEN ssubst])
+done
 
-  AC12_def "AC12 == \\<forall>A. (\\<forall>B \\<in> A. ~Finite(B)) -->   
-            (\\<exists>n \\<in> nat. 1 le n & (\\<exists>f. \\<forall>B \\<in> A. (pairwise_disjoint(f`B) &   
-            sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B)))"
+(* used only in Hartog.ML *)
+lemma ordertype_Int:
+     "well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)"
+apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst])
+apply (erule id_bij [THEN bij_ordertype_vimage])
+done
+
+lemma the_element: "(THE z. {x}={z}) = x"
+by (fast elim!: singleton_eq_iff [THEN iffD1, symmetric])
 
-  AC13_def "AC13(m) == \\<forall>A. 0\\<notin>A --> (\\<exists>f. \\<forall>B \\<in> A. f`B\\<noteq>0 &   
-                                          f`B \\<subseteq> B & f`B lepoll m)"
+lemma lam_sing_bij: "(\<lambda>x \<in> A. {x}) \<in> bij(A, {{x}. x \<in> A})"
+apply (rule_tac d = "%z. THE x. z={x}" in lam_bijective)
+apply (auto simp add: the_element) 
+done
+
+lemma inj_strengthen_type: 
+     "[| f \<in> inj(A, B);  !!a. a \<in> A ==> f`a \<in> C |] ==> f \<in> inj(A,C)"
+by (unfold inj_def, blast intro: Pi_type) 
+
+lemma nat_not_Finite: "~ Finite(nat)"
+by (unfold Finite_def, blast dest: eqpoll_imp_lepoll ltI lt_not_lepoll)
 
-  AC14_def "AC14 == \\<exists>m \\<in> nat. 1 le m & AC13(m)"
+lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll]
+
+(* ********************************************************************** *)
+(* Another elimination rule for \<exists>!                                       *)
+(* ********************************************************************** *)
+
+lemma ex1_two_eq: "[| \<exists>! x. P(x); P(x); P(y) |] ==> x=y"
+by blast
 
-  AC15_def "AC15 == \\<forall>A. 0\\<notin>A --> (\\<exists>m \\<in> nat. 1 le m & (\\<exists>f. \\<forall>B \\<in> A.   
-                                      f`B\\<noteq>0 & f`B \\<subseteq> B & f`B lepoll m))"
+(* ********************************************************************** *)
+(* image of a surjection                                                  *)
+(* ********************************************************************** *)
 
-  AC16_def "AC16(n, k)  == \\<forall>A. ~Finite(A) -->   
-            (\\<exists>T. T \\<subseteq> {X \\<in> Pow(A). X eqpoll succ(n)} &   
-            (\\<forall>X \\<in> {X \\<in> Pow(A). X eqpoll succ(k)}. \\<exists>! Y. Y \\<in> T & X \\<subseteq> Y))"
+lemma surj_image_eq: "f \<in> surj(A, B) ==> f``A = B"
+apply (unfold surj_def)
+apply (erule CollectE)
+apply (rule image_fun [THEN ssubst], (assumption), rule subset_refl)
+apply (blast dest: apply_type) 
+done
+
+
+(* ********************************************************************** *)
+(* Lemmas used in the proofs like WO? ==> AC?                             *)
+(* ********************************************************************** *)
 
-  AC17_def "AC17 == \\<forall>A. \\<forall>g \\<in> (Pow(A)-{0} -> A) -> Pow(A)-{0}.   
-                    \\<exists>f \\<in> Pow(A)-{0} -> A. f`(g`f) \\<in> g`f"
+lemma first_in_B:
+     "[| well_ord(Union(A),r); 0\<notin>A; B \<in> A |] ==> (THE b. first(b,B,r)) \<in> B"
+by (blast dest!: well_ord_imp_ex1_first
+                    [THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]])
+
+lemma ex_choice_fun: "[| well_ord(Union(A), R); 0\<notin>A |] ==> \<exists>f. f:(\<Pi>X \<in> A. X)"
+by (fast elim!: first_in_B intro!: lam_type)
 
-  AC18_def "AC18 == (!!X A B. A\\<noteq>0 & (\\<forall>a \\<in> A. B(a) \\<noteq> 0) -->   
-                 ((\\<Inter>a \\<in> A. \\<Union>b \\<in> B(a). X(a,b)) =   
-                 (\\<Union>f \\<in> \\<Pi>a \\<in> A. B(a). \\<Inter>a \\<in> A. X(a, f`a))))"
+lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \<exists>f. f:(\<Pi>X \<in> Pow(A)-{0}. X)"
+by (fast elim!: well_ord_subset [THEN ex_choice_fun])
+
 
-  AC19_def "AC19 == \\<forall>A. A\\<noteq>0 & 0\\<notin>A --> ((\\<Inter>a \\<in> A. \\<Union>b \\<in> a. b) =   
-                    (\\<Union>f \\<in> (\\<Pi>B \\<in> A. B). \\<Inter>a \\<in> A. f`a))"
+(* ********************************************************************** *)
+(* Lemmas needed to state when a finite relation is a function.           *)
+(*     The criteria are cardinalities of the relation and its domain.     *)
+(*     Used in WO6WO1.ML                                                  *)
+(* ********************************************************************** *)
 
-(* Auxiliary definitions used in the above definitions *)
-
-  pairwise_disjoint_def    "pairwise_disjoint(A)   
-                            == \\<forall>A1 \\<in> A. \\<forall>A2 \\<in> A. A1 Int A2 \\<noteq> 0 --> A1=A2"
+(*Using AC we could trivially prove, for all u, domain(u) \<lesssim> u*)
+lemma lepoll_m_imp_domain_lepoll_m: 
+     "[| m \<in> nat; u \<lesssim> m |] ==> domain(u) \<lesssim> m"
+apply (unfold lepoll_def)
+apply (erule exE)
+apply (rule_tac x = "\<lambda>x \<in> domain(u). LEAST i. \<exists>y. <x,y> \<in> u & f`<x,y> = i" 
+       in exI)
+apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective)
+apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] 
+                           inj_is_fun [THEN apply_type])
+apply (erule domainE)
+apply (frule inj_is_fun [THEN apply_type], (assumption))
+apply (rule LeastI2)
+apply (auto elim!: nat_into_Ord [THEN Ord_in_Ord])
+done
 
-  sets_of_size_between_def "sets_of_size_between(A,m,n)   
-                            == \\<forall>B \\<in> A. m lepoll B & B lepoll n"
-  
+lemma rel_domain_ex1: 
+    "[| succ(m) \<lesssim> domain(r); r \<lesssim> succ(m); m \<in> nat |] ==> function(r)"
+apply (unfold function_def, safe)
+apply (rule ccontr) 
+apply (fast elim!: lepoll_trans [THEN succ_lepoll_natE] 
+                   lepoll_m_imp_domain_lepoll_m [OF _ Diff_sing_lepoll]
+            elim: domain_Diff_eq [OF _ not_sym, THEN subst])
+done
+
+lemma rel_is_fun:
+     "[| succ(m) \<lesssim> domain(r);  r \<lesssim> succ(m);  m \<in> nat;   
+         r \<subseteq> A*B; A=domain(r) |] ==> r \<in> A->B"
+by (simp add: Pi_iff rel_domain_ex1)
+
 end
--- a/src/ZF/AC/Cardinal_aux.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,209 +0,0 @@
-(*  Title:      ZF/AC/Cardinal_aux.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Auxiliary lemmas concerning cardinalities
-*)
-
-(* ********************************************************************** *)
-(* Lemmas involving ordinals and cardinalities used in the proofs         *)
-(* concerning AC16 and DC                                                 *)
-(* ********************************************************************** *)
-
-(* j=|A| *)
-Goal "[| A lepoll i; Ord(i) |] ==> \\<exists>j. j le i & A eqpoll j";
-by (blast_tac (claset() addSIs [lepoll_cardinal_le, well_ord_Memrel,
-				well_ord_cardinal_eqpoll RS eqpoll_sym]
-                        addDs [lepoll_well_ord]) 1);
-qed "lepoll_imp_ex_le_eqpoll";
-
-(* j=|A| *)
-Goalw [lesspoll_def]
-    "[| A lesspoll i; Ord(i) |] ==> \\<exists>j. j<i & A eqpoll j";
-by (blast_tac (claset() addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
-qed "lesspoll_imp_ex_lt_eqpoll";
-
-Goalw [InfCard_def] "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
-by (rtac conjI 1);
-by (rtac Card_cardinal 1);
-by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1);
-by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll
-        RSN (2, well_ord_Memrel RS well_ord_lepoll_imp_Card_le)] 1 
-    THEN REPEAT (assume_tac 1));
-qed "Inf_Ord_imp_InfCard_cardinal";
-
-Goal "[| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |]  \
-\               ==> A Un B eqpoll i";
-by (rtac eqpollI 1);
-by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll
-        RS  lepoll_trans)] 2);
-by (Fast_tac 2);
-by (resolve_tac [Un_lepoll_sum RS lepoll_trans] 1);
-by (resolve_tac [lepoll_imp_sum_lepoll_prod RS lepoll_trans] 1);
-by (eresolve_tac [eqpoll_sym RSN (2, eqpoll_trans) RS eqpoll_imp_lepoll] 1
-        THEN (assume_tac 1));
-by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll RS 
-        (Ord_nat RS (nat_2I RS OrdmemD) RS subset_imp_lepoll RS lepoll_trans)
-        RS (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans))] 1 
-    THEN (REPEAT (assume_tac 1)));
-by (eresolve_tac [prod_eqpoll_cong RS eqpoll_imp_lepoll RS lepoll_trans] 1 
-    THEN (assume_tac 1));
-by (resolve_tac [Inf_Ord_imp_InfCard_cardinal RSN (2, well_ord_Memrel RS 
-        well_ord_InfCard_square_eq) RS eqpoll_imp_lepoll] 1 
-    THEN REPEAT (assume_tac 1));
-qed "Un_eqpoll_Inf_Ord";
-
-
-Goal "?f \\<in> bij({{y,z}. y \\<in> x}, x)";
-by (rtac RepFun_bijective 1);
-by (simp_tac (simpset() addsimps [doubleton_eq_iff]) 1);
-by (Blast_tac 1);
-qed "paired_bij";
-
-Goalw [eqpoll_def] "{{y,z}. y \\<in> x} eqpoll x";
-by (fast_tac (claset() addSIs [paired_bij]) 1);
-qed "paired_eqpoll";
-
-Goal "\\<exists>B. B eqpoll A & B Int C = 0";
-by (fast_tac (claset() addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
-qed "ex_eqpoll_disjoint";
-
-Goal "[| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |]  \
-\               ==> A Un B lepoll i";
-by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1);
-by (etac conjE 1);
-by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans)] 1);
-by (assume_tac 1);
-by (resolve_tac [Un_lepoll_Un RS lepoll_trans] 1 THEN (REPEAT (assume_tac 1)));
-by (eresolve_tac [eqpoll_refl RSN (2, Un_eqpoll_Inf_Ord) RS
-        eqpoll_imp_lepoll] 1
-        THEN (REPEAT (assume_tac 1)));
-qed "Un_lepoll_Inf_Ord";
-
-Goal "[| P(i); i \\<in> j; Ord(j) |] ==> (LEAST i. P(i)) \\<in> j";
-by (eresolve_tac [Least_le RS leE] 1);
-by (etac Ord_in_Ord 1 THEN (assume_tac 1));
-by (etac ltE 1);
-by (fast_tac (claset() addDs [OrdmemD]) 1);
-by (etac subst_elem 1 THEN (assume_tac 1));
-qed "Least_in_Ord";
-
-Goal "[| well_ord(x,r); y \\<subseteq> x; y lepoll succ(n); n \\<in> nat |]  \
-\       ==> y-{THE b. first(b,y,r)} lepoll n";
-by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1);
-by (fast_tac (claset() addSIs [Diff_sing_lepoll, the_first_in]) 1);
-by (res_inst_tac [("b","y-{THE b. first(b, y, r)}")] subst 1);
-by (rtac empty_lepollI 2);
-by (Fast_tac 1);
-qed "Diff_first_lepoll";
-
-Goal "(\\<Union>x \\<in> X. P(x)) \\<subseteq> (\\<Union>x \\<in> X. P(x)-Q(x)) Un (\\<Union>x \\<in> X. Q(x))";
-by (Fast_tac 1);
-qed "UN_subset_split";
-
-Goalw [lepoll_def] "Ord(a) ==> (\\<Union>x \\<in> a. {P(x)}) lepoll a";
-by (res_inst_tac [("x","\\<lambda>z \\<in> (\\<Union>x \\<in> a. {P(x)}). (LEAST i. P(i)=z)")] exI 1);
-by (res_inst_tac [("d","%z. P(z)")] lam_injective 1);
-by (fast_tac (claset() addSIs [Least_in_Ord]) 1);
-by (fast_tac (claset() addIs [LeastI] addSEs [Ord_in_Ord]) 1);
-qed "UN_sing_lepoll";
-
-Goal "[| well_ord(T, R); ~Finite(a); Ord(a); n \\<in> nat |] ==>  \
-\       \\<forall>f. (\\<forall>b \\<in> a. f`b lepoll n & f`b \\<subseteq> T) --> (\\<Union>b \\<in> a. f`b) lepoll a";
-by (induct_tac "n" 1);
-by (rtac allI 1);
-by (rtac impI 1);
-by (res_inst_tac [("b","\\<Union>b \\<in> a. f`b")] subst 1);
-by (rtac empty_lepollI 2);
-by (resolve_tac [equals0I RS sym] 1);
-by (REPEAT (eresolve_tac [UN_E, allE] 1));
-by (fast_tac (claset() addDs [lepoll_0_is_0 RS subst]) 1);
-by (rtac allI 1);
-by (rtac impI 1);
-by (eres_inst_tac [("x","\\<lambda>x \\<in> a. f`x - {THE b. first(b,f`x,R)}")] allE 1);
-by (etac impE 1);
-by (Asm_full_simp_tac 1);
-by (fast_tac (claset() addSIs [Diff_first_lepoll]) 1);
-by (Asm_full_simp_tac 1);
-by (resolve_tac [UN_subset_split RS subset_imp_lepoll RS lepoll_trans] 1);
-by (rtac Un_lepoll_Inf_Ord 1 THEN (REPEAT_FIRST assume_tac));
-by (etac UN_sing_lepoll 1);
-qed "UN_fun_lepoll_lemma";
-
-Goal "[| \\<forall>b \\<in> a. f`b lepoll n & f`b \\<subseteq> T; well_ord(T, R);  \
-\       ~Finite(a); Ord(a); n \\<in> nat |] ==> (\\<Union>b \\<in> a. f`b) lepoll a";
-by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1)));
-by (Fast_tac 1);
-qed "UN_fun_lepoll";
-
-Goal "[| \\<forall>b \\<in> a. F(b) lepoll n & F(b) \\<subseteq> T; well_ord(T, R);  \
-\       ~Finite(a); Ord(a); n \\<in> nat |] ==> (\\<Union>b \\<in> a. F(b)) lepoll a";
-by (rtac impE 1 THEN (assume_tac 3));
-by (res_inst_tac [("f","\\<lambda>b \\<in> a. F(b)")] (UN_fun_lepoll) 2 
-        THEN (TRYALL assume_tac));
-by Auto_tac;
-qed "UN_lepoll";
-
-Goal "Ord(a) ==> (\\<Union>b \\<in> a. F(b)) = (\\<Union>b \\<in> a. F(b) - (\\<Union>c \\<in> b. F(c)))";
-by (rtac equalityI 1);
-by (Fast_tac 2);
-by (rtac subsetI 1);
-by (etac UN_E 1);
-by (rtac UN_I 1);
-by (res_inst_tac [("P","%z. x \\<in> F(z)")] Least_in_Ord 1 THEN (REPEAT (assume_tac 1)));
-by (rtac DiffI 1);
-by (resolve_tac [Ord_in_Ord RSN (2, LeastI)] 1 THEN (REPEAT (assume_tac 1)));
-by (rtac notI 1);
-by (etac UN_E 1);
-by (eres_inst_tac [("P","%z. x \\<in> F(z)"),("i","c")] less_LeastE 1);
-by (eresolve_tac [Ord_Least RSN (2, ltI)] 1);
-qed "UN_eq_UN_Diffs";
-
-Goalw [lepoll_def, eqpoll_def]
-     "a lepoll X ==> \\<exists>Y. Y \\<subseteq> X & a eqpoll Y";
-by (etac exE 1);
-by (forward_tac [subset_refl RSN (2, restrict_bij)] 1);
-by (res_inst_tac [("x","f``a")] exI 1);
-by (fast_tac (claset() addSEs [inj_is_fun RS fun_is_rel RS image_subset]) 1);
-qed "lepoll_imp_eqpoll_subset";
-
-(* ********************************************************************** *)
-(* Diff_lesspoll_eqpoll_Card                                              *)
-(* ********************************************************************** *)
-
-Goal "[| A\\<approx>a; ~Finite(a); Card(a); B lesspoll a; A-B lesspoll a |] ==> P";
-by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE,
-        Card_is_Ord, conjE] 1));
-by (forw_inst_tac [("j","xa")] ([lt_Ord, lt_Ord] MRS Un_upper1_le) 1
-        THEN (assume_tac 1));
-by (forw_inst_tac [("j","xa")] ([lt_Ord, lt_Ord] MRS Un_upper2_le) 1
-        THEN (assume_tac 1));
-by (dtac Un_least_lt 1 THEN (assume_tac 1));
-by (dresolve_tac [le_imp_lepoll RSN
-        (2, eqpoll_imp_lepoll RS lepoll_trans)] 1
-        THEN (assume_tac 1));
-by (dresolve_tac [le_imp_lepoll RSN
-        (2, eqpoll_imp_lepoll RS lepoll_trans)] 1
-        THEN (assume_tac 1));
-by (res_inst_tac [("Q","Finite(x Un xa)")] (excluded_middle RS disjE) 1);
-by (dresolve_tac [[lepoll_Finite, lepoll_Finite] MRS Finite_Un] 2
-        THEN (REPEAT (assume_tac 2)));
-by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_Finite] 2);
-by (fast_tac (claset()
-        addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2);
-by (dresolve_tac [ Un_lepoll_Inf_Ord] 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (claset() addSEs [ltE, Ord_in_Ord]) 1);
-by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_trans RS
-         (lt_Card_imp_lesspoll RSN (2, lesspoll_trans1))] 1
-        THEN (TRYALL assume_tac));
-by (fast_tac (claset() addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
-qed "Diff_lesspoll_eqpoll_Card_lemma";
-
-Goal "[| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |]  \
-\       ==> A - B eqpoll a";
-by (rtac swap 1 THEN (Fast_tac 1));
-by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (claset() addSIs [lesspoll_def RS def_imp_iff RS iffD2,
-        subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
-qed "Diff_lesspoll_eqpoll_Card";
-
--- a/src/ZF/AC/Cardinal_aux.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/Cardinal_aux.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,212 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/Cardinal_aux.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+Auxiliary lemmas concerning cardinalities
+*)
+
+theory Cardinal_aux = AC_Equiv:
+
+lemma Diff_lepoll: "[| A \<lesssim> succ(m); B \<subseteq> A; B\<noteq>0 |] ==> A-B \<lesssim> m"
+apply (rule not_emptyE, (assumption))
+apply (blast intro: lepoll_trans [OF subset_imp_lepoll Diff_sing_lepoll])
+done
+
+
+(* ********************************************************************** *)
+(* Lemmas involving ordinals and cardinalities used in the proofs         *)
+(* concerning AC16 and DC                                                 *)
+(* ********************************************************************** *)
+
+
+(* j=|A| *)
+lemma lepoll_imp_ex_le_eqpoll:
+     "[| A \<lesssim> i; Ord(i) |] ==> \<exists>j. j le i & A \<approx> j"
+by (blast intro!: lepoll_cardinal_le well_ord_Memrel 
+                  well_ord_cardinal_eqpoll [THEN eqpoll_sym]
+          dest: lepoll_well_ord);
+
+(* j=|A| *)
+lemma lesspoll_imp_ex_lt_eqpoll: 
+     "[| A \<prec> i; Ord(i) |] ==> \<exists>j. j<i & A \<approx> j"
+by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE)
+
+lemma Inf_Ord_imp_InfCard_cardinal: "[| ~Finite(i); Ord(i) |] ==> InfCard(|i|)"
+apply (unfold InfCard_def)
+apply (rule conjI)
+apply (rule Card_cardinal)
+apply (rule Card_nat 
+            [THEN Card_def [THEN def_imp_iff, THEN iffD1, THEN ssubst]])
+  -- "rewriting would loop!"
+apply (rule well_ord_Memrel [THEN well_ord_lepoll_imp_Card_le], assumption) 
+apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+)
+done
+
+text{*An alternative and more general proof goes like this: A and B are both
+well-ordered (because they are injected into an ordinal), either A lepoll B
+or B lepoll A.  Also both are equipollent to their cardinalities, so
+(if A and B are infinite) then A Un B lepoll |A|+|B| = max(|A|,|B|) lepoll i.
+In fact, the correctly strengthened version of this theorem appears below.*}
+lemma Un_lepoll_Inf_Ord_weak:
+     "[|A \<approx> i; B \<approx> i; \<not> Finite(i); Ord(i)|] ==> A \<union> B \<lesssim> i"
+apply (rule Un_lepoll_sum [THEN lepoll_trans])
+apply (rule lepoll_imp_sum_lepoll_prod [THEN lepoll_trans])
+apply (erule eqpoll_trans [THEN eqpoll_imp_lepoll]) 
+apply (erule eqpoll_sym) 
+apply (rule subset_imp_lepoll [THEN lepoll_trans, THEN lepoll_trans]) 
+apply (rule nat_2I [THEN OrdmemD], rule Ord_nat) 
+apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption+) 
+apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) 
+apply (erule prod_eqpoll_cong [THEN eqpoll_imp_lepoll, THEN lepoll_trans],
+       assumption)
+apply (rule eqpoll_imp_lepoll) 
+apply (rule well_ord_Memrel [THEN well_ord_InfCard_square_eq], assumption) 
+apply (rule Inf_Ord_imp_InfCard_cardinal, assumption+) 
+done
+
+lemma Un_eqpoll_Inf_Ord:
+     "[| A \<approx> i; B \<approx> i; ~Finite(i); Ord(i) |] ==> A Un B \<approx> i"
+apply (rule eqpollI)
+apply (blast intro: Un_lepoll_Inf_Ord_weak) 
+apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]) 
+apply (rule Un_upper1 [THEN subset_imp_lepoll]) 
+done
+
+lemma paired_bij: "?f \<in> bij({{y,z}. y \<in> x}, x)"
+apply (rule RepFun_bijective)
+apply (simp add: doubleton_eq_iff, blast)
+done
+
+lemma paired_eqpoll: "{{y,z}. y \<in> x} \<approx> x"
+by (unfold eqpoll_def, fast intro!: paired_bij)
+
+lemma ex_eqpoll_disjoint: "\<exists>B. B \<approx> A & B Int C = 0"
+by (fast intro!: paired_eqpoll equals0I elim: mem_asym)
+
+(*Finally we reach this result.  Surely there's a simpler proof, as sketched
+  above?*)
+lemma Un_lepoll_Inf_Ord:
+     "[| A \<lesssim> i; B \<lesssim> i; ~Finite(i); Ord(i) |] ==> A Un B \<lesssim> i"
+apply (rule_tac A1 = "i" and C1 = "i" in ex_eqpoll_disjoint [THEN exE])
+apply (erule conjE)
+apply (drule lepoll_trans) 
+apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll])
+apply (rule Un_lepoll_Un [THEN lepoll_trans], (assumption+))
+apply (blast intro: eqpoll_refl Un_eqpoll_Inf_Ord eqpoll_imp_lepoll) 
+done
+
+lemma Least_in_Ord: "[| P(i); i \<in> j; Ord(j) |] ==> (LEAST i. P(i)) \<in> j"
+apply (erule Least_le [THEN leE])
+apply (erule Ord_in_Ord, assumption)
+apply (erule ltE)
+apply (fast dest: OrdmemD)
+apply (erule subst_elem, assumption)
+done
 
-Cardinal_aux = AC_Equiv
+lemma Diff_first_lepoll:
+     "[| well_ord(x,r); y \<subseteq> x; y \<lesssim> succ(n); n \<in> nat |] 
+      ==> y - {THE b. first(b,y,r)} \<lesssim> n"
+apply (case_tac "y=0", simp add: empty_lepollI) 
+apply (fast intro!: Diff_sing_lepoll the_first_in)
+done
+
+lemma UN_subset_split:
+     "(\<Union>x \<in> X. P(x)) \<subseteq> (\<Union>x \<in> X. P(x)-Q(x)) Un (\<Union>x \<in> X. Q(x))"
+by blast
+
+lemma UN_sing_lepoll: "Ord(a) ==> (\<Union>x \<in> a. {P(x)}) \<lesssim> a"
+apply (unfold lepoll_def)
+apply (rule_tac x = "\<lambda>z \<in> (\<Union>x \<in> a. {P (x) }) . (LEAST i. P (i) =z) " in exI)
+apply (rule_tac d = "%z. P (z) " in lam_injective)
+apply (fast intro!: Least_in_Ord)
+apply (fast intro: LeastI elim!: Ord_in_Ord)
+done
+
+lemma UN_fun_lepoll_lemma [rule_format]:
+     "[| well_ord(T, R); ~Finite(a); Ord(a); n \<in> nat |] 
+      ==> \<forall>f. (\<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T) --> (\<Union>b \<in> a. f`b) \<lesssim> a"
+apply (induct_tac "n")
+apply (rule allI)
+apply (rule impI)
+apply (rule_tac b = "\<Union>b \<in> a. f`b" in subst)
+apply (rule_tac [2] empty_lepollI)
+apply (rule equals0I [symmetric], clarify) 
+apply (fast dest: lepoll_0_is_0 [THEN subst])
+apply (rule allI)
+apply (rule impI)
+apply (erule_tac x = "\<lambda>x \<in> a. f`x - {THE b. first (b,f`x,R) }" in allE)
+apply (erule impE, simp)
+apply (fast intro!: Diff_first_lepoll, simp)
+apply (rule UN_subset_split [THEN subset_imp_lepoll, THEN lepoll_trans])
+apply (fast intro: Un_lepoll_Inf_Ord UN_sing_lepoll) 
+done
+
+lemma UN_fun_lepoll:
+     "[| \<forall>b \<in> a. f`b \<lesssim> n & f`b \<subseteq> T; well_ord(T, R);   
+         ~Finite(a); Ord(a); n \<in> nat |] ==> (\<Union>b \<in> a. f`b) \<lesssim> a"
+by (blast intro: UN_fun_lepoll_lemma); 
+
+lemma UN_lepoll:
+     "[| \<forall>b \<in> a. F(b) \<lesssim> n & F(b) \<subseteq> T; well_ord(T, R);   
+         ~Finite(a); Ord(a); n \<in> nat |] 
+      ==> (\<Union>b \<in> a. F(b)) \<lesssim> a"
+apply (rule rev_mp) 
+apply (rule_tac f="\<lambda>b \<in> a. F (b)" in UN_fun_lepoll);
+apply auto
+done
+
+lemma UN_eq_UN_Diffs:
+     "Ord(a) ==> (\<Union>b \<in> a. F(b)) = (\<Union>b \<in> a. F(b) - (\<Union>c \<in> b. F(c)))"
+apply (rule equalityI)
+ prefer 2 apply fast
+apply (rule subsetI)
+apply (erule UN_E)
+apply (rule UN_I)
+ apply (rule_tac P = "%z. x \<in> F (z) " in Least_in_Ord, (assumption+))
+apply (rule DiffI, best intro: Ord_in_Ord LeastI, clarify)
+apply (erule_tac P = "%z. x \<in> F (z) " and i = "c" in less_LeastE)
+apply (blast intro: Ord_Least ltI)
+done
+
+lemma lepoll_imp_eqpoll_subset: 
+     "a \<lesssim> X ==> \<exists>Y. Y \<subseteq> X & a \<approx> Y"
+apply (unfold lepoll_def eqpoll_def, clarify) 
+apply (blast intro: restrict_bij
+             dest: inj_is_fun [THEN fun_is_rel, THEN image_subset]) 
+done
+
+(* ********************************************************************** *)
+(* Diff_lesspoll_eqpoll_Card                                              *)
+(* ********************************************************************** *)
+
+lemma Diff_lesspoll_eqpoll_Card_lemma:
+     "[| A\<approx>a; ~Finite(a); Card(a); B \<prec> a; A-B \<prec> a |] ==> P"
+apply (elim lesspoll_imp_ex_lt_eqpoll [THEN exE] Card_is_Ord conjE)
+apply (frule_tac j=xa in Un_upper1_le [OF lt_Ord lt_Ord], assumption)
+apply (frule_tac j=xa in Un_upper2_le [OF lt_Ord lt_Ord], assumption)
+apply (drule Un_least_lt, assumption)
+apply (drule eqpoll_imp_lepoll [THEN lepoll_trans], 
+       rule le_imp_lepoll, assumption)+
+apply (case_tac "Finite(x Un xa)");
+txt{*finite case*}
+ apply (drule Finite_Un [OF lepoll_Finite lepoll_Finite], assumption+) 
+ apply (drule subset_Un_Diff [THEN subset_imp_lepoll, THEN lepoll_Finite])
+ apply (fast dest: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_Finite])
+txt{*infinite case*}
+apply (drule Un_lepoll_Inf_Ord, (assumption+))
+apply (blast intro: le_Ord2) 
+apply (drule lesspoll_trans1 
+             [OF subset_Un_Diff [THEN subset_imp_lepoll, THEN lepoll_trans] 
+                 lt_Card_imp_lesspoll], assumption+)
+apply (simp add: lesspoll_def) 
+done
+
+lemma Diff_lesspoll_eqpoll_Card:
+     "[| A \<approx> a; ~Finite(a); Card(a); B \<prec> a |] ==> A - B \<approx> a"
+apply (rule ccontr)
+apply (rule Diff_lesspoll_eqpoll_Card_lemma, (assumption+))
+apply (blast intro: lesspoll_def [THEN def_imp_iff, THEN iffD2] 
+                    subset_imp_lepoll eqpoll_imp_lepoll lepoll_trans)
+done
+
+end
--- a/src/ZF/AC/DC.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,533 +0,0 @@
-(*  Title:      ZF/AC/DC.ML
-    ID:  $Id$
-    Author:     Krzysztof Grabczewski
-
-The proofs concerning the Axiom of Dependent Choice
-*)
-
-(* ********************************************************************** *)
-(* DC ==> DC(omega)                                                       *)
-(*                                                                        *)
-(* The scheme of the proof:                                               *)
-(*                                                                        *)
-(* Assume DC. Let R and X satisfy the premise of DC(omega).               *)
-(*                                                                        *)
-(* Define XX and RR as follows:                                           *)
-(*                                                                        *)
-(*       XX = (\\<Union>n \\<in> nat. {f \\<in> n->X. \\<forall>k \\<in> n. <f``k, f`k> \\<in> R})              *)
-(*       f RR g iff domain(g)=succ(domain(f)) &                           *)
-(*              restrict(g, domain(f)) = f                                *)
-(*                                                                        *)
-(* Then RR satisfies the hypotheses of DC.                                *)
-(* So applying DC:                                                        *)
-(*                                                                        *)
-(*       \\<exists>f \\<in> nat->XX. \\<forall>n \\<in> nat. f`n RR f`succ(n)                        *)
-(*                                                                        *)
-(* Thence                                                                 *)
-(*                                                                        *)
-(*       ff = {<n, f`succ(n)`n>. n \\<in> nat}                                   *)
-(*                                                                        *)
-(* is the desired function.                                               *)
-(*                                                                        *)
-(* ********************************************************************** *)
-
-Open_locale "DC0_imp";
-
-val all_ex = thm "all_ex";
-val XX_def = thm "XX_def";
-val RR_def = thm "RR_def";
-
-Goalw [RR_def] "RR \\<subseteq> XX*XX";
-by (Fast_tac 1);
-qed "lemma1_1";
-
-Goalw [RR_def, XX_def] "RR \\<noteq> 0";
-by (rtac (all_ex RS ballE) 1);
-by (eresolve_tac [empty_subsetI RS PowI RSN (2, notE)] 2);
-by (eresolve_tac [nat_0I RS n_lesspoll_nat RSN (2, impE)] 1);
-by (etac bexE 1);
-by (res_inst_tac [("a","<0, {<0, x>}>")] not_emptyI 1);
-by (rtac CollectI 1);
-by (rtac SigmaI 1);
-by (rtac (nat_0I RS UN_I) 1);
-by (asm_simp_tac (simpset() addsimps [nat_0I RS UN_I]) 1);
-by (rtac (nat_1I RS UN_I) 1);
-by (asm_simp_tac (simpset() addsimps [singleton_0]) 2);
-by (force_tac (claset() addSIs [singleton_fun RS Pi_type],
-	       simpset() addsimps [singleton_0 RS sym]) 1);
-qed "lemma1_2";
-
-Goalw [RR_def, XX_def] "range(RR) \\<subseteq> domain(RR)";
-by (rtac range_subset_domain 1);
-by (Blast_tac 2);
-by (Clarify_tac 1);
-by (forward_tac [fun_is_rel RS image_subset RS PowI RS (all_ex RS bspec)] 1);
-by (eresolve_tac [[nat_into_Ord RSN (2, image_Ord_lepoll), n_lesspoll_nat]
-        MRS lesspoll_trans1 RSN (2, impE)] 1
-        THEN REPEAT (assume_tac 1));
-by (etac bexE 1);
-by (res_inst_tac [("x","cons(<n,x>, g)")] exI 1);
-by (rtac CollectI 1);
-by (force_tac (claset() addSEs [cons_fun_type2],
-	       simpset() addsimps [cons_image_n, cons_val_n, 
-				   cons_image_k, cons_val_k]) 1);
-by (asm_full_simp_tac (simpset()
-        addsimps [domain_of_fun, succ_def, restrict_cons_eq]) 1);
-qed "lemma1_3";
-
-
-Goal "[| \\<forall>n \\<in> nat. <f`n, f`succ(n)> \\<in> RR;  f \\<in> nat -> XX;  n \\<in> nat |]  \
-\     ==> \\<exists>k \\<in> nat. f`succ(n) \\<in> k -> X & n \\<in> k  \
-\                 & <f`succ(n)``n, f`succ(n)`n> \\<in> R";
-by (induct_tac "n" 1);
-by (dresolve_tac [nat_1I RSN (2, apply_type)] 1);
-by (dresolve_tac [nat_0I RSN (2, bspec)] 1);
-by (asm_full_simp_tac (simpset() addsimps [XX_def]) 1);
-by Safe_tac;
-by (rtac bexI 1 THEN (assume_tac 2));
-by (best_tac (claset() addIs [ltD]
-                        addSEs [nat_0_le RS leE]
-			addEs [sym RS trans RS succ_neq_0, domain_of_fun]
-	       addss (simpset() addsimps [RR_def])) 1);
-(** LEVEL 7, other subgoal **)
-by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (subgoal_tac "f ` succ(succ(x)) \\<in> succ(k)->X" 1);
-by (dresolve_tac [nat_succI RS nat_succI RSN (2, apply_type)] 1
-        THEN (assume_tac 1));
-by (full_simp_tac (simpset() addsimps [XX_def,RR_def]) 1);
-by Safe_tac;
-by (forw_inst_tac [("a","succ(k)")] (domain_of_fun RS sym RS trans) 1 THEN
-    (assume_tac 1));
-by (forw_inst_tac [("a","xa")] (domain_of_fun RS sym RS trans) 1 THEN
-    (assume_tac 1));
-by (fast_tac (claset() addSEs [nat_into_Ord RS succ_in_succ]
-        addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1);
-by (dtac domain_of_fun 1);
-by (full_simp_tac (simpset() addsimps [XX_def,RR_def]) 1);
-by (deepen_tac (claset() addDs [domain_of_fun RS sym RS trans]) 0 1);
-qed "lemma2";
-
-Goal "[| \\<forall>n \\<in> nat. <f`n, f`succ(n)> \\<in> RR;  f \\<in> nat -> XX;  m \\<in> nat |]  \
-\     ==>  {f`succ(x)`x. x \\<in> m} = {f`succ(m)`x. x \\<in> m}";
-by (subgoal_tac "\\<forall>x \\<in> m. f`succ(m)`x = f`succ(x)`x" 1);
-by (Asm_full_simp_tac 1);
-by (induct_tac "m" 1);
-by (Fast_tac 1);
-by (rtac ballI 1);
-by (etac succE 1);
-by (rtac restrict_eq_imp_val_eq 1);
-by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [RR_def]) 1);
-by (dtac lemma2 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSDs [domain_of_fun]) 1);
-by (dres_inst_tac [("x","xa")] bspec 1 THEN (assume_tac 1));
-by (eresolve_tac [sym RS trans RS sym] 1);
-by (resolve_tac [restrict_eq_imp_val_eq RS sym] 1);
-by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [RR_def]) 1);
-by (dtac lemma2 1 THEN REPEAT (assume_tac 1));
-by (blast_tac (claset() addSDs [domain_of_fun]
-                        addIs [nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1);
-qed "lemma3_1";
-
-Goal "[| \\<forall>n \\<in> nat. <f`n, f`succ(n)> \\<in> RR;  f \\<in> nat -> XX;  m \\<in> nat |] \
-\     ==> (\\<lambda>x \\<in> nat. f`succ(x)`x) `` m = f`succ(m)``m";
-by (etac natE 1);
-by (Asm_simp_tac 1);
-by (stac image_lam 1);
-by (fast_tac (claset() addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1);
-by (stac lemma3_1 1 THEN REPEAT (assume_tac 1));
-by (Fast_tac 1);
-by (fast_tac (claset() addSDs [lemma2]
-		       addSEs [nat_into_Ord RSN (2, OrdmemD) RSN 
-                            (2, image_fun RS sym)]) 1);
-qed "lemma3";
-
-Close_locale "DC0_imp";
-
-Goalw [DC_def, DC0_def] "DC0 ==> DC(nat)";
-by (Clarify_tac 1);
-by (REPEAT (etac allE 1));
-by (etac impE 1);
-   (*these three results comprise Lemma 1*)
-by (blast_tac (claset() addSIs (map export [lemma1_1, lemma1_2, lemma1_3])) 1);
-by (etac bexE 1);
-by (res_inst_tac [("x","\\<lambda>n \\<in> nat. f`succ(n)`n")] bexI 1);
-by (fast_tac (claset() addSIs [lam_type] addSDs [export lemma2]
-                       addSEs [fun_weaken_type, apply_type]) 2);
-by (rtac oallI 1);
-by (forward_tac [ltD RSN (3, export lemma2)] 1
-        THEN assume_tac 2);
-by (fast_tac (claset() addSEs [fun_weaken_type]) 1);
-(** LEVEL 10: last subgoal **)
-by (stac (ltD RSN (3, export lemma3)) 1);
-by (Force_tac 4);
-by (assume_tac 3);
-by (assume_tac 1);
-by (fast_tac (claset() addSEs [fun_weaken_type]) 1);
-qed "DC0_imp_DC_nat";
-
-
-(* ************************************************************************
-   DC(omega) ==> DC                                                       
-                                                                          
-   The scheme of the proof:                                               
-                                                                          
-   Assume DC(omega). Let R and x satisfy the premise of DC.               
-                                                                          
-   Define XX and RR as follows:                                           
-                                                                          
-    XX = (\\<Union>n \\<in> nat. {f \\<in> succ(n)->domain(R). \\<forall>k \\<in> n. <f`k, f`succ(k)> \\<in> R})
-
-    RR = {<z1,z2>:Fin(XX)*XX. 
-           (domain(z2)=succ(\\<Union>f \\<in> z1. domain(f)) &
-	    (\\<forall>f \\<in> z1. restrict(z2, domain(f)) = f)) |      
-	   (~ (\\<exists>g \\<in> XX. domain(g)=succ(\\<Union>f \\<in> z1. domain(f)) &
-	                (\\<forall>f \\<in> z1. restrict(g, domain(f)) = f)) &           
-	    z2={<0,x>})}                                          
-                                                                          
-   Then XX and RR satisfy the hypotheses of DC(omega).                    
-   So applying DC:                                                        
-                                                                          
-         \\<exists>f \\<in> nat->XX. \\<forall>n \\<in> nat. f``n RR f`n                             
-                                                                          
-   Thence                                                                 
-                                                                          
-         ff = {<n, f`n`n>. n \\<in> nat}                                         
-                                                                          
-   is the desired function.                                               
-                                                                          
-************************************************************************* *)
-
-Goal "n \\<in> nat ==> \\<forall>A. (A eqpoll n & A \\<subseteq> X) --> A \\<in> Fin(X)";
-by (induct_tac "n" 1);
-by (rtac allI 1);
-by (fast_tac (claset() addSIs [Fin.emptyI]
-        addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
-by (rtac allI 1);
-by (rtac impI 1);
-by (etac conjE 1);
-by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1
-        THEN (assume_tac 1));
-by (ftac Diff_sing_eqpoll 1 THEN (assume_tac 1));
-by (etac allE 1);
-by (etac impE 1);
-by (Fast_tac 1);
-by (dtac subsetD 1 THEN (assume_tac 1));
-by (dresolve_tac [Fin.consI] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [cons_Diff]) 1);
-qed "Finite_Fin_lemma";
-
-Goalw [Finite_def] "[| Finite(A); A \\<subseteq> X |] ==> A \\<in> Fin(X)";
-by (etac bexE 1);
-by (dtac Finite_Fin_lemma 1);
-by (etac allE 1);
-by (etac impE 1);
-by (assume_tac 2);
-by (Fast_tac 1);
-qed "Finite_Fin";
-
-Goal
- "x \\<in> X ==> {<0,x>}: (\\<Union>n \\<in> nat. {f \\<in> succ(n)->X. \\<forall>k \\<in> n. <f`k, f`succ(k)> \\<in> R})";
-by (rtac (nat_0I RS UN_I) 1);
-by (fast_tac (claset() addSIs [singleton_fun RS Pi_type]
-        addss (simpset() addsimps [singleton_0 RS sym])) 1);
-qed "singleton_in_funs";
-
-
-Open_locale "imp_DC0";
-
-val XX_def = thm "XX_def";
-val RR_def = thm "RR_def";
-val allRR_def = thm "allRR_def";
-
-Goal "[| range(R) \\<subseteq> domain(R);  x \\<in> domain(R) |]  \
-\     ==> RR \\<subseteq> Pow(XX)*XX &  \
-\            (\\<forall>Y \\<in> Pow(XX). Y lesspoll nat --> (\\<exists>x \\<in> XX. <Y,x>:RR))";
-by (rtac conjI 1);
-by (force_tac (claset() addSDs [FinD RS PowI], 
-	       simpset() addsimps [RR_def]) 1); 
-by (rtac (impI RS ballI) 1);
-by (dresolve_tac [[lesspoll_nat_is_Finite, PowD] MRS Finite_Fin] 1
-        THEN (assume_tac 1));
-by (excluded_middle_tac "\\<exists>g \\<in> XX. domain(g)=succ(\\<Union>f \\<in> Y. domain(f))  \
-\       & (\\<forall>f \\<in> Y. restrict(g, domain(f)) = f)" 1);
-by (fast_tac (claset() addss (simpset() addsimps [RR_def])) 2); 
-by (safe_tac (claset() delrules [domainE]));
-by (rewrite_goals_tac [XX_def,RR_def]);
-by (swap_res_tac [bexI] 1 THEN etac singleton_in_funs 2);
-by (asm_full_simp_tac (simpset() addsimps [nat_0I  RSN (2, bexI), 
-					   cons_fun_type2]) 1);
-qed "lemma4";
-
-Goal "[| f \\<in> nat->X; n \\<in> nat |] ==>  \
-\       (\\<Union>x \\<in> f``succ(n). P(x)) =  P(f`n) Un (\\<Union>x \\<in> f``n. P(x))";
-by (asm_full_simp_tac (simpset()
-        addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun),
-        [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
-qed "UN_image_succ_eq";
-
-Goal "[| (\\<Union>x \\<in> f``n. P(x)) = y; P(f`n) = succ(y);  \
-\        f \\<in> nat -> X; n \\<in> nat |] ==> (\\<Union>x \\<in> f``succ(n). P(x)) = succ(y)";
-by (asm_full_simp_tac (simpset() addsimps [UN_image_succ_eq]) 1);
-by (Fast_tac 1);
-qed "UN_image_succ_eq_succ";
-
-Goal "[| f \\<in> succ(n) -> D;  n \\<in> nat;  \
-\        domain(f)=succ(x); x=y |] ==> f`y \\<in> D";
-by (fast_tac (claset() addEs [apply_type]
-	      addSDs [[sym, domain_of_fun] MRS trans]) 1);
-qed "apply_domain_type";
-
-Goal "[| f \\<in> nat -> X; n \\<in> nat |] ==> f``succ(n) = cons(f`n, f``n)";
-by (asm_full_simp_tac (simpset()
-        addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1);
-qed "image_fun_succ";
-
-Goalw [XX_def] "[| domain(f`n) = succ(u); f \\<in> nat -> XX;  u=k;  n \\<in> nat |]   \
-\               ==> f`n \\<in> succ(k) -> domain(R)";
-by (dtac apply_type 1 THEN (assume_tac 1));
-by (fast_tac (claset() addEs [domain_eq_imp_fun_type]) 1);
-qed "f_n_type";
-
-Goalw [XX_def]
-     "[| f \\<in> nat -> XX;  domain(f`n) = succ(k);  n \\<in> nat |]  \
-\     ==> \\<forall>i \\<in> k. <f`n`i, f`n`succ(i)> \\<in> R";
-by (dtac apply_type 1 THEN (assume_tac 1));
-by (etac UN_E 1);
-by (etac CollectE 1);
-by (dresolve_tac [domain_of_fun RS sym RS trans] 1 THEN (assume_tac 1));
-by (Asm_full_simp_tac 1);
-qed "f_n_pairs_in_R";
-
-Goalw [restrict_def]
-     "[| restrict(f, domain(x))=x;  f \\<in> n->X;  domain(x) \\<subseteq> n |]  \
-\     ==> restrict(cons(<n, y>, f), domain(x)) = x";
-by (eresolve_tac [sym RS trans RS sym] 1);
-by (rtac fun_extension 1);
-by (fast_tac (claset() addSIs [lam_type]) 1);
-by (fast_tac (claset() addSIs [lam_type]) 1);
-by (asm_full_simp_tac (simpset() addsimps [subsetD RS cons_val_k]) 1);
-qed "restrict_cons_eq_restrict";
-
-Goal "[| \\<forall>x \\<in> f``n. restrict(f`n, domain(x))=x;  \
-\        f \\<in> nat -> XX;  \
-\        n \\<in> nat;  domain(f`n) = succ(n);  \
-\        (\\<Union>x \\<in> f``n. domain(x)) \\<subseteq> n |] \
-\     ==> \\<forall>x \\<in> f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x)) = x";
-by (rtac ballI 1);
-by (asm_full_simp_tac (simpset() addsimps [image_fun_succ]) 1);
-by (dtac f_n_type 1 THEN REPEAT (ares_tac [refl] 1));
-by (etac disjE 1);
-by (asm_full_simp_tac (simpset() addsimps [domain_of_fun,restrict_cons_eq]) 1);
-by (dtac bspec 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [restrict_cons_eq_restrict]) 1);
-qed "all_in_image_restrict_eq";
-
-Goalw [RR_def, allRR_def]
-     "[| \\<forall>b<nat. <f``b, f`b> \\<in> RR;  \
-\        f \\<in> nat -> XX; range(R) \\<subseteq> domain(R); x \\<in> domain(R)|]   \
-\     ==> allRR";
-by (rtac oallI 1);
-by (dtac ltD 1);
-by (etac nat_induct 1);
-by (dresolve_tac [[nat_0I, Ord_nat] MRS ltI RSN (2, ospec)] 1);
-by (fast_tac (FOL_cs addss
-              (simpset() addsimps [singleton_fun RS domain_of_fun,
-                                  singleton_0, singleton_in_funs])) 1);
-(*induction step*) (** LEVEL 5 **)
-by (full_simp_tac (*prevent simplification of ~\\<exists>to \\<forall>~*)
-		  (FOL_ss addsimps [separation, split]) 1);
-by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1
-        THEN (assume_tac 1));
-by (REPEAT (eresolve_tac [conjE, disjE] 1));
-by (force_tac (FOL_cs addSEs [trans, subst_context]
-                     addSIs [UN_image_succ_eq_succ], simpset()) 1);
-by (etac conjE 1);
-by (etac notE 1);
-by (asm_lr_simp_tac (simpset() addsimps [XX_def, UN_image_succ_eq_succ]) 1);
-(** LEVEL 12 **)
-by (REPEAT (eresolve_tac [conjE, bexE] 1));
-by (dtac apply_domain_type 1 THEN REPEAT (assume_tac 1));
-by (etac domainE 1);
-by (etac domainE 1);
-by (forward_tac [export f_n_type] 1 THEN REPEAT (assume_tac 1));
-by (rtac bexI 1);
-by (etac nat_succI 2);
-by (res_inst_tac [("x","cons(<succ(xa), ya>, f`xa)")] bexI 1);
-by (rtac conjI 1);
-by (fast_tac (FOL_cs
-        addEs [subst_context RSN (2, trans) RS domain_cons_eq_succ,
-	       subst_context, export all_in_image_restrict_eq, 
-	       trans, equalityD1]) 2);
-by (eresolve_tac [rangeI RSN (2, subsetD) RSN (2, cons_fun_type2)] 2
-        THEN REPEAT (assume_tac 2));
-by (rtac ballI 1);
-by (etac succE 1);
-(** LEVEL 25 **)
-by (EVERY [dtac (domain_of_fun RSN (2, export f_n_pairs_in_R)) 2,
-  REPEAT (assume_tac 2), dtac bspec 2, assume_tac 2]);
-by (asm_full_simp_tac (simpset()
-        addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 2);
-by (asm_full_simp_tac (simpset() addsimps [cons_val_n, cons_val_k]) 1);
-qed "simplify_recursion";
-
-
-Goalw [allRR_def]
-     "[| allRR; f \\<in> nat -> XX; range(R) \\<subseteq> domain(R); x \\<in> domain(R); n \\<in> nat |]   \
-\     ==> f`n \\<in> succ(n) -> domain(R) & (\\<forall>i \\<in> n. <f`n`i, f`n`succ(i)>:R)";
-by (dtac ospec 1);
-by (eresolve_tac [Ord_nat RSN (2, ltI)] 1);
-by (etac CollectE 1);
-by (Asm_full_simp_tac 1);
-by (rtac conjI 1);
-by (fast_tac (FOL_cs addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 2);
-by (rewtac XX_def);
-by (fast_tac (claset()
-        addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1);
-qed "lemma2";
-
-Goal "[| allRR;  f \\<in> nat -> XX;  n \\<in> nat;  range(R) \\<subseteq> domain(R);  x \\<in> domain(R)  \
-\       |] ==> f`n`n = f`succ(n)`n";
-by (forward_tac [lemma2 RS conjunct1 RS domain_of_fun] 1
-        THEN REPEAT (assume_tac 1));
-by (rewtac allRR_def);
-by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1
-        THEN (assume_tac 1));
-by (Asm_full_simp_tac 1);
-by (REPEAT (etac conjE 1));
-by (etac ballE 1);
-by (eresolve_tac [restrict_eq_imp_val_eq RS sym] 1);
-by (fast_tac (claset() addSEs [ssubst]) 1);
-by (asm_full_simp_tac (simpset()
-        addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
-qed "lemma3";
-
-Close_locale "imp_DC0";
-
-
-Goalw [DC_def, DC0_def] "DC(nat) ==> DC0";
-by (REPEAT (resolve_tac [allI, impI] 1));
-by (REPEAT (eresolve_tac [asm_rl, conjE, ex_in_domain RS exE, allE] 1));
-by (eresolve_tac [export lemma4 RSN (2, impE)] 1
-        THEN REPEAT (assume_tac 1));
-by (etac bexE 1);
-by (dresolve_tac [export simplify_recursion] 1
-        THEN REPEAT (assume_tac 1));
-by (res_inst_tac [("x","\\<lambda>n \\<in> nat. f`n`n")] bexI 1);
-by (rtac lam_type 2);
-by (eresolve_tac [[export lemma2 RS conjunct1, succI1] MRS apply_type] 2
-        THEN REPEAT (assume_tac 2));
-by (rtac ballI 1);
-by (forward_tac [export (nat_succI RSN (5,lemma2)) RS conjunct2] 1
-        THEN REPEAT (assume_tac 1));
-by (dresolve_tac [export lemma3] 1 THEN REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [nat_succI]) 1);
-qed "DC_nat_imp_DC0";
-
-(* ********************************************************************** *)
-(* \\<forall>K. Card(K) --> DC(K) ==> WO3                                       *)
-(* ********************************************************************** *)
-
-Goalw [lesspoll_def]
-        "[| ~ A lesspoll B; C lesspoll B |] ==> A - C \\<noteq> 0";
-by (fast_tac (claset() addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll]
-        addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1);
-qed "lesspoll_lemma";
-
-val [f_type, Ord_a, not_eq] = goalw thy [inj_def]
-        "[| f \\<in> a->X; Ord(a); (!!b c. [| b<c; c \\<in> a |] ==> f`b\\<noteq>f`c) |]   \
-\        ==> f \\<in> inj(a, X)";
-by (resolve_tac [f_type RS CollectI] 1);
-by (REPEAT (resolve_tac [ballI,impI] 1));
-by (resolve_tac [Ord_a RS Ord_in_Ord RS Ord_linear_lt] 1
-        THEN (assume_tac 1));
-by (eres_inst_tac [("j","x")] (Ord_a RS Ord_in_Ord) 1);
-by (REPEAT (fast_tac (claset() addDs [not_eq, not_eq RS not_sym]) 1));
-qed "fun_Ord_inj";
-
-Goal "[| f \\<in> X->Y; A \\<subseteq> X; a \\<in> A |] ==> f`a \\<in> f``A";
-by (fast_tac (claset() addSEs [image_fun RS ssubst]) 1);
-qed "value_in_image";
-
-Goalw [DC_def, WO3_def] "\\<forall>K. Card(K) --> DC(K) ==> WO3";
-by (rtac allI 1);
-by (excluded_middle_tac "A lesspoll Hartog(A)" 1);
-by (fast_tac (claset() addSDs [lesspoll_imp_ex_lt_eqpoll]
-        addSIs [Ord_Hartog, leI RS le_imp_subset]) 2);
-by (REPEAT (eresolve_tac [allE, impE] 1));
-by (rtac Card_Hartog 1);
-by (eres_inst_tac [("x","A")] allE 1);
-by (eres_inst_tac [("x","{<z1,z2>:Pow(A)*A . z1  \
-\               lesspoll Hartog(A) & z2 \\<notin> z1}")] allE 1);
-by (Asm_full_simp_tac 1);
-by (etac impE 1);
-by (fast_tac (claset() addEs [lesspoll_lemma RS not_emptyE]) 1);
-by (etac bexE 1);
-by (resolve_tac [exI RS (lepoll_def RS (def_imp_iff RS iffD2))
-        RS (HartogI RS notE)] 1);
-by (resolve_tac [Ord_Hartog RSN (2, fun_Ord_inj)] 1 THEN (assume_tac 1));
-by (dresolve_tac [Ord_Hartog RSN (2, OrdmemD) RSN (2,
-        ltD RSN (3, value_in_image))] 1 
-        THEN REPEAT (assume_tac 1));
-by (force_tac (claset() addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)], 
-	       simpset()) 1);
-qed "DC_WO3";
-
-(* ********************************************************************** *)
-(* WO1 ==> \\<forall>K. Card(K) --> DC(K)                                       *)
-(* ********************************************************************** *)
-
-Goal "[| Ord(a); b \\<in> a |] ==> (\\<lambda>x \\<in> a. P(x))``b = (\\<lambda>x \\<in> b. P(x))``b";
-by (rtac images_eq 1);
-by (REPEAT (fast_tac (claset() addSEs [RepFunI, OrdmemD]
-        addSIs [lam_type]) 2));
-by (rtac ballI 1);
-by (dresolve_tac [OrdmemD RS subsetD] 1
-        THEN REPEAT (assume_tac 1));
-by (Asm_full_simp_tac 1);
-qed "lam_images_eq";
-
-Goalw [lesspoll_def] "[| Card(K); b \\<in> K |] ==> b lesspoll K";
-by (asm_full_simp_tac (simpset() addsimps [Card_iff_initial]) 1);
-by (fast_tac (claset() addSIs [le_imp_lepoll, ltI, leI]) 1);
-qed "in_Card_imp_lesspoll";
-
-Goal "(\\<lambda>b \\<in> a. P(b)) \\<in> a -> {P(b). b \\<in> a}";
-by (fast_tac (claset() addSIs [lam_type, RepFunI]) 1);
-qed "lam_type_RepFun";
-
-Goal "[| \\<forall>Y \\<in> Pow(X). Y lesspoll K --> (\\<exists>x \\<in> X. <Y, x> \\<in> R);  \
-\        b \\<in> K; Z \\<in> Pow(X); Z lesspoll K |]  \
-\     ==> {x \\<in> X. <Z,x> \\<in> R} \\<noteq> 0";
-by (Blast_tac 1);
-qed "lemmaX";
-
-Goal "[| Card(K); well_ord(X,Q);  \
-\       \\<forall>Y \\<in> Pow(X). Y lesspoll K --> (\\<exists>x \\<in> X. <Y, x> \\<in> R); b \\<in> K |]  \
-\       ==> ff(b, X, Q, R) \\<in> {x \\<in> X. <(\\<lambda>c \\<in> b. ff(c, X, Q, R))``b, x> \\<in> R}";
-by (res_inst_tac [("P","b \\<in> K")] impE 1 THEN TRYALL assume_tac);
-by (res_inst_tac [("i","b")] (Card_is_Ord RS Ord_in_Ord RS trans_induct) 1
-        THEN REPEAT (assume_tac 1));
-by (rtac impI 1);
-by (resolve_tac [ff_def RS def_transrec RS ssubst] 1);
-by (etac the_first_in 1);
-by (Fast_tac 1);
-by (asm_full_simp_tac (simpset()
-        addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1);
-by (etac lemmaX 1 THEN assume_tac 1);
-by (blast_tac (claset() addIs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1);
-by (blast_tac (claset() addIs [lesspoll_trans1, in_Card_imp_lesspoll, 
-                               RepFun_lepoll]) 1); 
-qed "lemma";
-
-Goalw [DC_def, WO1_def] "WO1 ==> \\<forall>K. Card(K) --> DC(K)";
-by (REPEAT (resolve_tac [allI,impI] 1));
-by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
-by (res_inst_tac [("x","\\<lambda>b \\<in> K. ff(b, X, Ra, R)")] bexI 1);
-by (rtac lam_type 2);
-by (resolve_tac [lemma RS CollectD1] 2 THEN REPEAT (assume_tac 2));
-by (asm_full_simp_tac (simpset()
-        addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1);
-by (fast_tac (claset() addSEs [ltE, lemma RS CollectD2]) 1);
-qed "WO1_DC_Card";
--- a/src/ZF/AC/DC.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/DC.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -2,69 +2,578 @@
     ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-Theory file for the proofs concernind the Axiom of Dependent Choice
+The proofs concerning the Axiom of Dependent Choice
 *)
 
-DC  =  AC_Equiv + Hartog + Cardinal_aux + DC_lemmas + 
+theory DC = AC_Equiv + Hartog + Cardinal_aux:
+
+lemma RepFun_lepoll: "Ord(a) ==> {P(b). b \<in> a} \<lesssim> a"
+apply (unfold lepoll_def)
+apply (rule_tac x = "\<lambda>z \<in> RepFun (a,P) . LEAST i. z=P (i) " in exI)
+apply (rule_tac d="%z. P (z)" in lam_injective)
+ apply (fast intro!: Least_in_Ord)
+apply (rule sym) 
+apply (fast intro: LeastI Ord_in_Ord) 
+done
 
-consts
+text{*Trivial in the presence of AC, but here we need a wellordering of X*}
+lemma image_Ord_lepoll: "[| f \<in> X->Y; Ord(X) |] ==> f``X \<lesssim> X"
+apply (unfold lepoll_def)
+apply (rule_tac x = "\<lambda>x \<in> f``X. LEAST y. f`y = x" in exI)
+apply (rule_tac d = "%z. f`z" in lam_injective)
+apply (fast intro!: Least_in_Ord apply_equality, clarify) 
+apply (rule LeastI) 
+ apply (erule apply_equality, assumption+) 
+apply (blast intro: Ord_in_Ord)
+done
 
-  DC  :: i => o
-  DC0 :: o
-  ff  :: [i, i, i, i] => i
+lemma range_subset_domain: 
+      "[| R \<subseteq> X*X;   !!g. g \<in> X ==> \<exists>u. <g,u> \<in> R |] 
+       ==> range(R) \<subseteq> domain(R)"
+by (blast ); 
+
+lemma cons_fun_type: "g \<in> n->X ==> cons(<n,x>, g) \<in> succ(n) -> cons(x, X)"
+apply (unfold succ_def)
+apply (fast intro!: fun_extend elim!: mem_irrefl)
+done
 
-rules
+lemma cons_fun_type2:
+     "[| g \<in> n->X; x \<in> X |] ==> cons(<n,x>, g) \<in> succ(n) -> X"
+by (erule cons_absorb [THEN subst], erule cons_fun_type)
+
+lemma cons_image_n: "n \<in> nat ==> cons(<n,x>, g)``n = g``n"
+by (fast elim!: mem_irrefl)
+
+lemma cons_val_n: "g \<in> n->X ==> cons(<n,x>, g)`n = x"
+by (fast intro!: apply_equality elim!: cons_fun_type)
+
+lemma cons_image_k: "k \<in> n ==> cons(<n,x>, g)``k = g``k"
+by (fast elim: mem_asym)
 
-  DC_def  "DC(a) ==
-	     \\<forall>X R. R \\<subseteq> Pow(X)*X &
-             (\\<forall>Y \\<in> Pow(X). Y lesspoll a --> (\\<exists>x \\<in> X. <Y,x> \\<in> R)) 
-             --> (\\<exists>f \\<in> a->X. \\<forall>b<a. <f``b,f`b> \\<in> R)"
+lemma cons_val_k: "[| k \<in> n; g \<in> n->X |] ==> cons(<n,x>, g)`k = g`k"
+by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair)
+
+lemma domain_cons_eq_succ: "domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)"
+by (simp add: domain_cons succ_def)
+
+lemma restrict_cons_eq: "g \<in> n->X ==> restrict(cons(<n,x>, g), n) = g"
+apply (unfold restrict_def)
+apply (rule fun_extension)
+apply (rule lam_type)
+apply (erule cons_fun_type [THEN apply_type])
+apply (erule succI2, assumption)
+apply (simp add: cons_val_k)
+done
+
+lemma succ_in_succ: "[| Ord(k); i \<in> k |] ==> succ(i) \<in> succ(k)"
+apply (rule Ord_linear [of "succ(i)" "succ(k)", THEN disjE])
+apply (fast elim: Ord_in_Ord mem_irrefl mem_asym)+
+done
 
-  DC0_def "DC0 == \\<forall>A B R. R \\<subseteq> A*B & R\\<noteq>0 & range(R) \\<subseteq> domain(R) 
-                  --> (\\<exists>f \\<in> nat->domain(R). \\<forall>n \\<in> nat. <f`n,f`succ(n)>:R)"
+lemma restrict_eq_imp_val_eq: 
+      "[| restrict(f, domain(g)) = g; x \<in> domain(g) |] ==> f`x = g`x"
+apply (unfold restrict_def) 
+apply (erule subst, simp)
+done
+
+lemma domain_eq_imp_fun_type: "[| domain(f)=A; f \<in> B->C |] ==> f \<in> A->C"
+by (frule domain_of_fun, fast)
+
+lemma ex_in_domain: "[| R \<subseteq> A * B; R \<noteq> 0 |] ==> \<exists>x. x \<in> domain(R)"
+by (fast elim!: not_emptyE)
+
 
-  ff_def  "ff(b, X, Q, R) ==
-	   transrec(b, %c r. THE x. first(x, {x \\<in> X. <r``c, x> \\<in> R}, Q))"
-  
+constdefs
+
+  DC  :: "i => o"
+    "DC(a) == \<forall>X R. R \<subseteq> Pow(X)*X  &
+		    (\<forall>Y \<in> Pow(X). Y \<prec> a --> (\<exists>x \<in> X. <Y,x> \<in> R)) 
+		    --> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
+
+  DC0 :: o
+    "DC0 == \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R) 
+                    --> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)"
+
+  ff  :: "[i, i, i, i] => i"
+    "ff(b, X, Q, R) ==
+	   transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
+
 
 locale DC0_imp =
-  fixes 
-    XX	:: i
-    RR	:: i
-    X	:: i
-    R	:: i
+  fixes XX and RR and X and R
+
+  assumes all_ex: "\<forall>Y \<in> Pow(X). Y \<prec> nat --> (\<exists>x \<in> X. <Y, x> \<in> R)"
+
+  defines XX_def: "XX == (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})"
+     and RR_def:  "RR == {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  
+                                       & restrict(z2, domain(z1)) = z1}"
+
+
+(* ********************************************************************** *)
+(* DC ==> DC(omega)                                                       *)
+(*                                                                        *)
+(* The scheme of the proof:                                               *)
+(*                                                                        *)
+(* Assume DC. Let R and X satisfy the premise of DC(omega).               *)
+(*                                                                        *)
+(* Define XX and RR as follows:                                           *)
+(*                                                                        *)
+(*       XX = (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})           *)
+(*       f RR g iff domain(g)=succ(domain(f)) &                           *)
+(*              restrict(g, domain(f)) = f                                *)
+(*                                                                        *)
+(* Then RR satisfies the hypotheses of DC.                                *)
+(* So applying DC:                                                        *)
+(*                                                                        *)
+(*       \<exists>f \<in> nat->XX. \<forall>n \<in> nat. f`n RR f`succ(n)                        *)
+(*                                                                        *)
+(* Thence                                                                 *)
+(*                                                                        *)
+(*       ff = {<n, f`succ(n)`n>. n \<in> nat}                                 *)
+(*                                                                        *)
+(* is the desired function.                                               *)
+(*                                                                        *)
+(* ********************************************************************** *)
+
+lemma (in DC0_imp) lemma1_1: "RR \<subseteq> XX*XX"
+by (unfold RR_def, fast)
+
+lemma (in DC0_imp) lemma1_2: "RR \<noteq> 0"
+apply (unfold RR_def XX_def)
+apply (rule all_ex [THEN ballE])
+apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]])
+apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]])
+apply (erule bexE)
+apply (rule_tac a = "<0, {<0, x>}>" in not_emptyI)
+apply (rule CollectI)
+apply (rule SigmaI)
+apply (rule nat_0I [THEN UN_I])
+apply (simp (no_asm_simp) add: nat_0I [THEN UN_I])
+apply (rule nat_1I [THEN UN_I])
+apply (force intro!: singleton_fun [THEN Pi_type]
+             simp add: singleton_0 [symmetric])
+apply (simp add: singleton_0)
+done
+
+lemma (in DC0_imp) lemma1_3: "range(RR) \<subseteq> domain(RR)"
+apply (unfold RR_def XX_def)
+apply (rule range_subset_domain, blast, clarify)
+apply (frule fun_is_rel [THEN image_subset, THEN PowI, 
+                         THEN all_ex [THEN bspec]])
+apply (erule impE[OF _ lesspoll_trans1[OF image_Ord_lepoll 
+                                          [OF _ nat_into_Ord] n_lesspoll_nat]],
+       assumption+)
+apply (erule bexE)
+apply (rule_tac x = "cons (<n,x>, g) " in exI)
+apply (rule CollectI)
+apply (force elim!: cons_fun_type2 
+             simp add: cons_image_n cons_val_n cons_image_k cons_val_k)
+apply (simp add: domain_of_fun succ_def restrict_cons_eq)
+done
 
-  assumes
-    all_ex    "\\<forall>Y \\<in> Pow(X). Y lesspoll nat --> (\\<exists>x \\<in> X. <Y, x> \\<in> R)"
+lemma (in DC0_imp) lemma2:
+     "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR;  f \<in> nat -> XX;  n \<in> nat |]   
+      ==> \<exists>k \<in> nat. f`succ(n) \<in> k -> X & n \<in> k   
+                  & <f`succ(n)``n, f`succ(n)`n> \<in> R"
+apply (induct_tac "n")
+apply (drule apply_type [OF _ nat_1I])
+apply (drule bspec [OF _ nat_0I])
+apply (simp add: XX_def, safe)
+apply (rule rev_bexI, assumption)
+apply (subgoal_tac "0 \<in> x", force)
+apply (force simp add: RR_def
+	     intro: ltD elim!: nat_0_le [THEN leE])
+(** LEVEL 7, other subgoal **)
+apply (drule bspec [OF _ nat_succI], assumption)
+apply (subgoal_tac "f ` succ (succ (x)) \<in> succ (k) ->X")
+apply (drule apply_type [OF _ nat_succI [THEN nat_succI]], assumption)
+apply (simp (no_asm_use) add: XX_def RR_def)
+apply safe
+apply (frule_tac a="succ(k)" in domain_of_fun [symmetric, THEN trans], 
+       assumption)
+apply (frule_tac a="xa" in domain_of_fun [symmetric, THEN trans], 
+       assumption)
+apply (fast elim!: nat_into_Ord [THEN succ_in_succ] 
+            dest!: bspec [OF _ nat_into_Ord [THEN succ_in_succ]])
+apply (drule domain_of_fun)
+apply (simp add: XX_def RR_def, clarify) 
+apply (blast dest: domain_of_fun [symmetric, THEN trans] )
+done
+
+lemma (in DC0_imp) lemma3_1:
+     "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR;  f \<in> nat -> XX;  m \<in> nat |]   
+      ==>  {f`succ(x)`x. x \<in> m} = {f`succ(m)`x. x \<in> m}"
+apply (subgoal_tac "\<forall>x \<in> m. f`succ (m) `x = f`succ (x) `x")
+apply simp
+apply (induct_tac "m", blast)
+apply (rule ballI)
+apply (erule succE)
+ apply (rule restrict_eq_imp_val_eq)
+  apply (drule bspec [OF _ nat_succI], assumption)
+  apply (simp add: RR_def)
+ apply (drule lemma2, assumption+)
+ apply (fast dest!: domain_of_fun)
+apply (drule_tac x = "xa" in bspec, assumption)
+apply (erule sym [THEN trans, symmetric])
+apply (rule restrict_eq_imp_val_eq [symmetric])
+ apply (drule bspec [OF _ nat_succI], assumption)
+ apply (simp add: RR_def)
+apply (drule lemma2, assumption+)
+apply (blast dest!: domain_of_fun 
+             intro: nat_into_Ord OrdmemD [THEN subsetD])
+done
 
-  defines
-    XX_def    "XX == (\\<Union>n \\<in> nat. {f \\<in> n->X. \\<forall>k \\<in> n. <f``k, f`k> \\<in> R})"
-    RR_def    "RR == {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  
-                                  & restrict(z2, domain(z1)) = z1}"
+lemma (in DC0_imp) lemma3:
+     "[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR;  f \<in> nat -> XX;  m \<in> nat |]  
+      ==> (\<lambda>x \<in> nat. f`succ(x)`x) `` m = f`succ(m)``m"
+apply (erule natE, simp)
+apply (subst image_lam)
+ apply (fast elim!: OrdmemD [OF nat_succI Ord_nat])
+apply (subst lemma3_1, assumption+)
+ apply fast
+apply (fast dest!: lemma2 
+            elim!: image_fun [symmetric, OF _ OrdmemD [OF _ nat_into_Ord]])
+done
+
+
+theorem DC0_imp_DC_nat: "DC0 ==> DC(nat)"
+apply (unfold DC_def DC0_def, clarify)
+apply (elim allE)
+apply (erule impE)
+   (*these three results comprise Lemma 1*)
+apply (blast intro!: DC0_imp.lemma1_1 DC0_imp.lemma1_2 DC0_imp.lemma1_3)
+apply (erule bexE)
+apply (rule_tac x = "\<lambda>n \<in> nat. f`succ (n) `n" in rev_bexI)
+ apply (rule lam_type, blast dest!: DC0_imp.lemma2 intro: fun_weaken_type)
+apply (rule oallI)
+apply (frule DC0_imp.lemma2, assumption)
+  apply (blast intro: fun_weaken_type)
+ apply (erule ltD) 
+(** LEVEL 11: last subgoal **)
+apply (subst DC0_imp.lemma3, assumption+) 
+  apply (fast elim!: fun_weaken_type)
+ apply (erule ltD, force) 
+done
+
+
+(* ************************************************************************
+   DC(omega) ==> DC                                                       
+                                                                          
+   The scheme of the proof:                                               
+                                                                          
+   Assume DC(omega). Let R and x satisfy the premise of DC.               
+                                                                          
+   Define XX and RR as follows:                                           
+                                                                          
+    XX = (\<Union>n \<in> nat. {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})
+
+    RR = {<z1,z2>:Fin(XX)*XX. 
+           (domain(z2)=succ(\<Union>f \<in> z1. domain(f)) &
+	    (\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |      
+	   (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) &
+	                (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) &           
+	    z2={<0,x>})}                                          
+                                                                          
+   Then XX and RR satisfy the hypotheses of DC(omega).                    
+   So applying DC:                                                        
+                                                                          
+         \<exists>f \<in> nat->XX. \<forall>n \<in> nat. f``n RR f`n                             
+                                                                          
+   Thence                                                                 
+                                                                          
+         ff = {<n, f`n`n>. n \<in> nat}                                         
+                                                                          
+   is the desired function.                                               
+                                                                          
+************************************************************************* *)
+
+lemma singleton_in_funs: 
+ "x \<in> X ==> {<0,x>} \<in> 
+            (\<Union>n \<in> nat. {f \<in> succ(n)->X. \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
+apply (rule nat_0I [THEN UN_I])
+apply (force simp add: singleton_0 [symmetric]
+	     intro!: singleton_fun [THEN Pi_type])
+done
 
 
 locale imp_DC0 =
-  fixes 
-    XX	:: i
-    RR	:: i
-    x	:: i
-    R	:: i
-    f	:: i
-    allRR :: o
+  fixes XX and RR and x and R and f and allRR
+  defines XX_def: "XX == (\<Union>n \<in> nat.
+		      {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
+      and RR_def:
+	 "RR == {<z1,z2>:Fin(XX)*XX. 
+		  (domain(z2)=succ(\<Union>f \<in> z1. domain(f))  
+		    & (\<forall>f \<in> z1. restrict(z2, domain(f)) = f))
+		  | (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))  
+		     & (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
+      and allRR_def:
+	"allRR == \<forall>b<nat.
+		   <f``b, f`b> \<in>  
+		    {<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
+				    & (\<Union>f \<in> z1. domain(f)) = b  
+				    & (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
+
+lemma (in imp_DC0) lemma4:
+     "[| range(R) \<subseteq> domain(R);  x \<in> domain(R) |]   
+      ==> RR \<subseteq> Pow(XX)*XX &   
+             (\<forall>Y \<in> Pow(XX). Y \<prec> nat --> (\<exists>x \<in> XX. <Y,x>:RR))"
+apply (rule conjI)
+apply (force dest!: FinD [THEN PowI] simp add: RR_def)
+apply (rule impI [THEN ballI])
+apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption)
+apply (case_tac
+       "\<exists>g \<in> XX. domain (g) =
+             succ(\<Union>f \<in> Y. domain(f)) & (\<forall>f\<in>Y. restrict(g, domain(f)) = f)")
+apply (simp add: RR_def, blast)
+apply (safe del: domainE)
+apply (unfold XX_def RR_def)
+apply (rule rev_bexI, erule singleton_in_funs)
+apply (simp add: nat_0I [THEN rev_bexI] cons_fun_type2)
+done
+
+lemma (in imp_DC0) UN_image_succ_eq:
+     "[| f \<in> nat->X; n \<in> nat |] 
+      ==> (\<Union>x \<in> f``succ(n). P(x)) =  P(f`n) Un (\<Union>x \<in> f``n. P(x))"
+by (simp add: image_fun OrdmemD) 
+
+lemma (in imp_DC0) UN_image_succ_eq_succ:
+     "[| (\<Union>x \<in> f``n. P(x)) = y; P(f`n) = succ(y);   
+         f \<in> nat -> X; n \<in> nat |] ==> (\<Union>x \<in> f``succ(n). P(x)) = succ(y)"
+by (simp add: UN_image_succ_eq, blast)
+
+lemma (in imp_DC0) apply_domain_type:
+     "[| h \<in> succ(n) -> D;  n \<in> nat; domain(h)=succ(y) |] ==> h`y \<in> D"
+by (fast elim: apply_type dest!: trans [OF sym domain_of_fun])
+
+lemma (in imp_DC0) image_fun_succ:
+     "[| h \<in> nat -> X; n \<in> nat |] ==> h``succ(n) = cons(h`n, h``n)"
+by (simp add: image_fun OrdmemD) 
+
+lemma (in imp_DC0) f_n_type:
+     "[| domain(f`n) = succ(k); f \<in> nat -> XX;  n \<in> nat |]    
+      ==> f`n \<in> succ(k) -> domain(R)"
+apply (unfold XX_def)
+apply (drule apply_type, assumption)
+apply (fast elim: domain_eq_imp_fun_type)
+done
+
+lemma (in imp_DC0) f_n_pairs_in_R [rule_format]: 
+     "[| h \<in> nat -> XX;  domain(h`n) = succ(k);  n \<in> nat |]   
+      ==> \<forall>i \<in> k. <h`n`i, h`n`succ(i)> \<in> R"
+apply (unfold XX_def)
+apply (drule apply_type, assumption)
+apply (elim UN_E CollectE)
+apply (drule domain_of_fun [symmetric, THEN trans], assumption)
+apply simp
+done
+
+lemma (in imp_DC0) restrict_cons_eq_restrict: 
+     "[| restrict(h, domain(u))=u;  h \<in> n->X;  domain(u) \<subseteq> n |]   
+      ==> restrict(cons(<n, y>, h), domain(u)) = u"
+apply (unfold restrict_def)
+apply (erule sym [THEN trans, symmetric])
+apply (rule fun_extension)
+apply (fast intro!: lam_type)
+apply (fast intro!: lam_type)
+apply (simp add: subsetD [THEN cons_val_k])
+done
+
+lemma (in imp_DC0) all_in_image_restrict_eq:
+     "[| \<forall>x \<in> f``n. restrict(f`n, domain(x))=x;   
+         f \<in> nat -> XX;   
+         n \<in> nat;  domain(f`n) = succ(n);   
+         (\<Union>x \<in> f``n. domain(x)) \<subseteq> n |]  
+      ==> \<forall>x \<in> f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x)) = x"
+apply (rule ballI)
+apply (simp add: image_fun_succ)
+apply (drule f_n_type, assumption+)
+apply (erule disjE)
+ apply (simp add: domain_of_fun restrict_cons_eq) 
+apply (blast intro!: restrict_cons_eq_restrict)
+done
+
+lemma (in imp_DC0) simplify_recursion: 
+     "[| \<forall>b<nat. <f``b, f`b> \<in> RR;   
+         f \<in> nat -> XX; range(R) \<subseteq> domain(R); x \<in> domain(R)|]    
+      ==> allRR"
+apply (unfold RR_def allRR_def)
+apply (rule oallI, drule ltD)
+apply (erule nat_induct)
+apply (drule_tac x="0" in ospec, blast intro: Limit_has_0) 
+apply (force simp add: singleton_fun [THEN domain_of_fun] singleton_in_funs) 
+(*induction step*) (** LEVEL 5 **)
+(*prevent simplification of ~\<exists> to \<forall>~ *)
+apply (simp only: separation split)
+apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI);
+apply (elim conjE disjE)
+apply (force elim!: trans subst_context
+             intro!: UN_image_succ_eq_succ)
+apply (erule notE)
+apply (simp add: XX_def UN_image_succ_eq_succ)
+apply (elim conjE bexE)
+apply (drule apply_domain_type, assumption+)
+apply (erule domainE)+
+apply (frule f_n_type)
+apply (simp add: XX_def, assumption+)
+apply (rule rev_bexI, erule nat_succI)
+apply (rule_tac x = "cons (<succ (xa), ya>, f`xa) " in bexI)
+prefer 2 apply (blast intro: cons_fun_type2) 
+apply (rule conjI)
+prefer 2 apply (fast del: ballI subsetI
+		 elim: trans [OF _ subst_context, THEN domain_cons_eq_succ]
+		       subst_context
+		       all_in_image_restrict_eq [simplified XX_def]
+		       trans equalityD1)
+(*one remaining subgoal*)
+apply (rule ballI)
+apply (erule succE)
+(** LEVEL 25 **)
+ apply (simp add: cons_val_n cons_val_k)
+(*assumption+ will not perform the required backtracking!*)
+apply (drule f_n_pairs_in_R [simplified XX_def, OF _ domain_of_fun], 
+       assumption, assumption, assumption)
+apply (simp add: nat_into_Ord [THEN succ_in_succ] succI2 cons_val_k)
+done
+
 
-  defines
-    XX_def    "XX == (\\<Union>n \\<in> nat.
-		      {f \\<in> succ(n)->domain(R). \\<forall>k \\<in> n. <f`k, f`succ(k)> \\<in> R})"
-    RR_def
-     "RR == {<z1,z2>:Fin(XX)*XX. 
-              (domain(z2)=succ(\\<Union>f \\<in> z1. domain(f))  
-                & (\\<forall>f \\<in> z1. restrict(z2, domain(f)) = f))
-	      | (~ (\\<exists>g \\<in> XX. domain(g)=succ(\\<Union>f \\<in> z1. domain(f))  
-                 & (\\<forall>f \\<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}"
-    allRR_def
-     "allRR == \\<forall>b<nat.
-                <f``b, f`b> \\<in>  
-                 {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(\\<Union>f \\<in> z1. domain(f))  
-                                    & (\\<Union>f \\<in> z1. domain(f)) = b  
-                                    & (\\<forall>f \\<in> z1. restrict(z2,domain(f)) = f))}"
+lemma (in imp_DC0) lemma2: 
+     "[| allRR; f \<in> nat->XX; range(R) \<subseteq> domain(R); x \<in> domain(R); n \<in> nat |]
+      ==> f`n \<in> succ(n) -> domain(R) & (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)"
+apply (unfold allRR_def)
+apply (drule ospec)
+apply (erule ltI [OF _ Ord_nat])
+apply (erule CollectE, simp)
+apply (rule conjI)
+prefer 2 apply (fast elim!: f_n_pairs_in_R trans subst_context)
+apply (unfold XX_def)
+apply (fast elim!: trans [THEN domain_eq_imp_fun_type] subst_context)
+done
+
+lemma (in imp_DC0) lemma3:
+     "[| allRR; f \<in> nat->XX; n\<in>nat; range(R) \<subseteq> domain(R);  x \<in> domain(R) |]
+      ==> f`n`n = f`succ(n)`n"
+apply (frule lemma2 [THEN conjunct1, THEN domain_of_fun], assumption+)
+apply (unfold allRR_def)
+apply (drule ospec) 
+apply (drule ltI [OF nat_succI Ord_nat], assumption)
+apply simp
+apply (elim conjE ballE)
+apply (erule restrict_eq_imp_val_eq [symmetric], force) 
+apply (simp add: image_fun OrdmemD) 
+done
+
+
+theorem DC_nat_imp_DC0: "DC(nat) ==> DC0"
+apply (unfold DC_def DC0_def)
+apply (intro allI impI)
+apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+
+apply (erule impE [OF _ imp_DC0.lemma4], assumption+)
+apply (erule bexE)
+apply (drule imp_DC0.simplify_recursion, assumption+)
+apply (rule_tac x = "\<lambda>n \<in> nat. f`n`n" in bexI)
+apply (rule_tac [2] lam_type)
+apply (erule_tac [2] apply_type [OF imp_DC0.lemma2 [THEN conjunct1] succI1])
+apply (rule ballI)
+apply (frule_tac n="succ(n)" in imp_DC0.lemma2, 
+       (assumption|erule nat_succI)+)
+apply (drule imp_DC0.lemma3, auto)
+done
+
+(* ********************************************************************** *)
+(* \<forall>K. Card(K) --> DC(K) ==> WO3                                       *)
+(* ********************************************************************** *)
+
+lemma fun_Ord_inj:
+      "[| f \<in> a->X;  Ord(a); 
+          !!b c. [| b<c; c \<in> a |] ==> f`b\<noteq>f`c |]    
+       ==> f \<in> inj(a, X)"
+apply (unfold inj_def, simp) 
+apply (intro ballI impI)
+apply (rule_tac j=x in Ord_in_Ord [THEN Ord_linear_lt], assumption+)
+apply (blast intro: Ord_in_Ord, auto) 
+apply (atomize, blast dest: not_sym) 
+done
+
+lemma value_in_image: "[| f \<in> X->Y; A \<subseteq> X; a \<in> A |] ==> f`a \<in> f``A"
+by (fast elim!: image_fun [THEN ssubst]);
+
+theorem DC_WO3: "(\<forall>K. Card(K) --> DC(K)) ==> WO3"
+apply (unfold DC_def WO3_def)
+apply (rule allI)
+apply (case_tac "A \<prec> Hartog (A)");
+apply (fast dest!: lesspoll_imp_ex_lt_eqpoll 
+            intro!: Ord_Hartog leI [THEN le_imp_subset])
+apply (erule allE impE)+
+apply (rule Card_Hartog)
+apply (erule_tac x = "A" in allE)
+apply (erule_tac x = "{<z1,z2> \<in> Pow (A) *A . z1 \<prec> Hartog (A) & z2 \<notin> z1}" 
+                 in allE)
+apply simp
+apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE])
+apply (erule bexE)
+apply (rule Hartog_lepoll_selfE) 
+apply (rule lepoll_def [THEN def_imp_iff, THEN iffD2])
+apply (rule exI, rule fun_Ord_inj, assumption, rule Ord_Hartog)
+apply (drule value_in_image) 
+apply (drule OrdmemD, rule Ord_Hartog, assumption+, erule ltD) 
+apply (drule ospec)
+apply (blast intro: ltI Ord_Hartog, force) 
+done
+
+(* ********************************************************************** *)
+(* WO1 ==> \<forall>K. Card(K) --> DC(K)                                       *)
+(* ********************************************************************** *)
+
+lemma images_eq:
+     "[| \<forall>x \<in> A. f`x=g`x; f \<in> Df->Cf; g \<in> Dg->Cg; A \<subseteq> Df; A \<subseteq> Dg |] 
+      ==> f``A = g``A"
+apply (simp (no_asm_simp) add: image_fun)
+done
+
+lemma lam_images_eq:
+     "[| Ord(a); b \<in> a |] ==> (\<lambda>x \<in> a. h(x))``b = (\<lambda>x \<in> b. h(x))``b"
+apply (rule images_eq)
+    apply (rule ballI)
+    apply (drule OrdmemD [THEN subsetD], assumption+)
+    apply simp
+   apply (fast elim!: RepFunI OrdmemD intro!: lam_type)+
+done
+
+lemma lam_type_RepFun: "(\<lambda>b \<in> a. h(b)) \<in> a -> {h(b). b \<in> a}"
+by (fast intro!: lam_type RepFunI)
+
+lemma lemmaX:
+     "[| \<forall>Y \<in> Pow(X). Y \<prec> K --> (\<exists>x \<in> X. <Y, x> \<in> R);   
+         b \<in> K; Z \<in> Pow(X); Z \<prec> K |]   
+      ==> {x \<in> X. <Z,x> \<in> R} \<noteq> 0"
+by blast
+
+
+lemma WO1_DC_lemma:
+     "[| Card(K); well_ord(X,Q);   
+         \<forall>Y \<in> Pow(X). Y \<prec> K --> (\<exists>x \<in> X. <Y, x> \<in> R); b \<in> K |]   
+      ==> ff(b, X, Q, R) \<in> {x \<in> X. <(\<lambda>c \<in> b. ff(c, X, Q, R))``b, x> \<in> R}"
+apply (rule_tac P = "b \<in> K" in impE, (erule_tac [2] asm_rl)+)
+apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], 
+       assumption+)
+apply (rule impI)
+apply (rule ff_def [THEN def_transrec, THEN ssubst])
+apply (erule the_first_in, fast)
+apply (simp add: image_fun [OF lam_type_RepFun subset_refl])
+apply (erule lemmaX, assumption)
+ apply (blast intro: Card_is_Ord OrdmemD [THEN subsetD])
+apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll)
+done
+
+theorem WO1_DC_Card: "WO1 ==> \<forall>K. Card(K) --> DC(K)"
+apply (unfold DC_def WO1_def)
+apply (rule allI impI)+
+apply (erule allE exE conjE)+
+apply (rule_tac x = "\<lambda>b \<in> K. ff (b, X, Ra, R) " in bexI)
+ apply (simp add: lam_images_eq [OF Card_is_Ord ltD])
+ apply (fast elim!: ltE WO1_DC_lemma [THEN CollectD2])
+apply (rule_tac lam_type)
+apply (rule WO1_DC_lemma [THEN CollectD1], assumption+)
+done
+
 end
--- a/src/ZF/AC/DC_lemmas.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-(*  Title:      ZF/AC/DC_lemmas.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-More general lemmas used in the proofs concerning DC
-
-*)
-
-val [prem] = goalw thy [lepoll_def]
-        "Ord(a) ==> {P(b). b \\<in> a} lepoll a";
-by (res_inst_tac [("x","\\<lambda>z \\<in> RepFun(a,P). LEAST i. z=P(i)")] exI 1);
-by (res_inst_tac [("d","%z. P(z)")] (sym RSN (2, lam_injective)) 1);
-by (fast_tac (claset() addSIs [Least_in_Ord, prem]) 1);
-by (REPEAT (eresolve_tac [RepFunE, LeastI, prem RS Ord_in_Ord] 1));
-qed "RepFun_lepoll";
-
-Goalw [lesspoll_def] "n \\<in> nat ==> n lesspoll nat";
-by (rtac conjI 1);
-by (eresolve_tac [Ord_nat RSN (2, OrdmemD) RS subset_imp_lepoll] 1);
-by (rtac notI 1);
-by (etac eqpollE 1);
-by (rtac succ_lepoll_natE 1 THEN (assume_tac 2));
-by (eresolve_tac [nat_succI RS (Ord_nat RSN (2, OrdmemD) RS
-        subset_imp_lepoll) RS lepoll_trans] 1
-        THEN (assume_tac 1));
-qed "n_lesspoll_nat";
-
-Goalw [lepoll_def]
-        "[| f \\<in> X->Y; Ord(X) |] ==> f``X lepoll X";
-by (res_inst_tac [("x","\\<lambda>x \\<in> f``X. LEAST y. f`y = x")] exI 1);
-by (res_inst_tac [("d","%z. f`z")] lam_injective 1);
-by (fast_tac (claset() addSIs [Least_in_Ord, apply_equality]) 1);
-by (fast_tac (claset() addSEs [Ord_in_Ord] addSIs [LeastI, apply_equality]) 1);
-qed "image_Ord_lepoll";
-
-val [major, minor] = goal thy
-        "[| (!!g. g \\<in> X ==> \\<exists>u. <g,u>:R); R \\<subseteq> X*X  \
-\       |] ==> range(R) \\<subseteq> domain(R)";
-by (rtac subsetI 1);
-by (etac rangeE 1);
-by (dresolve_tac [minor RS subsetD RS SigmaD2 RS major] 1);
-by (Fast_tac 1);
-qed "range_subset_domain";
-
-val prems = goal thy "!!k. k \\<in> n ==> k\\<noteq>n";
-by (fast_tac (claset() addSEs [mem_irrefl]) 1);
-qed "mem_not_eq";
-
-Goalw [succ_def] "g \\<in> n->X ==> cons(<n,x>, g) \\<in> succ(n) -> cons(x, X)";
-by (fast_tac (claset() addSIs [fun_extend] addSEs [mem_irrefl]) 1);
-qed "cons_fun_type";
-
-Goal "[| g \\<in> n->X; x \\<in> X |] ==> cons(<n,x>, g) \\<in> succ(n) -> X";
-by (etac (cons_absorb RS subst) 1 THEN etac cons_fun_type 1);
-qed "cons_fun_type2";
-
-Goal "n \\<in> nat ==> cons(<n,x>, g)``n = g``n";
-by (fast_tac (claset() addSEs [mem_irrefl]) 1);
-qed "cons_image_n";
-
-Goal "g \\<in> n->X ==> cons(<n,x>, g)`n = x";
-by (fast_tac (claset() addSIs [apply_equality] addSEs [cons_fun_type]) 1);
-qed "cons_val_n";
-
-Goal "k \\<in> n ==> cons(<n,x>, g)``k = g``k";
-by (fast_tac (claset() addEs [mem_asym]) 1);
-qed "cons_image_k";
-
-Goal "[| k \\<in> n; g \\<in> n->X |] ==> cons(<n,x>, g)`k = g`k";
-by (fast_tac (claset() addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1);
-qed "cons_val_k";
-
-Goal "domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
-by (asm_full_simp_tac (simpset() addsimps [domain_cons, succ_def]) 1);
-qed "domain_cons_eq_succ";
-
-Goalw [restrict_def] "g \\<in> n->X ==> restrict(cons(<n,x>, g), n)=g";
-by (rtac fun_extension 1);
-by (rtac lam_type 1);
-by (eresolve_tac [cons_fun_type RS apply_type] 1);
-by (etac succI2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [cons_val_k]) 1);
-qed "restrict_cons_eq";
-
-Goal "[| Ord(k); i \\<in> k |] ==> succ(i) \\<in> succ(k)";
-by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3));
-by (REPEAT (fast_tac (claset() addEs [Ord_in_Ord, mem_irrefl, mem_asym]) 1));
-qed "succ_in_succ";
-
-Goalw [restrict_def]
-        "[| restrict(f, domain(g)) = g; x \\<in> domain(g) |] ==> f`x = g`x";
-by (etac subst 1);
-by (Asm_full_simp_tac 1);
-qed "restrict_eq_imp_val_eq";
-
-Goal "[| domain(f)=A; f \\<in> B->C |] ==> f \\<in> A->C";
-by (ftac domain_of_fun 1);
-by (Fast_tac 1);
-qed "domain_eq_imp_fun_type";
-
-Goal "[| R \\<subseteq> A * B; R \\<noteq> 0 |] ==> \\<exists>x. x \\<in> domain(R)";
-by (fast_tac (claset() addSEs [not_emptyE]) 1);
-qed "ex_in_domain";
-
--- a/src/ZF/AC/DC_lemmas.thy	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-(*Dummy theory to document dependencies *)
-
-DC_lemmas = AC_Equiv + Cardinal_aux
--- a/src/ZF/AC/HH.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,212 +0,0 @@
-(*  Title:      ZF/AC/HH.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Some properties of the recursive definition of HH used in the proofs of
-  AC17 ==> AC1
-  AC1 ==> WO2
-  AC15 ==> WO6
-*)
-
-(* ********************************************************************** *)
-(* Lemmas useful in each of the three proofs                              *)
-(* ********************************************************************** *)
-
-Goal "HH(f,x,a) =  \
-\       (let z = x - (\\<Union>b \\<in> a. HH(f,x,b))  \
-\       in  if(f`z \\<in> Pow(z)-{0}, f`z, {x}))";
-by (resolve_tac [HH_def RS def_transrec RS trans] 1);
-by (Simp_tac 1);
-qed "HH_def_satisfies_eq";
-
-Goal "HH(f,x,a) \\<in> Pow(x)-{0} | HH(f,x,a)={x}";
-by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
-by (Fast_tac 1);
-qed "HH_values";
-
-Goal "B \\<subseteq> A ==> X-(\\<Union>a \\<in> A. P(a)) = X-(\\<Union>a \\<in> A-B. P(a))-(\\<Union>b \\<in> B. P(b))";
-by (Fast_tac 1);
-qed "subset_imp_Diff_eq";
-
-Goal "[| c \\<in> a-b; b<a |] ==> c=b | b<c & c<a";
-by (etac ltE 1);
-by (dtac Ord_linear 1);
-by (fast_tac (claset() addSIs [ltI] addIs [Ord_in_Ord]) 2);
-by (fast_tac (claset() addEs [Ord_in_Ord]) 1);
-qed "Ord_DiffE";
-
-val prems = goal thy "(!!y. y \\<in> A ==> P(y) = {x}) ==> x - (\\<Union>y \\<in> A. P(y)) = x";
-by (asm_full_simp_tac (simpset() addsimps prems) 1);
-by (fast_tac (claset() addSDs [prem] addSEs [mem_irrefl]) 1);
-qed "Diff_UN_eq_self";
-
-Goal "x - (\\<Union>b \\<in> a. HH(f,x,b)) = x - (\\<Union>b \\<in> a1. HH(f,x,b))  \
-\               ==> HH(f,x,a) = HH(f,x,a1)";
-by (resolve_tac [HH_def_satisfies_eq RS
-                (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1);
-by (etac subst_context 1);
-qed "HH_eq";
-
-Goal "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}";
-by (res_inst_tac [("P","b<a")] impE 1 THEN REPEAT (assume_tac 2));
-by (eresolve_tac [lt_Ord2 RS trans_induct] 1);
-by (rtac impI 1);
-by (resolve_tac [HH_eq RS trans] 1 THEN (assume_tac 2));
-by (resolve_tac [leI RS le_imp_subset RS subset_imp_Diff_eq RS ssubst] 1
-        THEN (assume_tac 1));
-by (res_inst_tac [("t","%z. z-?X")] subst_context 1);
-by (rtac Diff_UN_eq_self 1);
-by (dtac Ord_DiffE 1 THEN (assume_tac 1));
-by (fast_tac (claset() addEs [ltE]) 1);
-qed "HH_is_x_gt_too";
-
-Goal "[| HH(f,x,a) \\<in> Pow(x)-{0}; b<a |] ==> HH(f,x,b) \\<in> Pow(x)-{0}";
-by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
-by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1));
-by (dtac subst 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [mem_irrefl]) 1);
-qed "HH_subset_x_lt_too";
-
-Goal "HH(f,x,a) \\<in> Pow(x)-{0}   \
-\               ==> HH(f,x,a) \\<in> Pow(x - (\\<Union>b \\<in> a. HH(f,x,b)))-{0}";
-by (dresolve_tac [HH_def_satisfies_eq RS subst] 1);
-by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
-by (dresolve_tac [split_if RS iffD1] 1);
-by (Simp_tac 1);
-by (fast_tac (subset_cs addSEs [mem_irrefl]) 1);
-qed "HH_subset_x_imp_subset_Diff_UN";
-
-Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v \\<in> w |] ==> P";
-by (forw_inst_tac [("P","%y. y \\<in> Pow(x)-{0}")] subst 1 THEN (assume_tac 1));
-by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1);
-by (dtac subst_elem 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSIs [singleton_iff RS iffD2, equals0I]) 1);
-qed "HH_eq_arg_lt";
-
-Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0};  \
-\               Ord(v); Ord(w) |] ==> v=w";
-by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac);
-by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2
-        THEN REPEAT (assume_tac 2));
-by (dtac subst_elem 1 THEN (assume_tac 1));
-by (fast_tac (FOL_cs addDs [ltD] addSEs [HH_eq_arg_lt]) 1);
-qed "HH_eq_imp_arg_eq";
-
-Goalw [lepoll_def, inj_def]
-        "[| HH(f, x, i) \\<in> Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}";
-by (res_inst_tac [("x","\\<lambda>j \\<in> i. HH(f, x, j)")] exI 1);
-by (Asm_simp_tac 1);
-by (fast_tac (FOL_cs addSEs [HH_eq_imp_arg_eq, Ord_in_Ord, HH_subset_x_lt_too]
-                addSIs [lam_type, ballI, ltI] addIs [bexI]) 1);
-qed "HH_subset_x_imp_lepoll";
-
-Goal "HH(f, x, Hartog(Pow(x)-{0})) = {x}";
-by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 2));
-by (fast_tac (FOL_cs addSDs [HH_subset_x_imp_lepoll]
-                addSIs [Ord_Hartog] addSEs [HartogE]) 1);
-qed "HH_Hartog_is_x";
-
-Goal "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}";
-by (fast_tac (claset() addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1);
-qed "HH_Least_eq_x";
-
-Goal "a \\<in> (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \\<in> Pow(x)-{0}";
-by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
-by (rtac less_LeastE 1);
-by (eresolve_tac [Ord_Least RSN (2, ltI)] 2);
-by (assume_tac 1);
-qed "less_Least_subset_x";
-
-(* ********************************************************************** *)
-(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1              *)
-(* ********************************************************************** *)
-
-Goalw [inj_def]
-        "(\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
-\        \\<in> inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})";
-by (Asm_full_simp_tac 1);
-by (fast_tac (claset()  addSIs [lam_type] addDs [less_Least_subset_x]
-                addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1);
-qed "lam_Least_HH_inj_Pow";
-
-Goal "\\<forall>a \\<in> (LEAST i. HH(f,x,i)={x}). \\<exists>z \\<in> x. HH(f,x,a) = {z}  \
-\               ==> (\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
-\                       \\<in> inj(LEAST i. HH(f,x,i)={x}, {{y}. y \\<in> x})";
-by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1);
-by (Asm_full_simp_tac 1);
-qed "lam_Least_HH_inj";
-
-Goalw [surj_def]
-        "[| x - (\\<Union>a \\<in> A. F(a)) = 0;  \
-\               \\<forall>a \\<in> A. \\<exists>z \\<in> x. F(a) = {z} |]  \
-\               ==> (\\<lambda>a \\<in> A. F(a)) \\<in> surj(A, {{y}. y \\<in> x})";
-by (asm_full_simp_tac (simpset() addsimps [lam_type, Diff_eq_0_iff]) 1);
-by Safe_tac;
-by (set_mp_tac 1);
-by (deepen_tac (claset() addSIs [bexI] addSEs [equalityE]) 4 1);
-qed "lam_surj_sing";
-
-Goal "y \\<in> Pow(x)-{0} ==> x \\<noteq> 0";
-by Auto_tac;
-qed "not_emptyI2";
-
-Goal "f`(x - (\\<Union>j \\<in> i. HH(f,x,j))): Pow(x - (\\<Union>j \\<in> i. HH(f,x,j)))-{0}  \
-\       ==> HH(f, x, i) \\<in> Pow(x) - {0}";
-by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI,
-                not_emptyI2 RS if_P]) 1);
-by (Fast_tac 1);
-qed "f_subset_imp_HH_subset";
-
-val [prem] = goal thy "(!!z. z \\<in> Pow(x)-{0} ==> f`z \\<in> Pow(z)-{0}) ==>  \
-\       x - (\\<Union>j \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0";
-by (excluded_middle_tac "?P \\<in> {0}" 1);
-by (Fast_tac 2);
-by (dresolve_tac [Diff_subset RS PowI RS DiffI RS prem RS
-                f_subset_imp_HH_subset] 1);
-by (fast_tac (claset() addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)]
-                addSEs [mem_irrefl]) 1);
-qed "f_subsets_imp_UN_HH_eq_x";
-
-Goal "HH(f,x,i)=f`(x - (\\<Union>j \\<in> i. HH(f,x,j))) | HH(f,x,i)={x}";
-by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1);
-qed "HH_values2";
-
-Goal "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\\<Union>j \\<in> i. HH(f,x,j)))";
-by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSEs [equalityE, mem_irrefl]
-        addSDs [singleton_subsetD]) 1);
-qed "HH_subset_imp_eq";
-
-Goal "[| f \\<in> (Pow(x)-{0}) -> {{z}. z \\<in> x};  \
-\       a \\<in> (LEAST i. HH(f,x,i)={x}) |] ==> \\<exists>z \\<in> x. HH(f,x,a) = {z}";
-by (dtac less_Least_subset_x 1);
-by (ftac HH_subset_imp_eq 1);
-by (dtac apply_type 1);
-by (resolve_tac [Diff_subset RS PowI RS DiffI] 1);
-by (fast_tac 
-    (claset() addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
-by (fast_tac (claset() addss (simpset())) 1);
-qed "f_sing_imp_HH_sing";
-
-Goalw [bij_def] 
-        "[| x - (\\<Union>j \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0;  \
-\       f \\<in> (Pow(x)-{0}) -> {{z}. z \\<in> x} |]  \
-\       ==> (\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))  \
-\                       \\<in> bij(LEAST i. HH(f,x,i)={x}, {{y}. y \\<in> x})";
-by (fast_tac (claset() addSIs [lam_Least_HH_inj, lam_surj_sing,
-                              f_sing_imp_HH_sing]) 1);
-qed "f_sing_lam_bij";
-
-Goal "f \\<in> (\\<Pi>X \\<in> Pow(x)-{0}. F(X))  \
-\       ==> (\\<lambda>X \\<in> Pow(x)-{0}. {f`X}) \\<in> (\\<Pi>X \\<in> Pow(x)-{0}. {{z}. z \\<in> F(X)})";
-by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2]
-                     addDs [apply_type]) 1);
-qed "lam_singI";
-
-val bij_Least_HH_x = 
-    (lam_singI RSN (2, [f_sing_lam_bij, lam_sing_bij RS bij_converse_bij]
-                    MRS comp_bij)) |> standard;
--- a/src/ZF/AC/HH.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/HH.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -2,22 +2,255 @@
     ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-The theory file for the proofs of
+Some properties of the recursive definition of HH used in the proofs of
   AC17 ==> AC1
   AC1 ==> WO2
   AC15 ==> WO6
 *)
 
-HH = AC_Equiv + Hartog + WO_AC + Let +
+theory HH = AC_Equiv + Hartog:
 
-consts
+constdefs
  
-  HH                      :: [i, i, i] => i
+  HH :: "[i, i, i] => i"
+    "HH(f,x,a) == transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
+                                    in  if f`z \<in> Pow(z)-{0} then f`z else {x})"
+
+
+(* ********************************************************************** *)
+(* Lemmas useful in each of the three proofs                              *)
+(* ********************************************************************** *)
+
+lemma HH_def_satisfies_eq:
+     "HH(f,x,a) = (let z = x - (\<Union>b \<in> a. HH(f,x,b))   
+                   in  if f`z \<in> Pow(z)-{0} then f`z else {x})"
+by (rule HH_def [THEN def_transrec, THEN trans], simp)
+
+lemma HH_values: "HH(f,x,a) \<in> Pow(x)-{0} | HH(f,x,a)={x}"
+apply (rule HH_def_satisfies_eq [THEN ssubst])
+apply (simp add: Let_def Diff_subset [THEN PowI], fast)
+done
+
+lemma subset_imp_Diff_eq:
+     "B \<subseteq> A ==> X-(\<Union>a \<in> A. P(a)) = X-(\<Union>a \<in> A-B. P(a))-(\<Union>b \<in> B. P(b))"
+by fast
+
+lemma Ord_DiffE: "[| c \<in> a-b; b<a |] ==> c=b | b<c & c<a"
+apply (erule ltE)
+apply (drule Ord_linear [of _ c])
+apply (fast elim: Ord_in_Ord)
+apply (fast intro!: ltI intro: Ord_in_Ord)
+done
+
+lemma Diff_UN_eq_self:
+     "(!!y. y \<in> A ==> P(y) = {x}) ==> x - (\<Union>y \<in> A. P(y)) = x" 
+apply (simp, fast elim!: mem_irrefl)
+done
+
+lemma HH_eq: "x - (\<Union>b \<in> a. HH(f,x,b)) = x - (\<Union>b \<in> a1. HH(f,x,b))   
+              ==> HH(f,x,a) = HH(f,x,a1)"
+apply (subst HH_def_satisfies_eq) 
+apply (rule HH_def_satisfies_eq [THEN trans], simp) 
+done
+
+lemma HH_is_x_gt_too: "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}"
+apply (rule_tac P = "b<a" in impE)
+prefer 2 apply assumption+
+apply (erule lt_Ord2 [THEN trans_induct])
+apply (rule impI)
+apply (rule HH_eq [THEN trans])
+prefer 2 apply assumption+
+apply (rule leI [THEN le_imp_subset, THEN subset_imp_Diff_eq, THEN ssubst], 
+       assumption)
+apply (rule_tac t = "%z. z-?X" in subst_context)
+apply (rule Diff_UN_eq_self)
+apply (drule Ord_DiffE, assumption) 
+apply (fast elim: ltE, auto) 
+done
+
+lemma HH_subset_x_lt_too:
+     "[| HH(f,x,a) \<in> Pow(x)-{0}; b<a |] ==> HH(f,x,b) \<in> Pow(x)-{0}"
+apply (rule HH_values [THEN disjE], assumption)
+apply (drule HH_is_x_gt_too, assumption)
+apply (drule subst, assumption)
+apply (fast elim!: mem_irrefl)
+done
+
+lemma HH_subset_x_imp_subset_Diff_UN:
+    "HH(f,x,a) \<in> Pow(x)-{0} ==> HH(f,x,a) \<in> Pow(x - (\<Union>b \<in> a. HH(f,x,b)))-{0}"
+apply (drule HH_def_satisfies_eq [THEN subst])
+apply (rule HH_def_satisfies_eq [THEN ssubst])
+apply (simp add: Let_def Diff_subset [THEN PowI])
+apply (drule split_if [THEN iffD1])
+apply (fast elim!: mem_irrefl)
+done
+
+lemma HH_eq_arg_lt:
+     "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v) \<in> Pow(x)-{0}; v \<in> w |] ==> P"
+apply (frule_tac P = "%y. y \<in> Pow (x) -{0}" in subst, assumption)
+apply (drule_tac a = "w" in HH_subset_x_imp_subset_Diff_UN)
+apply (drule subst_elem, assumption)
+apply (fast intro!: singleton_iff [THEN iffD2] equals0I)
+done
+
+lemma HH_eq_imp_arg_eq:
+  "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w) \<in> Pow(x)-{0}; Ord(v); Ord(w) |] ==> v=w"
+apply (rule_tac j = "w" in Ord_linear_lt)
+apply (simp_all (no_asm_simp))
+ apply (drule subst_elem, assumption) 
+ apply (blast dest: ltD HH_eq_arg_lt)
+apply (blast dest: HH_eq_arg_lt [OF sym] ltD) 
+done
+
+lemma HH_subset_x_imp_lepoll: 
+     "[| HH(f, x, i) \<in> Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"
+apply (unfold lepoll_def inj_def)
+apply (rule_tac x = "\<lambda>j \<in> i. HH (f, x, j) " in exI)
+apply (simp (no_asm_simp))
+apply (fast del: DiffE
+	    elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too 
+            intro!: lam_type ballI ltI intro: bexI)
+done
+
+lemma HH_Hartog_is_x: "HH(f, x, Hartog(Pow(x)-{0})) = {x}"
+apply (rule HH_values [THEN disjE])
+prefer 2 apply assumption 
+apply (fast del: DiffE
+            intro!: Ord_Hartog 
+	    dest!: HH_subset_x_imp_lepoll 
+            elim!: Hartog_lepoll_selfE)
+done
+
+lemma HH_Least_eq_x: "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}"
+by (fast intro!: Ord_Hartog HH_Hartog_is_x LeastI)
+
+lemma less_Least_subset_x:
+     "a \<in> (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \<in> Pow(x)-{0}"
+apply (rule HH_values [THEN disjE], assumption)
+apply (rule less_LeastE)
+apply (erule_tac [2] ltI [OF _ Ord_Least], assumption)
+done
 
-defs
+(* ********************************************************************** *)
+(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1              *)
+(* ********************************************************************** *)
+
+lemma lam_Least_HH_inj_Pow: 
+        "(\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))   
+         \<in> inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"
+apply (unfold inj_def, simp)
+apply (fast intro!: lam_type dest: less_Least_subset_x 
+            elim!: HH_eq_imp_arg_eq Ord_Least [THEN Ord_in_Ord])
+done
+
+lemma lam_Least_HH_inj:
+     "\<forall>a \<in> (LEAST i. HH(f,x,i)={x}). \<exists>z \<in> x. HH(f,x,a) = {z}   
+      ==> (\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))   
+          \<in> inj(LEAST i. HH(f,x,i)={x}, {{y}. y \<in> x})"
+by (rule lam_Least_HH_inj_Pow [THEN inj_strengthen_type], simp)
+
+lemma lam_surj_sing: 
+        "[| x - (\<Union>a \<in> A. F(a)) = 0;  \<forall>a \<in> A. \<exists>z \<in> x. F(a) = {z} |]   
+         ==> (\<lambda>a \<in> A. F(a)) \<in> surj(A, {{y}. y \<in> x})"
+apply (simp add: surj_def lam_type Diff_eq_0_iff)
+apply (blast elim: equalityE) 
+done
+
+lemma not_emptyI2: "y \<in> Pow(x)-{0} ==> x \<noteq> 0"
+by auto
+
+lemma f_subset_imp_HH_subset:
+     "f`(x - (\<Union>j \<in> i. HH(f,x,j))) \<in> Pow(x - (\<Union>j \<in> i. HH(f,x,j)))-{0}   
+      ==> HH(f, x, i) \<in> Pow(x) - {0}"
+apply (rule HH_def_satisfies_eq [THEN ssubst])
+apply (simp add: Let_def Diff_subset [THEN PowI] not_emptyI2 [THEN if_P], fast)
+done
+
+lemma f_subsets_imp_UN_HH_eq_x:
+     "\<forall>z \<in> Pow(x)-{0}. f`z \<in> Pow(z)-{0}
+      ==> x - (\<Union>j \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"
+apply (case_tac "?P \<in> {0}", fast)
+apply (drule Diff_subset [THEN PowI, THEN DiffI])
+apply (drule bspec, assumption) 
+apply (drule f_subset_imp_HH_subset) 
+apply (blast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] 
+             elim!: mem_irrefl)
+done
+
+lemma HH_values2: "HH(f,x,i) = f`(x - (\<Union>j \<in> i. HH(f,x,j))) | HH(f,x,i)={x}"
+apply (rule HH_def_satisfies_eq [THEN ssubst])
+apply (simp add: Let_def Diff_subset [THEN PowI])
+done
+
+lemma HH_subset_imp_eq:
+     "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\<Union>j \<in> i. HH(f,x,j)))"
+apply (rule HH_values2 [THEN disjE], assumption)
+apply (fast elim!: equalityE mem_irrefl dest!: singleton_subsetD)
+done
 
-  HH_def  "HH(f,x,a) == transrec(a, %b r. let z = x - (\\<Union>c \\<in> b. r`c)
-                        in  if(f`z \\<in> Pow(z)-{0}, f`z, {x}))"
+lemma f_sing_imp_HH_sing:
+     "[| f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x};   
+         a \<in> (LEAST i. HH(f,x,i)={x}) |] ==> \<exists>z \<in> x. HH(f,x,a) = {z}"
+apply (drule less_Least_subset_x)
+apply (frule HH_subset_imp_eq)
+apply (drule apply_type)
+apply (rule Diff_subset [THEN PowI, THEN DiffI])
+apply (fast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2], force) 
+done
+
+lemma f_sing_lam_bij: 
+     "[| x - (\<Union>j \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0;   
+         f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x} |]   
+      ==> (\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))   
+          \<in> bij(LEAST i. HH(f,x,i)={x}, {{y}. y \<in> x})"
+apply (unfold bij_def)
+apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing)
+done
+
+lemma lam_singI:
+     "f \<in> (\<Pi>X \<in> Pow(x)-{0}. F(X))   
+      ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X}) \<in> (\<Pi>X \<in> Pow(x)-{0}. {{z}. z \<in> F(X)})"
+by (fast del: DiffI DiffE
+	    intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type)
+
+(*FIXME: both uses have the form ...[THEN bij_converse_bij], so 
+  simplification is needed!*)
+lemmas bij_Least_HH_x =  
+    comp_bij [OF f_sing_lam_bij [OF _ lam_singI] 
+              lam_sing_bij [THEN bij_converse_bij], standard]
+
+
+(* ********************************************************************** *)
+(*                     The proof of AC1 ==> WO2                           *)
+(* ********************************************************************** *)
+
+(*Establishing the existence of a bijection, namely
+converse
+ (converse(\<lambda>x\<in>x. {x}) O
+  Lambda
+   (LEAST i. HH(\<lambda>X\<in>Pow(x) - {0}. {f ` X}, x, i) = {x},
+    HH(\<lambda>X\<in>Pow(x) - {0}. {f ` X}, x)))
+Perhaps it could be simplified. *)
+
+lemma bijection:
+     "f \<in> (\<Pi>X \<in> Pow(x) - {0}. X) 
+      ==> \<exists>g. g \<in> bij(x, LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x})"
+apply (rule exI) 
+apply (rule bij_Least_HH_x [THEN bij_converse_bij])
+apply (rule f_subsets_imp_UN_HH_eq_x)
+apply (intro ballI apply_type) 
+apply (fast intro: lam_type apply_type del: DiffE)
+apply assumption; 
+apply (fast intro: Pi_weaken_type)
+done
+
+lemma AC1_WO2: "AC1 ==> WO2"
+apply (unfold AC1_def WO2_def eqpoll_def)
+apply (intro allI) 
+apply (drule_tac x = "Pow(A) - {0}" in spec) 
+apply (blast dest: bijection)
+done
 
 end
 
+
--- a/src/ZF/AC/Hartog.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,82 +0,0 @@
-(*  Title:      ZF/AC/Hartog.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-  Some proofs on the Hartogs function.
-*)
-
-Goal "\\<forall>a. Ord(a) --> a \\<in> X ==> P";
-by (res_inst_tac [("X1","{y \\<in> X. Ord(y)}")] (ON_class RS revcut_rl) 1);
-by (Fast_tac 1);
-qed "Ords_in_set";
-
-Goalw [lepoll_def] "[| Ord(a); a lepoll X |] ==>  \
-\               \\<exists>Y. Y \\<subseteq> X & (\\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)";
-by (etac exE 1);
-by (REPEAT (resolve_tac [exI, conjI] 1));
-by (eresolve_tac [inj_is_fun RS fun_is_rel RS image_subset] 1);
-by (rtac exI 1);
-by (rtac conjI 1);
-by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, 
-        restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1
-        THEN (assume_tac 1));
-by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS
-        (well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS 
-        (ordertype_Memrel RSN (2, trans))] 1
-        THEN (REPEAT (assume_tac 1)));
-qed "Ord_lepoll_imp_ex_well_ord";
-
-Goal "[| Ord(a); a lepoll X |] ==>  \
-\               \\<exists>Y. Y \\<subseteq> X & (\\<exists>R. R \\<subseteq> X*X & ordertype(Y,R)=a)";
-by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1));
-by Safe_tac;
-by (REPEAT (ares_tac [exI, conjI] 1));
-by (etac ordertype_Int 2);
-by (Fast_tac 1);
-qed "Ord_lepoll_imp_eq_ordertype";
-
-Goal "\\<forall>a. Ord(a) --> a lepoll X ==>  \
-\       \\<forall>a. Ord(a) -->  \
-\       a:{a. Z \\<in> Pow(X)*Pow(X*X), \\<exists>Y R. Z=<Y,R> & ordertype(Y,R)=a}";
-by (REPEAT (resolve_tac [allI,impI] 1));
-by (REPEAT (eresolve_tac [allE, impE] 1));
-by (assume_tac 1);
-by (dtac Ord_lepoll_imp_eq_ordertype 1 THEN (assume_tac 1));
-by (fast_tac (claset() addSIs [ReplaceI] addEs [sym]) 1);
-qed "Ords_lepoll_set_lemma";
-
-Goal "\\<forall>a. Ord(a) --> a lepoll X ==> P";
-by (eresolve_tac [Ords_lepoll_set_lemma RS Ords_in_set] 1);
-qed "Ords_lepoll_set";
-
-Goal "\\<exists>a. Ord(a) & ~a lepoll X";
-by (rtac swap 1);
-by (Fast_tac 1);
-by (rtac Ords_lepoll_set 1);
-by (Fast_tac 1);
-qed "ex_Ord_not_lepoll";
-
-Goalw [Hartog_def] "~ Hartog(A) lepoll A";
-by (resolve_tac [ex_Ord_not_lepoll RS exE] 1);
-by (rtac LeastI 1);
-by (REPEAT (Fast_tac 1));
-qed "HartogI";
-
-val HartogE = HartogI RS notE;
-
-Goalw [Hartog_def] "Ord(Hartog(A))";
-by (rtac Ord_Least 1);
-qed "Ord_Hartog";
-
-Goalw [Hartog_def] "[| i < Hartog(A); ~ i lepoll A |] ==> P";
-by (fast_tac (claset() addEs [less_LeastE]) 1);
-qed "less_HartogE1";
-
-Goal "[| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
-by (fast_tac (claset() addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
-                RS lepoll_trans RS HartogE]) 1);
-qed "less_HartogE";
-
-Goal "Card(Hartog(A))";
-by (fast_tac (claset() addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1);
-qed "Card_Hartog";
--- a/src/ZF/AC/Hartog.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/Hartog.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -5,14 +5,79 @@
 Hartog's function.
 *)
 
-Hartog = Cardinal +
+theory Hartog = AC_Equiv:
+
+constdefs
+  Hartog :: "i => i"
+   "Hartog(X) == LEAST i. ~ i \<lesssim> X"
+
+lemma Ords_in_set: "\<forall>a. Ord(a) --> a \<in> X ==> P"
+apply (rule_tac X1 = "{y \<in> X. Ord (y) }" in ON_class [THEN revcut_rl])
+apply fast
+done
 
-consts
+lemma Ord_lepoll_imp_ex_well_ord:
+     "[| Ord(a); a \<lesssim> X |] 
+      ==> \<exists>Y. Y \<subseteq> X & (\<exists>R. well_ord(Y,R) & ordertype(Y,R)=a)"
+apply (unfold lepoll_def)
+apply (erule exE)
+apply (intro exI conjI)
+  apply (erule inj_is_fun [THEN fun_is_rel, THEN image_subset])
+ apply (rule well_ord_rvimage [OF bij_is_inj well_ord_Memrel]) 
+  apply (erule restrict_bij [THEN bij_converse_bij]) 
+apply (rule subset_refl, assumption); 
+apply (rule trans) 
+apply (rule bij_ordertype_vimage) 
+apply (erule restrict_bij [THEN bij_converse_bij]) 
+apply (rule subset_refl) 
+apply (erule well_ord_Memrel) 
+apply (erule ordertype_Memrel) 
+done
+
+lemma Ord_lepoll_imp_eq_ordertype:
+     "[| Ord(a); a \<lesssim> X |] ==> \<exists>Y. Y \<subseteq> X & (\<exists>R. R \<subseteq> X*X & ordertype(Y,R)=a)"
+apply (drule Ord_lepoll_imp_ex_well_ord, (assumption))
+apply clarify
+apply (intro exI conjI)
+apply (erule_tac [3] ordertype_Int, auto) 
+done
 
-  Hartog :: i => i
+lemma Ords_lepoll_set_lemma:
+     "(\<forall>a. Ord(a) --> a \<lesssim> X) ==>   
+       \<forall>a. Ord(a) -->   
+        a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=<Y,R> & ordertype(Y,R)=b}"
+apply (intro allI impI)
+apply (elim allE impE, assumption)
+apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) 
+done
+
+lemma Ords_lepoll_set: "\<forall>a. Ord(a) --> a \<lesssim> X ==> P"
+by (erule Ords_lepoll_set_lemma [THEN Ords_in_set])
+
+lemma ex_Ord_not_lepoll: "\<exists>a. Ord(a) & ~a \<lesssim> X"
+apply (rule ccontr)
+apply (best intro: Ords_lepoll_set) 
+done
 
-defs
+lemma not_Hartog_lepoll_self: "~ Hartog(A) \<lesssim> A"
+apply (unfold Hartog_def)
+apply (rule ex_Ord_not_lepoll [THEN exE])
+apply (rule LeastI, auto) 
+done
+
+lemmas Hartog_lepoll_selfE = not_Hartog_lepoll_self [THEN notE, standard]
 
-  Hartog_def "Hartog(X) == LEAST i. ~i lepoll X"
+lemma Ord_Hartog: "Ord(Hartog(A))"
+by (unfold Hartog_def, rule Ord_Least)
+
+lemma less_HartogE1: "[| i < Hartog(A); ~ i \<lesssim> A |] ==> P"
+by (unfold Hartog_def, fast elim: less_LeastE)
+
+lemma less_HartogE: "[| i < Hartog(A); i \<approx> Hartog(A) |] ==> P"
+by (blast intro: less_HartogE1 eqpoll_sym eqpoll_imp_lepoll 
+                 lepoll_trans [THEN Hartog_lepoll_selfE]);
+
+lemma Card_Hartog: "Card(Hartog(A))"
+by (fast intro!: CardI Ord_Hartog elim: less_HartogE)
 
 end
--- a/src/ZF/AC/ROOT.ML	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/ROOT.ML	Wed Jan 16 17:52:06 2002 +0100
@@ -6,15 +6,12 @@
 Executes the proofs of the AC-equivalences, due to Krzysztof Grabczewski
 *)
 
-time_use_thy "AC_Equiv";
-
 time_use_thy "WO6_WO1";
 time_use_thy "WO1_WO7";
 
-time_use     "AC7_AC9.ML";
+time_use_thy "AC7_AC9";
 
 time_use_thy "WO1_AC";
-time_use_thy "AC1_WO2";
 
 time_use_thy "AC15_WO6";
 
--- a/src/ZF/AC/WO1_AC.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-(*  Title:      ZF/AC/WO1_AC.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
-
-The latter proof is referred to as clear by the Rubins.
-However it seems to be quite complicated.
-The formal proof presented below is a mechanisation of the proof 
-by Lawrence C. Paulson which is the following:
-
-Assume WO1.  Let s be a set of infinite sets.
- 
-Suppose x \\<in> s.  Then x is equipollent to |x| (by WO1), an infinite cardinal; 
-call it K.  Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an 
-isomorphism h \\<in> bij(K+K, x).  (Here + means disjoint sum.)
- 
-So there is a partition of x into 2-element sets, namely
- 
-        {{h(Inl(i)), h(Inr(i))} . i \\<in> K}
- 
-So for all x \\<in> s the desired partition exists.  By AC1 (which follows from WO1) 
-there exists a function f that chooses a partition for each x \\<in> s.  Therefore we 
-have AC10(2).
-
-*)
-
-(* ********************************************************************** *)
-(* WO1 ==> AC1                                                            *)
-(* ********************************************************************** *)
-
-Goalw [AC1_def, WO1_def] "WO1 ==> AC1";
-by (fast_tac (claset() addSEs [ex_choice_fun]) 1);
-qed "WO1_AC1";
-
-(* ********************************************************************** *)
-(* WO1 ==> AC10(n) (n >= 1)                                               *)
-(* ********************************************************************** *)
-
-Goalw [WO1_def] "[| WO1; \\<forall>B \\<in> A. \\<exists>C \\<in> D(B). P(C,B) |]  \
-\               ==> \\<exists>f. \\<forall>B \\<in> A. P(f`B,B)";
-by (eres_inst_tac [("x","Union({{C \\<in> D(B). P(C,B)}. B \\<in> A})")] allE 1);
-by (etac exE 1);
-by (dtac ex_choice_fun 1);
-by (Fast_tac 1);
-by (etac exE 1);
-by (res_inst_tac [("x","\\<lambda>x \\<in> A. f`{C \\<in> D(x). P(C,x)}")] exI 1);
-by (Asm_full_simp_tac 1);
-by (blast_tac (claset() addSDs [RepFunI RSN (2, apply_type)]) 1);
-val lemma1 = result();
-
-Goalw [WO1_def] "[| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
-by (rtac eqpoll_trans 1);
-by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 2);
-by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
-by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1);
-by (fold_tac [cadd_def]);
-by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS 
-		 InfCard_cdouble_eq RS ssubst] 1);
-by (rtac eqpoll_refl 2);
-by (rtac notI 1);
-by (etac notE 1);
-by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1
-        THEN (assume_tac 2));
-by (fast_tac (claset() addSEs [well_ord_cardinal_eqpoll]) 1);
-val lemma2_1 = result();
-
-Goal "f \\<in> bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \\<in> D} \\<in> Pow(Pow(B))";
-by (fast_tac (claset() addSEs [bij_is_fun RS apply_type]) 1);
-val lemma2_2 = result();
-
-Goal "[| f \\<in> inj(A,B); f`a = f`b; a \\<in> A; b \\<in> A |] ==> a=b";
-by (rtac inj_equality 1);
-by (TRYALL (fast_tac (claset() addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
-val lemma = result();
-
-Goalw AC_aux_defs
-        "f \\<in> bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \\<in> D})";
-by (blast_tac (claset() addDs [bij_is_inj RS lemma]) 1);
-val lemma2_3 = result();
-
-val [major, minor] = goalw thy AC_aux_defs 
-        "[| f \\<in> bij(D+D, B); 1 le n |] ==>  \
-\       sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \\<in> D}, 2, succ(n))";
-by (rewtac succ_def);
-by (fast_tac (claset() 
-        addSIs [cons_lepoll_cong, minor, lepoll_refl] 
-        addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
-                le_imp_subset RS subset_imp_lepoll]
-        addDs [major RS bij_is_inj RS lemma]
-        addSEs [mem_irrefl]) 1);
-val lemma2_4 = result();
-
-Goalw [bij_def, surj_def]
-        "f \\<in> bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i \\<in> D})=B";
-by (fast_tac (claset() addSEs [inj_is_fun RS apply_type]) 1);
-val lemma2_5 = result();
-
-Goal "[| WO1; ~Finite(B); 1 le n  |]  \
-\       ==> \\<exists>C \\<in> Pow(Pow(B)). pairwise_disjoint(C) &  \
-\               sets_of_size_between(C, 2, succ(n)) &  \
-\               Union(C)=B";
-by (eresolve_tac [lemma2_1 RS (eqpoll_def RS def_imp_iff RS iffD1 RS exE)] 1
-        THEN (assume_tac 1));
-by (fast_tac (FOL_cs addSIs [bexI]
-                addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1);
-val lemma2 = result();
-
-Goalw AC_defs "[| WO1; 1 le n |] ==> AC10(n)";
-by (fast_tac (claset() addSIs [lemma1] addSEs [lemma2]) 1);
-qed "WO1_AC10";
--- a/src/ZF/AC/WO1_AC.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/WO1_AC.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,105 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/WO1_AC.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
+
+The latter proof is referred to as clear by the Rubins.
+However it seems to be quite complicated.
+The formal proof presented below is a mechanisation of the proof 
+by Lawrence C. Paulson which is the following:
+
+Assume WO1.  Let s be a set of infinite sets.
+ 
+Suppose x \<in> s.  Then x is equipollent to |x| (by WO1), an infinite cardinal
+call it K.  Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an 
+isomorphism h \<in> bij(K+K, x).  (Here + means disjoint sum.)
+ 
+So there is a partition of x into 2-element sets, namely
+ 
+        {{h(Inl(i)), h(Inr(i))} . i \<in> K}
+ 
+So for all x \<in> s the desired partition exists.  By AC1 (which follows from WO1) 
+there exists a function f that chooses a partition for each x \<in> s.  Therefore we 
+have AC10(2).
+
+*)
+
+theory WO1_AC = AC_Equiv:
+
+(* ********************************************************************** *)
+(* WO1 ==> AC1                                                            *)
+(* ********************************************************************** *)
+
+theorem WO1_AC1: "WO1 ==> AC1"
+by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun)
+
+(* ********************************************************************** *)
+(* WO1 ==> AC10(n) (n >= 1)                                               *)
+(* ********************************************************************** *)
+
+lemma lemma1: "[| WO1; \<forall>B \<in> A. \<exists>C \<in> D(B). P(C,B) |] ==> \<exists>f. \<forall>B \<in> A. P(f`B,B)"
+apply (unfold WO1_def)
+apply (erule_tac x = "Union ({{C \<in> D (B) . P (C,B) }. B \<in> A}) " in allE)
+apply (erule exE, drule ex_choice_fun, fast)
+apply (erule exE)
+apply (rule_tac x = "\<lambda>x \<in> A. f`{C \<in> D (x) . P (C,x) }" in exI)
+apply (simp, blast dest!: apply_type [OF _ RepFunI])
+done
 
-WO1_AC = AC_Equiv + WO_AC
+lemma lemma2_1: "[| ~Finite(B); WO1 |] ==> |B| + |B| \<approx>  B"
+apply (unfold WO1_def)
+apply (rule eqpoll_trans)
+prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll)
+apply (rule eqpoll_sym [THEN eqpoll_trans])
+apply (fast elim!: well_ord_cardinal_eqpoll)
+apply (drule spec [of _ B]) 
+apply (clarify dest!: eqpoll_imp_Finite_iff [OF well_ord_cardinal_eqpoll]) 
+apply (simp add: cadd_def [symmetric] 
+            eqpoll_refl InfCard_cdouble_eq Card_cardinal Inf_Card_is_InfCard) 
+done
+
+lemma lemma2_2:
+     "f \<in> bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \<in> D} \<in> Pow(Pow(B))"
+by (fast elim!: bij_is_fun [THEN apply_type])
+
+
+lemma lemma2_3: 
+        "f \<in> bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \<in> D})"
+apply (unfold pairwise_disjoint_def)
+apply (blast dest: bij_is_inj [THEN inj_apply_equality])
+done
+
+lemma lemma2_4:
+     "[| f \<in> bij(D+D, B); 1\<le>n |] 
+      ==> sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \<in> D}, 2, succ(n))"
+apply (simp (no_asm_simp) add: sets_of_size_between_def succ_def)
+apply (blast intro!: cons_lepoll_cong 
+            intro: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll]  
+                   le_imp_subset [THEN subset_imp_lepoll]  lepoll_trans 
+            dest: bij_is_inj [THEN inj_apply_equality] elim!: mem_irrefl)
+done
+
+lemma lemma2_5: 
+     "f \<in> bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i \<in> D})=B"
+apply (unfold bij_def surj_def)
+apply (fast elim!: inj_is_fun [THEN apply_type])
+done
+
+lemma lemma2:
+     "[| WO1; ~Finite(B); 1\<le>n  |]   
+      ==> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) &   
+                sets_of_size_between(C, 2, succ(n)) &   
+                Union(C)=B"
+apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], 
+       assumption)
+apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5)
+done
+
+theorem WO1_AC10: "[| WO1; 1\<le>n |] ==> AC10(n)"
+apply (unfold AC10_def)
+apply (fast intro!: lemma1 elim!: lemma2)
+done
+
+end
+
--- a/src/ZF/AC/WO1_WO7.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-(*  Title:      ZF/AC/WO1_WO7.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5)
-LEMMA is the sentence denoted by (**)
-
-Also, WO1 <-> WO8
-*)
-
-(* ********************************************************************** *)
-(* It is easy to see, that WO7 is equivallent to (**)                     *)
-(* ********************************************************************** *)
-
-Goalw [WO7_def, LEMMA_def] 
-  "WO7 <-> LEMMA";
-by (fast_tac (claset() addSEs [Finite_well_ord_converse]) 1);
-qed "WO7_iff_LEMMA";
-
-(* ********************************************************************** *)
-(* It is also easy to show that LEMMA implies WO1.                        *)
-(* ********************************************************************** *)
-
-Goalw [WO1_def, LEMMA_def] "LEMMA ==> WO1";
-by (rtac allI 1);
-by (etac allE 1);
-by (excluded_middle_tac "Finite(A)" 1);
-by (Fast_tac 1);
-by (rewrite_goals_tac [Finite_def, eqpoll_def]);
-by (fast_tac (claset() addSIs [[bij_is_inj, nat_implies_well_ord] MRS
-                                 well_ord_rvimage]) 1);
-qed "LEMMA_imp_WO1";
-
-(* ********************************************************************** *)
-(* The Rubins' proof of the other implication is contained within the     *)
-(* following sentence \\<in>                                                   *)
-(* "... each infinite ordinal is well ordered by < but not by >."         *)
-(* This statement can be proved by the following two theorems.            *)
-(* But moreover we need to show similar property for any well ordered     *)
-(* infinite set. It is not very difficult thanks to Isabelle order types  *)
-(* We show that if a set is well ordered by some relation and by its     *)
-(* converse, then apropriate order type is well ordered by the converse   *)
-(* of it's membership relation, which in connection with the previous     *)
-(* gives the conclusion.                                                  *)
-(* ********************************************************************** *)
-
-Goalw [wf_on_def, wf_def] 
-    "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))";
-by (dresolve_tac [nat_le_infinite_Ord RS le_imp_subset] 1 
-    THEN (assume_tac 1));
-by (rtac notI 1);
-by (eres_inst_tac [("x","nat")] allE 1);
-by (Blast_tac 1);
-qed "converse_Memrel_not_wf_on";
-
-Goalw [well_ord_def] 
-    "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))";
-by (fast_tac (claset() addSDs [converse_Memrel_not_wf_on]) 1);
-qed "converse_Memrel_not_well_ord";
-
-Goal "[| well_ord(A,r); well_ord(A,converse(r)) |]  \
-\       ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))";
-by (rtac ([ordertype_ord_iso RS ord_iso_sym RS ord_iso_rvimage_eq, 
-                Memrel_type RS (subset_Int_iff RS iffD1)] 
-                MRS trans RS subst) 1
-        THEN (assume_tac 1));
-by (rtac (rvimage_converse RS subst) 1);
-by (etac (ordertype_ord_iso RS ord_iso_sym RS ord_iso_is_bij RS
-                bij_is_inj RS well_ord_rvimage) 1
-        THEN (assume_tac 1));
-qed "well_ord_converse_Memrel";
-
-Goalw [WO1_def, LEMMA_def] "WO1 ==> LEMMA";
-by (REPEAT (resolve_tac [allI,impI] 1));
-by (REPEAT (eresolve_tac [allE,exE] 1));
-by (REPEAT (ares_tac [exI,conjI,notI] 1));
-by (ftac well_ord_converse_Memrel 1 THEN (assume_tac 1));
-by (forward_tac [Ord_ordertype RS converse_Memrel_not_well_ord] 1);
-by (contr_tac 2);
-by (fast_tac (empty_cs addSEs [ordertype_ord_iso RS ord_iso_is_bij RS 
-                bij_is_inj RS (exI RS (lepoll_def RS def_imp_iff RS iffD2))
-                RS lepoll_Finite]
-                addSIs [notI] addEs [notE]) 1);
-qed "WO1_imp_LEMMA";
-
-
-Goal "WO1 <-> WO7";
-by (simp_tac (simpset() addsimps [WO7_iff_LEMMA]) 1);
-by (blast_tac (claset() addIs [LEMMA_imp_WO1, WO1_imp_LEMMA]) 1);
-qed "WO1_iff_WO7";
-
-
-
-(* ********************************************************************** *)
-
-(*            The proof of WO8 <-> WO1 (Rubin & Rubin p. 6)               *)
-
-(* ********************************************************************** *)
-
-Goalw WO_defs "WO1 ==> WO8";
-by (Fast_tac 1);
-qed "WO1_WO8";
-
-
-(* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof   *)
-Goalw WO_defs "WO8 ==> WO1";
-by (rtac allI 1);
-by (eres_inst_tac [("x","{{x}. x \\<in> A}")] allE 1);
-by (etac impE 1);
-by (fast_tac (claset() addSEs [lam_sing_bij RS bij_is_inj RS
-                        well_ord_rvimage]) 2);
-by (res_inst_tac [("x","\\<lambda>a \\<in> {{x}. x \\<in> A}. THE x. a={x}")] exI 1);
-by (force_tac (claset() addSIs [lam_type],
-               simpset() addsimps [singleton_eq_iff, the_equality]) 1);
-qed "WO8_WO1";
--- a/src/ZF/AC/WO1_WO7.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/WO1_WO7.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -5,13 +5,110 @@
 
 WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5)
 LEMMA is the sentence denoted by (**)
+
+Also, WO1 <-> WO8
 *)
 
-WO1_WO7 = AC_Equiv +
+theory WO1_WO7 = AC_Equiv:
 
 constdefs
   LEMMA :: o
     "LEMMA ==
-     \\<forall>X. ~Finite(X) --> (\\<exists>R. well_ord(X,R) & ~well_ord(X,converse(R)))"
+     \<forall>X. ~Finite(X) --> (\<exists>R. well_ord(X,R) & ~well_ord(X,converse(R)))"
+
+(* ********************************************************************** *)
+(* It is easy to see that WO7 is equivalent to (**)                       *)
+(* ********************************************************************** *)
+
+lemma WO7_iff_LEMMA: "WO7 <-> LEMMA"
+apply (unfold WO7_def LEMMA_def)
+apply (blast intro: Finite_well_ord_converse)
+done
+
+(* ********************************************************************** *)
+(* It is also easy to show that LEMMA implies WO1.                        *)
+(* ********************************************************************** *)
+
+lemma LEMMA_imp_WO1: "LEMMA ==> WO1"
+apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def)
+apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord])
+done
+
+(* ********************************************************************** *)
+(* The Rubins' proof of the other implication is contained within the     *)
+(* following sentence \<in>                                                   *)
+(* "... each infinite ordinal is well ordered by < but not by >."         *)
+(* This statement can be proved by the following two theorems.            *)
+(* But moreover we need to show similar property for any well ordered     *)
+(* infinite set. It is not very difficult thanks to Isabelle order types  *)
+(* We show that if a set is well ordered by some relation and by its     *)
+(* converse, then apropriate order type is well ordered by the converse   *)
+(* of it's membership relation, which in connection with the previous     *)
+(* gives the conclusion.                                                  *)
+(* ********************************************************************** *)
+
+lemma converse_Memrel_not_wf_on: 
+    "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))"
+apply (unfold wf_on_def wf_def)
+apply (drule nat_le_infinite_Ord [THEN le_imp_subset], (assumption))
+apply (rule notI)
+apply (erule_tac x = "nat" in allE, blast)
+done
+
+lemma converse_Memrel_not_well_ord: 
+    "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))"
+apply (unfold well_ord_def)
+apply (blast dest: converse_Memrel_not_wf_on)
+done
+
+lemma well_ord_rvimage_ordertype:
+     "well_ord(A,r) \<Longrightarrow>
+       rvimage (ordertype(A,r), converse(ordermap(A,r)),r) =
+       Memrel(ordertype(A,r))" 
+by (blast intro: ordertype_ord_iso [THEN ord_iso_sym] ord_iso_rvimage_eq
+             Memrel_type [THEN subset_Int_iff [THEN iffD1]] trans)
+
+lemma well_ord_converse_Memrel:
+     "[| well_ord(A,r); well_ord(A,converse(r)) |]   
+      ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))" 
+apply (subst well_ord_rvimage_ordertype [symmetric], assumption) 
+apply (rule rvimage_converse [THEN subst])
+apply (blast intro: ordertype_ord_iso ord_iso_sym ord_iso_is_bij
+                    bij_is_inj well_ord_rvimage)
+done
+
+lemma WO1_imp_LEMMA: "WO1 ==> LEMMA"
+apply (unfold WO1_def LEMMA_def, clarify) 
+apply (blast dest: well_ord_converse_Memrel
+                   Ord_ordertype [THEN converse_Memrel_not_well_ord]
+	     intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite
+                    lepoll_def [THEN def_imp_iff, THEN iffD2] )
+done
+
+lemma WO1_iff_WO7: "WO1 <-> WO7"
+apply (simp add: WO7_iff_LEMMA)
+apply (blast intro: LEMMA_imp_WO1 WO1_imp_LEMMA)
+done
+
+
+
+(* ********************************************************************** *)
+(*            The proof of WO8 <-> WO1 (Rubin & Rubin p. 6)               *)
+(* ********************************************************************** *)
+
+lemma WO1_WO8: "WO1 ==> WO8"
+by (unfold WO1_def WO8_def, fast)
+
+
+(* The implication "WO8 ==> WO1": a faithful image of Rubin & Rubin's proof*)
+lemma WO8_WO1: "WO8 ==> WO1"
+apply (unfold WO1_def WO8_def)
+apply (rule allI)
+apply (erule_tac x = "{{x}. x \<in> A}" in allE)
+apply (erule impE)
+ apply (rule_tac x = "\<lambda>a \<in> {{x}. x \<in> A}. THE x. a={x}" in exI)
+ apply (force intro!: lam_type simp add: singleton_eq_iff the_equality)
+apply (blast intro: lam_sing_bij bij_is_inj well_ord_rvimage)
+done
 
 end
--- a/src/ZF/AC/WO2_AC16.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,591 +0,0 @@
-(*  Title:      ZF/AC/WO2_AC16.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-  The proof of WO2 ==> AC16(k #+ m, k)
-  
-  The main part of the proof is the inductive reasoning concerning
-  properties of constructed family T_gamma.
-  The proof deals with three cases for ordinals: 0, succ and limit ordinal.
-  The first instance is trivial, the third not difficult, but the second
-  is very complicated requiring many lemmas.
-  We also need to prove that at any stage gamma the set 
-  (s - Union(...) - k_gamma)   (Rubin & Rubin page 15)
-  contains m distinct elements (in fact is equipollent to s)
-*)
-
-(* ********************************************************************** *)
-(* case of limit ordinal                                                  *)
-(* ********************************************************************** *)
-
-
-Goal "[| \\<forall>y<x. \\<forall>z<a. z<y | (\\<exists>Y \\<in> F(y). f(z)<=Y)  \
-\               --> (\\<exists>! Y. Y \\<in> F(y) & f(z)<=Y);  \
-\               \\<forall>i j. i le j --> F(i) \\<subseteq> F(j); j le i; i<x; z<a;  \
-\               V \\<in> F(i); f(z)<=V; W \\<in> F(j); f(z)<=W |]  \
-\               ==> V = W";
-by (REPEAT (eresolve_tac [asm_rl, allE, impE] 1));
-by (dtac subsetD 1 THEN (assume_tac 1));
-by (REPEAT (dtac ospec 1 THEN (assume_tac 1)));
-by (eresolve_tac [disjI2 RSN (2, impE)] 1);
-by (fast_tac (FOL_cs addSIs [bexI]) 1);
-by (etac ex1_two_eq 1 THEN (REPEAT (ares_tac [conjI] 1)));
-val lemma3_1 = result();
-
-
-Goal "[| \\<forall>y<x. \\<forall>z<a. z<y | (\\<exists>Y \\<in> F(y). f(z)<=Y)  \
-\               --> (\\<exists>! Y. Y \\<in> F(y) & f(z)<=Y);  \
-\               \\<forall>i j. i le j --> F(i) \\<subseteq> F(j); i<x; j<x; z<a;  \
-\               V \\<in> F(i); f(z)<=V; W \\<in> F(j); f(z)<=W |]  \
-\               ==> V = W";
-by (res_inst_tac [("j","j")] ([lt_Ord, lt_Ord] MRS Ord_linear_le) 1
-    THEN (REPEAT (assume_tac 1)));
-by (eresolve_tac [lemma3_1 RS sym] 1 THEN (REPEAT (assume_tac 1)));
-by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1)));
-val lemma3 = result();
-
-
-Goal "[| \\<forall>y<x. F(y) \\<subseteq> X &  \
-\               (\\<forall>x<a. x < y | (\\<exists>Y \\<in> F(y). fa(x) \\<subseteq> Y) -->  \
-\                       (\\<exists>! Y. Y \\<in> F(y) & fa(x) \\<subseteq> Y)); x < a |]  \
-\               ==> \\<forall>y<x. \\<forall>z<a. z < y | (\\<exists>Y \\<in> F(y). fa(z) \\<subseteq> Y) -->  \
-\                       (\\<exists>! Y. Y \\<in> F(y) & fa(z) \\<subseteq> Y)";
-by (REPEAT (resolve_tac [oallI, impI] 1));
-by (dtac ospec 1 THEN (assume_tac 1));
-by (fast_tac (FOL_cs addSEs [oallE]) 1);
-val lemma4 = result();
-
-
-Goal "[| \\<forall>y<x. F(y) \\<subseteq> X &  \
-\               (\\<forall>x<a. x < y | (\\<exists>Y \\<in> F(y). fa(x) \\<subseteq> Y) -->  \
-\                       (\\<exists>! Y. Y \\<in> F(y) & fa(x) \\<subseteq> Y)); \
-\               x < a; Limit(x); \\<forall>i j. i le j --> F(i) \\<subseteq> F(j) |]  \
-\               ==> (\\<Union>x<x. F(x)) \\<subseteq> X &  \
-\               (\\<forall>xa<a. xa < x | (\\<exists>x \\<in> \\<Union>x<x. F(x). fa(xa) \\<subseteq> x)  \
-\               --> (\\<exists>! Y. Y \\<in> (\\<Union>x<x. F(x)) & fa(xa) \\<subseteq> Y))";
-by (rtac conjI 1);
-by (rtac subsetI 1);
-by (etac OUN_E 1);
-by (dtac ospec 1 THEN (assume_tac 1));
-by (Fast_tac 1);
-by (dtac lemma4 1 THEN (assume_tac 1));
-by (rtac oallI 1);
-by (rtac impI 1);
-by (etac disjE 1);
-by (forward_tac [Limit_has_succ RSN (2, ospec)] 1 THEN (REPEAT (assume_tac 1)));
-by (dres_inst_tac [("A","a"),("x","xa")] ospec 1 THEN (assume_tac 1));
-by (eresolve_tac [lt_Ord RS le_refl RSN (2, disjI1 RSN (2, impE))] 1
-        THEN (assume_tac 1));
-by (REPEAT (eresolve_tac [ex1E, conjE] 1));
-by (rtac ex1I 1);
-by (rtac conjI 1 THEN (assume_tac 2));
-by (eresolve_tac [Limit_has_succ RS OUN_I] 1 THEN (TRYALL assume_tac));
-by (REPEAT (eresolve_tac [conjE, OUN_E] 1));
-by (etac lemma3 1 THEN (TRYALL assume_tac));
-by (etac Limit_has_succ 1 THEN (assume_tac 1));
-by (etac bexE 1);
-by (rtac ex1I 1);
-by (etac conjI 1 THEN (assume_tac 1));
-by (REPEAT (eresolve_tac [conjE, OUN_E] 1));
-by (etac lemma3 1 THEN (TRYALL assume_tac));
-val lemma5 = result();
-
-(* ********************************************************************** *)
-(* case of successor ordinal                                              *)
-(* ********************************************************************** *)
-
-(*
-  First quite complicated proof of the fact used in the recursive construction
-  of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage
-  gamma the set (s - Union(...) - k_gamma) is equipollent to s
-  (Rubin & Rubin page 15).
-*)
-
-(* ********************************************************************** *)
-(* dbl_Diff_eqpoll_Card                                                   *)
-(* ********************************************************************** *)
-
-
-Goal "[| A eqpoll a; Card(a); ~Finite(a); B lesspoll a;  \
-\       C lesspoll a |] ==> A - B - C eqpoll a";
-by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
-by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1)));
-qed "dbl_Diff_eqpoll_Card";
-
-(* ********************************************************************** *)
-(* Case of finite ordinals                                                *)
-(* ********************************************************************** *)
-
-
-Goalw [lesspoll_def]
-        "[| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a";
-by (rtac conjI 1);
-by (dresolve_tac [nat_le_infinite_Ord RS le_imp_lepoll] 1
-        THEN (assume_tac 1));
-by (rewtac Finite_def);
-by (fast_tac (claset() addSEs [eqpoll_sym RS eqpoll_trans]) 2);
-by (rtac lepoll_trans 1 THEN (assume_tac 2));
-by (fast_tac (claset() addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_subset RS 
-        subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1);
-qed "Finite_lesspoll_infinite_Ord";
-
-Goal "[| \\<forall>x \\<in> X. x lepoll n & x \\<subseteq> T; well_ord(T, R); X lepoll b;  \
-\       b<a; ~Finite(a); Card(a); n \\<in> nat |]  \
-\       ==> Union(X) lesspoll a";
-by (excluded_middle_tac "Finite(X)" 1);
-by (resolve_tac [Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord)] 2
-        THEN (REPEAT (assume_tac 3)));
-by (fast_tac (claset() addSEs [lepoll_nat_imp_Finite]
-                addSIs [Finite_Union]) 2);
-by (dresolve_tac [lt_Ord RSN (2, lepoll_imp_ex_le_eqpoll)] 1 THEN (assume_tac 1));
-by (REPEAT (eresolve_tac [exE, conjE] 1));
-by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1));
-by (eresolve_tac [eqpoll_sym RS (eqpoll_def RS def_imp_iff RS iffD1) RS
-                exE] 1);
-by (forward_tac [bij_is_surj RS surj_image_eq] 1);
-by (dresolve_tac [[bij_is_fun, subset_refl] MRS image_fun] 1);
-by (dresolve_tac [sym RS trans] 1 THEN (assume_tac 1));
-by (blast_tac (claset() addIs [lesspoll_trans1, UN_lepoll, lt_Ord, 
-                               lt_trans1 RSN (2, lt_Card_imp_lesspoll)]) 1); 
-qed "Union_lesspoll";
-
-(* ********************************************************************** *)
-(* recfunAC16_lepoll_index                                                *)
-(* ********************************************************************** *)
-
-
-Goal "A Un {a} = cons(a, A)";
-by (Fast_tac 1);
-qed "Un_sing_eq_cons";
-
-
-Goal "A lepoll B ==> A Un {a} lepoll succ(B)";
-by (asm_simp_tac (simpset() addsimps [Un_sing_eq_cons, succ_def]) 1);
-by (eresolve_tac [mem_not_refl RSN (2, cons_lepoll_cong)] 1);
-qed "Un_lepoll_succ";
-
-
-Goal "Ord(a) ==> F(a) - (\\<Union>b<succ(a). F(b)) = 0";
-by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
-qed "Diff_UN_succ_empty";
-
-
-Goal "Ord(a) ==> F(a) Un X - (\\<Union>b<succ(a). F(b)) \\<subseteq> X";
-by (fast_tac (claset() addSIs [OUN_I, le_refl]) 1);
-qed "Diff_UN_succ_subset";
-
-
-Goal "Ord(x) ==>  \
-\       recfunAC16(f, g, x, a) - (\\<Union>i<x. recfunAC16(f, g, i, a)) lepoll 1";
-by (etac Ord_cases 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_0,
-				      empty_subsetI RS subset_imp_lepoll]) 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit, Diff_cancel, 
-				      empty_subsetI RS subset_imp_lepoll]) 2);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1);
-by (rtac conjI 1);
-by (fast_tac (claset() addSIs [empty_subsetI RS subset_imp_lepoll]
-                      addSEs [Diff_UN_succ_empty RS ssubst]) 1);
-by (fast_tac (claset() addSEs [Diff_UN_succ_subset RS subset_imp_lepoll RS
-        (singleton_eqpoll_1 RS eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
-qed "recfunAC16_Diff_lepoll_1";
-
-
-Goal "[| z \\<in> F(x); Ord(x) |]  \
-\       ==> z \\<in> F(LEAST i. z \\<in> F(i)) - (\\<Union>j<(LEAST i. z \\<in> F(i)). F(j))";
-by (fast_tac (claset() addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1);
-qed "in_Least_Diff";
-
-
-Goal "[| (LEAST i. w \\<in> F(i)) = (LEAST i. z \\<in> F(i));  \
-\       w \\<in> (\\<Union>i<a. F(i)); z \\<in> (\\<Union>i<a. F(i)) |]  \
-\       ==> \\<exists>b<a. w \\<in> (F(b) - (\\<Union>c<b. F(c))) & z \\<in> (F(b) - (\\<Union>c<b. F(c)))";
-by (REPEAT (etac OUN_E 1));
-by (dresolve_tac [lt_Ord RSN (2, in_Least_Diff)] 1 THEN (assume_tac 1));
-by (forward_tac [lt_Ord RSN (2, in_Least_Diff)] 1 THEN (assume_tac 1));
-by (rtac oexI 1);
-by (rtac conjI 1 THEN (assume_tac 2));
-by (etac subst 1 THEN (assume_tac 1));
-by (eresolve_tac [lt_Ord RSN (2, Least_le) RS lt_trans1] 1
-        THEN (REPEAT (assume_tac 1)));
-qed "Least_eq_imp_ex";
-
-
-Goal "[| A lepoll 1; a \\<in> A; b \\<in> A |] ==> a=b";
-by (fast_tac (claset() addSDs [lepoll_1_is_sing]) 1);
-qed "two_in_lepoll_1";
-
-
-Goal "[| \\<forall>i<a. F(i)-(\\<Union>j<i. F(j)) lepoll 1; Limit(a) |]  \
-\       ==> (\\<Union>x<a. F(x)) lepoll a";
-by (resolve_tac [lepoll_def RS (def_imp_iff RS iffD2)] 1);
-by (res_inst_tac [("x","\\<lambda>z \\<in> (\\<Union>x<a. F(x)). LEAST i. z \\<in> F(i)")] exI 1);
-by (rewtac inj_def);
-by (rtac CollectI 1);
-by (rtac lam_type 1);
-by (etac OUN_E 1);
-by (etac Least_in_Ord 1);
-by (etac ltD 1);
-by (etac lt_Ord2 1);
-by (rtac ballI 1);
-by (rtac ballI 1);
-by (Asm_simp_tac 1);
-by (rtac impI 1);
-by (dtac Least_eq_imp_ex 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (claset() addSEs [two_in_lepoll_1]) 1);
-qed "UN_lepoll_index";
-
-
-Goal "Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
-by (etac trans_induct 1);
-by (etac Ord_cases 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_0, lepoll_refl]) 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_succ]) 1);
-by (fast_tac (claset() 
-        addSDs [succI1 RSN (2, bspec)]
-        addSEs [subset_succI RS subset_imp_lepoll RSN (2, lepoll_trans),
-                Un_lepoll_succ]) 1);
-by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit]) 1);
-by (fast_tac (claset() addSEs [lt_Ord RS recfunAC16_Diff_lepoll_1]
-                       addSIs [UN_lepoll_index]) 1);
-qed "recfunAC16_lepoll_index";
-
-
-Goal "[| recfunAC16(f,g,y,a) \\<subseteq> {X \\<in> Pow(A). X eqpoll n};  \
-\        A eqpoll a;  y<a;  ~Finite(a);  Card(a);  n \\<in> nat |]  \
-\     ==> Union(recfunAC16(f,g,y,a)) lesspoll a";
-by (eresolve_tac [eqpoll_def RS def_imp_iff RS iffD1 RS exE] 1);
-by (rtac Union_lesspoll 1 THEN (TRYALL assume_tac));
-by (eresolve_tac [lt_Ord RS recfunAC16_lepoll_index] 3);
-by (eresolve_tac [[bij_is_inj, Card_is_Ord RS well_ord_Memrel] MRS
-		  well_ord_rvimage] 2 
-    THEN (assume_tac 2));
-by (fast_tac (claset() addSEs [eqpoll_imp_lepoll]) 1);
-qed "Union_recfunAC16_lesspoll";
-
-
-Goal "[| recfunAC16(f, fa, y, a) \\<subseteq> {X \\<in> Pow(A) . X eqpoll succ(k #+ m)};  \
-\       Card(a); ~ Finite(a); A eqpoll a;  \
-\       k \\<in> nat;  y<a;  \
-\       fa \\<in> bij(a, {Y \\<in> Pow(A). Y eqpoll succ(k)}) |]  \
-\       ==> A - Union(recfunAC16(f, fa, y, a)) - fa`y eqpoll a";
-by (rtac dbl_Diff_eqpoll_Card 1 THEN (TRYALL assume_tac));
-by (rtac Union_recfunAC16_lesspoll 1 THEN (REPEAT (assume_tac 1)));
-by (Simp_tac 1);
-by (resolve_tac [nat_succI RSN 
-		 (2, bexI RS (Finite_def RS def_imp_iff RS iffD2)) RS 
-		 (Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord))] 1
-        THEN (TRYALL assume_tac));
-by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type) RS CollectE] 1
-        THEN (TRYALL assume_tac));
-qed "dbl_Diff_eqpoll";
-
-(* back to the proof *)
-
-val disj_Un_eqpoll_nat_sum = 
-    [disj_Un_eqpoll_sum, sum_eqpoll_cong, nat_sum_eqpoll_sum] MRS 
-    (eqpoll_trans RS eqpoll_trans) |> standard;
-
-
-Goal "[| x \\<in> Pow(A - B - fa`i); x eqpoll m;  \
-\       fa \\<in> bij(a, {x \\<in> Pow(A) . x eqpoll k}); i<a; k \\<in> nat; m \\<in> nat |]  \
-\       ==> fa ` i Un x \\<in> {x \\<in> Pow(A) . x eqpoll k #+ m}";
-by (rtac CollectI 1);
-by (fast_tac (claset() 
-        addSEs [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)]) 1);
-by (rtac disj_Un_eqpoll_nat_sum 1
-        THEN (TRYALL assume_tac));
-by (fast_tac (claset() addSIs [equals0I]) 1);
-by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)] 1
-        THEN (REPEAT (assume_tac 1)));
-qed "Un_in_Collect";
-
-(* ********************************************************************** *)
-(* Lemmas simplifying assumptions                                         *)
-(* ********************************************************************** *)
-
-
-Goal "[| \\<forall>y<succ(j). F(y)<=X & (\\<forall>x<a. x<y | P(x,y)  \
-\       --> Q(x,y)); succ(j)<a |]  \
-\       ==> F(j)<=X & (\\<forall>x<a. x<j | P(x,j) --> Q(x,j))";
-by (dtac ospec 1);
-by (resolve_tac [lt_Ord RS (succI1 RS ltI RS lt_Ord RS le_refl)] 1
-        THEN (REPEAT (assume_tac 1)));
-val lemma6 = result();
-
-
-Goal "[| \\<forall>x<a. x<j | P(x,j) --> Q(x,j);  succ(j)<a |]  \
-\     ==> P(j,j) --> (\\<forall>x<a. x le j | P(x,j) --> Q(x,j))";
-by (fast_tac (claset() addSEs [leE]) 1);
-val lemma7 = result();
-
-(* ********************************************************************** *)
-(* Lemmas needded to prove ex_next_set which means that for any successor *)
-(* ordinal there is a set satisfying certain properties                   *)
-(* ********************************************************************** *)
-
-
-Goal "[| A eqpoll a; ~ Finite(a); Ord(a); m \\<in> nat |]  \
-\       ==> \\<exists>X \\<in> Pow(A). X eqpoll m";
-by (eresolve_tac [Ord_nat RSN (2, ltI) RS 
-                (nat_le_infinite_Ord RSN (2, lt_trans2)) RS 
-                leI RS le_imp_lepoll RS 
-                ((eqpoll_sym RS eqpoll_imp_lepoll) RSN (2, lepoll_trans)) RS 
-                lepoll_imp_eqpoll_subset RS exE] 1 
-        THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSEs [eqpoll_sym]) 1);
-qed "ex_subset_eqpoll";
-
-
-Goal "[| A \\<subseteq> B Un C; A Int C = 0 |] ==> A \\<subseteq> B";
-by (Blast_tac 1);
-qed "subset_Un_disjoint";
-
-
-Goal "[| X \\<in> Pow(A - Union(B) -C); T \\<in> B; F \\<subseteq> T |] ==> F Int X = 0";
-by (Blast_tac 1);
-qed "Int_empty";
-
-(* ********************************************************************** *)
-(* equipollent subset (and finite) is the whole set                       *)
-(* ********************************************************************** *)
-
-
-Goal "[| A \\<subseteq> B; a \\<in> A; A - {a} = B - {a} |] ==> A = B";
-by (fast_tac (claset() addSEs [equalityE]) 1);
-qed "Diffs_eq_imp_eq";
-
-
-Goal "m \\<in> nat ==> \\<forall>A B. A \\<subseteq> B & m lepoll A & B lepoll m --> A=B";
-by (induct_tac "m" 1);
-by (fast_tac (claset() addSDs [lepoll_0_is_0]) 1);
-by (REPEAT (resolve_tac [allI, impI] 1));
-by (REPEAT (etac conjE 1));
-by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1
-        THEN (assume_tac 1));
-by (forward_tac [subsetD RS Diff_sing_lepoll] 1
-        THEN REPEAT (assume_tac 1));
-by (ftac lepoll_Diff_sing 1);
-by (REPEAT (eresolve_tac [allE, impE] 1));
-by (rtac conjI 1);
-by (Fast_tac 2);
-by (Fast_tac 1);
-by (etac Diffs_eq_imp_eq 1
-        THEN REPEAT (assume_tac 1));
-qed "subset_imp_eq_lemma";
-
-
-Goal "[| A \\<subseteq> B; m lepoll A; B lepoll m; m \\<in> nat |] ==> A=B";
-by (fast_tac (FOL_cs addSDs [subset_imp_eq_lemma]) 1);
-qed "subset_imp_eq";
-
-
-Goal "[| f \\<in> bij(a, {Y \\<in> X. Y eqpoll succ(k)}); k \\<in> nat; f`b \\<subseteq> f`y; b<a;  \
-\       y<a |] ==> b=y";
-by (dtac subset_imp_eq 1);
-by (etac nat_succI 3);
-by (fast_tac (claset() addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
-                CollectE, eqpoll_sym RS eqpoll_imp_lepoll]) 1);
-by (fast_tac (claset() addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
-        CollectE, eqpoll_imp_lepoll]) 1);
-by (rewrite_goals_tac [bij_def, inj_def]);
-by (fast_tac (claset() addSDs [ltD]) 1);
-qed "bij_imp_arg_eq";
-
-
-Goal "[| recfunAC16(f, fa, y, a) \\<subseteq> {X \\<in> Pow(A) . X eqpoll succ(k #+ m)};  \
-\       Card(a); ~ Finite(a); A eqpoll a;  \
-\       k \\<in> nat; m \\<in> nat; y<a;  \
-\       fa \\<in> bij(a, {Y \\<in> Pow(A). Y eqpoll succ(k)});  \
-\       ~ (\\<exists>Y \\<in> recfunAC16(f, fa, y, a). fa`y \\<subseteq> Y) |]  \
-\       ==> \\<exists>X \\<in> {Y \\<in> Pow(A). Y eqpoll succ(k #+ m)}. fa`y \\<subseteq> X &  \
-\               (\\<forall>b<a. fa`b \\<subseteq> X -->  \
-\               (\\<forall>T \\<in> recfunAC16(f, fa, y, a). ~ fa`b \\<subseteq> T))";
-by (eresolve_tac [dbl_Diff_eqpoll RS ex_subset_eqpoll RS bexE] 1
-        THEN REPEAT (assume_tac 1));
-by (etac Card_is_Ord 1);
-by (ftac Un_in_Collect 2 THEN REPEAT (assume_tac 2));
-by (etac CollectE 4);
-by (rtac bexI 4);
-by (rtac CollectI 5);
-by (assume_tac 5);
-by (eresolve_tac [add_succ RS subst] 5);
-by (assume_tac 1);
-by (etac nat_succI 1);
-by (assume_tac 1);
-by (rtac conjI 1);
-by (Fast_tac 1);
-by (REPEAT (resolve_tac [ballI, impI, oallI, notI] 1));
-by (dresolve_tac [Int_empty RSN (2, subset_Un_disjoint)] 1
-        THEN REPEAT (assume_tac 1));
-by (dtac bij_imp_arg_eq 1 THEN REPEAT (assume_tac 1));
-by (hyp_subst_tac 1);
-by (eresolve_tac [bexI RSN (2, notE)] 1 THEN TRYALL assume_tac);
-qed "ex_next_set";
-
-(* ********************************************************************** *)
-(* Lemma ex_next_Ord states that for any successor                        *)
-(* ordinal there is a number of the set satisfying certain properties     *)
-(* ********************************************************************** *)
-
-
-Goal "[| recfunAC16(f, fa, y, a) \\<subseteq> {X \\<in> Pow(A) . X eqpoll succ(k #+ m)};  \
-\       Card(a); ~ Finite(a); A eqpoll a;  \
-\       k \\<in> nat; m \\<in> nat; y<a;  \
-\       fa \\<in> bij(a, {Y \\<in> Pow(A). Y eqpoll succ(k)});  \
-\       f \\<in> bij(a, {Y \\<in> Pow(A). Y eqpoll succ(k #+ m)});  \
-\       ~ (\\<exists>Y \\<in> recfunAC16(f, fa, y, a). fa`y \\<subseteq> Y) |]  \
-\       ==> \\<exists>c<a. fa`y \\<subseteq> f`c &  \
-\               (\\<forall>b<a. fa`b \\<subseteq> f`c -->  \
-\               (\\<forall>T \\<in> recfunAC16(f, fa, y, a). ~ fa`b \\<subseteq> T))";
-by (dtac ex_next_set 1 THEN REPEAT (assume_tac 1));
-by (etac bexE 1);
-by (resolve_tac [bij_converse_bij RS bij_is_fun RS apply_type RS ltI RSN
-        (2, oexI)] 1);
-by (resolve_tac [right_inverse_bij RS ssubst] 1
-        THEN REPEAT (ares_tac [Card_is_Ord] 1));
-qed "ex_next_Ord";
-
-
-Goal "[| \\<exists>! Y. Y \\<in> Z & P(Y); ~P(W) |] ==> \\<exists>! Y. Y \\<in> (Z Un {W}) & P(Y)";
-by (Fast_tac 1);
-qed "ex1_in_Un_sing";
-
-(* ********************************************************************** *)
-(* Lemma simplifying assumptions                                          *)
-(* ********************************************************************** *)
-
-
-Goal "[| \\<forall>x<a. x<j | (\\<exists>xa \\<in> F(j). P(x, xa))  \
-\       --> (\\<exists>! Y. Y \\<in> F(j) & P(x, Y)); F(j) \\<subseteq> X;  \
-\       L \\<in> X; P(j, L) & (\\<forall>x<a. P(x, L) --> (\\<forall>xa \\<in> F(j). ~P(x, xa))) |]  \
-\       ==> F(j) Un {L} \\<subseteq> X &  \
-\       (\\<forall>x<a. x le j | (\\<exists>xa \\<in> (F(j) Un {L}). P(x, xa)) -->  \
-\               (\\<exists>! Y. Y \\<in> (F(j) Un {L}) & P(x, Y)))";
-by (rtac conjI 1);
-by (fast_tac (claset() addSIs [singleton_subsetI]) 1);
-by (rtac oallI 1);
-by (etac oallE 1 THEN (contr_tac 2));
-by (blast_tac (claset() addSEs [leE]) 1);
-val lemma8 = result();
-
-(* ********************************************************************** *)
-(* The main part of the proof: inductive proof of the property of T_gamma *)
-(* lemma main_induct                                                      *)
-(* ********************************************************************** *)
-
-
-Goal "[| b < a; f \\<in> bij(a, {Y \\<in> Pow(A) . Y eqpoll succ(k #+ m)});  \
-\       fa \\<in> bij(a, {Y \\<in> Pow(A) . Y eqpoll succ(k)});  \
-\       ~Finite(a); Card(a); A eqpoll a; k \\<in> nat; m \\<in> nat |]  \
-\       ==> recfunAC16(f, fa, b, a) \\<subseteq> {X \\<in> Pow(A) . X eqpoll succ(k #+ m)} &  \
-\       (\\<forall>x<a. x < b | (\\<exists>Y \\<in> recfunAC16(f, fa, b, a). fa ` x \\<subseteq> Y) -->  \
-\       (\\<exists>! Y. Y \\<in> recfunAC16(f, fa, b, a) & fa ` x \\<subseteq> Y))";
-by (etac lt_induct 1);
-by (ftac lt_Ord 1);
-by (etac Ord_cases 1);
-(* case 0 *)
-by (asm_simp_tac (simpset() addsimps [recfunAC16_0]) 1);
-(* case Limit *)
-by (asm_simp_tac (simpset() addsimps [recfunAC16_Limit]) 2);
-by (rtac lemma5 2 THEN (REPEAT (assume_tac 2)));
-by (fast_tac (FOL_cs addSEs [recfunAC16_mono]) 2);
-(* case succ *)
-by (hyp_subst_tac 1);
-by (eresolve_tac [lemma6 RS conjE] 1 THEN (assume_tac 1));
-by (asm_simp_tac (simpset() delsplits [split_if]
-			    addsimps [recfunAC16_succ]) 1);
-by (resolve_tac [conjI RS (split_if RS iffD2)] 1);
-by (Asm_simp_tac 1);
-by (etac lemma7 1 THEN  assume_tac 1);
-by (rtac impI 1);
-by (resolve_tac [ex_next_Ord RS oexE] 1 
-    THEN REPEAT (ares_tac [le_refl RS lt_trans] 1));
-by (etac lemma8 1 THEN (assume_tac 1));
-by (resolve_tac [bij_is_fun RS apply_type] 1 THEN (assume_tac 1));
-by (eresolve_tac [Least_le RS lt_trans2 RS ltD] 1 
-        THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1));
-by (rtac (lt_Ord RSN (2, LeastI)) 1 THEN REPEAT (assume_tac 1));
-qed "main_induct";
-
-(* ********************************************************************** *)
-(* Lemma to simplify the inductive proof                                  *)
-(*   - the desired property is a consequence of the inductive assumption  *)
-(* ********************************************************************** *)
-
-val [prem1, prem2, prem3, prem4] = goal thy
-        "[| (!!b. b<a ==> F(b) \\<subseteq> S & (\\<forall>x<a. (x<b | (\\<exists>Y \\<in> F(b). f`x \\<subseteq> Y)) \
-\       --> (\\<exists>! Y. Y \\<in> F(b) & f`x \\<subseteq> Y)));  \
-\       f \\<in> a->f``(a); Limit(a); (!!i j. i le j ==> F(i) \\<subseteq> F(j)) |]  \
-\       ==> (\\<Union>j<a. F(j)) \\<subseteq> S &  \
-\       (\\<forall>x \\<in> f``a. \\<exists>! Y. Y \\<in> (\\<Union>j<a. F(j)) & x \\<subseteq> Y)";
-by (rtac conjI 1);
-by (rtac subsetI 1);
-by (etac OUN_E 1);
-by (dtac prem1 1);
-by (Fast_tac 1);
-(** LEVEL 5 **)
-by (rtac ballI 1);
-by (etac imageE 1);
-by (dresolve_tac [prem3 RS Limit_is_Ord RSN (2, ltI) RS
-        (prem3 RS Limit_has_succ)] 1);
-by (ftac prem1 1);
-by (etac conjE 1);
-(** LEVEL 10 **)
-by (dresolve_tac [leI RS succ_leE RSN (2, ospec)] 1 THEN (assume_tac 1));
-by (etac impE 1);
-by (fast_tac (claset() addSEs [leI RS succ_leE RS lt_Ord RS le_refl]) 1);
-by (dresolve_tac [prem2 RSN (2, apply_equality)] 1);
-by (REPEAT (eresolve_tac [conjE, ex1E] 1));
-(** LEVEL 15 **)
-by (rtac ex1I 1);
-by (fast_tac (claset() addSIs [OUN_I]) 1);
-by (REPEAT (eresolve_tac [conjE, OUN_E] 1));
-by (eresolve_tac [lt_Ord RSN (2, lt_Ord RS Ord_linear_le)] 1 
-    THEN assume_tac 1);
-by (dresolve_tac [prem4 RS subsetD] 2 THEN (assume_tac 2));
-(** LEVEL 20 **)
-by (fast_tac FOL_cs 2);
-by (ftac prem1 1);
-by (ftac succ_leE 1);
-by (dresolve_tac [prem4 RS subsetD] 1 THEN (assume_tac 1));
-by (etac conjE 1);
-(** LEVEL 25 **)
-by (dresolve_tac [lt_trans RSN (2, ospec)] 1 THEN (TRYALL assume_tac));
-by (dresolve_tac [disjI1 RSN (2, mp)] 1 THEN (assume_tac 1));
-by (etac ex1_two_eq 1);
-by (REPEAT (Fast_tac 1));
-qed "lemma_simp_induct";
-
-(* ********************************************************************** *)
-(* The target theorem                                                     *)
-(* ********************************************************************** *)
-
-
-Goalw [AC16_def] "[| WO2; 0<m; k \\<in> nat; m \\<in> nat |] ==> AC16(k #+ m,k)";
-by (rtac allI 1);
-by (rtac impI 1);
-by (ftac WO2_infinite_subsets_eqpoll_X 1 
-    THEN (REPEAT (assume_tac 1)));
-by (forw_inst_tac [("n","k #+ m")] (WO2_infinite_subsets_eqpoll_X) 1
-        THEN (REPEAT (ares_tac [add_type] 1)));
-by (ftac WO2_imp_ex_Card 1);
-by (REPEAT (eresolve_tac [exE,conjE] 1));
-by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS
-        def_imp_iff RS iffD1)] 1 THEN (assume_tac 1));
-by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS
-        def_imp_iff RS iffD1)] 1 THEN (assume_tac 1));
-by (REPEAT (etac exE 1));
-by (res_inst_tac [("x","\\<Union>j<a. recfunAC16(fa,f,j,a)")] exI 1);
-by (res_inst_tac [("P","%z. ?Y & (\\<forall>x \\<in> z. ?Z(x))")] 
-        (bij_is_surj RS surj_image_eq RS subst) 1
-        THEN (assume_tac 1));
-by (rtac lemma_simp_induct 1);
-by (eresolve_tac [bij_is_fun RS surj_image RS surj_is_fun] 2);
-by (eresolve_tac [eqpoll_imp_lepoll RS lepoll_infinite RS
-        infinite_Card_is_InfCard RS InfCard_is_Limit] 2 
-        THEN REPEAT (assume_tac 2));
-by (etac recfunAC16_mono 2);
-by (rtac main_induct 1 
-        THEN REPEAT (ares_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1));
-qed "WO2_AC16";
--- a/src/ZF/AC/WO2_AC16.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/WO2_AC16.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -1,3 +1,588 @@
-(*Dummy theory to document dependencies *)
+(*  Title:      ZF/AC/WO2_AC16.thy
+    ID:         $Id$
+    Author:     Krzysztof Grabczewski
+
+  The proof of WO2 ==> AC16(k #+ m, k)
+  
+  The main part of the proof is the inductive reasoning concerning
+  properties of constructed family T_gamma.
+  The proof deals with three cases for ordinals: 0, succ and limit ordinal.
+  The first instance is trivial, the third not difficult, but the second
+  is very complicated requiring many lemmas.
+  We also need to prove that at any stage gamma the set 
+  (s - Union(...) - k_gamma)   (Rubin & Rubin page 15)
+  contains m distinct elements (in fact is equipollent to s)
+*)
+
+theory WO2_AC16 = AC_Equiv + AC16_lemmas + Cardinal_aux:
+
+(**** A recursive definition used in the proof of WO2 ==> AC16 ****)
+
+constdefs
+  recfunAC16 :: "[i,i,i,i] => i"
+    "recfunAC16(f,h,i,a) == 
+         transrec2(i, 0, 
+              %g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
+                    else r Un {f`(LEAST i. h`g \<subseteq> f`i & 
+                         (\<forall>b<a. (h`b \<subseteq> f`i --> (\<forall>t \<in> r. ~ h`b \<subseteq> t))))})"
+
+(* ********************************************************************** *)
+(* Basic properties of recfunAC16                                         *)
+(* ********************************************************************** *)
+
+lemma recfunAC16_0: "recfunAC16(f,h,0,a) = 0"
+apply (simp add: recfunAC16_def);
+done
+
+lemma recfunAC16_succ: 
+     "recfunAC16(f,h,succ(i),a) =   
+      (if (\<exists>y \<in> recfunAC16(f,h,i,a). h ` i \<subseteq> y) then recfunAC16(f,h,i,a)  
+       else recfunAC16(f,h,i,a) Un  
+            {f ` (LEAST j. h ` i \<subseteq> f ` j &   
+             (\<forall>b<a. (h`b \<subseteq> f`j   
+              --> (\<forall>t \<in> recfunAC16(f,h,i,a). ~ h`b \<subseteq> t))))})"
+apply (simp add: recfunAC16_def);
+done
+
+lemma recfunAC16_Limit: "Limit(i)   
+        ==> recfunAC16(f,h,i,a) = (\<Union>j<i. recfunAC16(f,h,j,a))"
+by (simp add: recfunAC16_def transrec2_Limit);
+
+(* ********************************************************************** *)
+(* Monotonicity of transrec2                                              *)
+(* ********************************************************************** *)
+
+lemma transrec2_mono_lemma [rule_format]:
+     "[| !!g r. r \<subseteq> B(g,r);  Ord(i) |]   
+      ==> j<i --> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
+apply (erule trans_induct);
+apply (rule Ord_cases , assumption+);
+apply fast
+apply (simp (no_asm_simp))
+apply (blast elim!: leE); 
+apply (simp add: transrec2_Limit); 
+apply (blast intro: OUN_I ltI Ord_in_Ord [THEN le_refl]
+             elim!: Limit_has_succ [THEN ltE])
+done
+
+lemma transrec2_mono:
+     "[| !!g r. r \<subseteq> B(g,r); j\<le>i |] 
+      ==> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)"
+apply (erule leE);
+apply (rule transrec2_mono_lemma)
+apply (auto intro: lt_Ord2 ); 
+done
+
+(* ********************************************************************** *)
+(* Monotonicity of recfunAC16                                             *)
+(* ********************************************************************** *)
+
+lemma recfunAC16_mono: 
+       "i\<le>j ==> recfunAC16(f, g, i, a) \<subseteq> recfunAC16(f, g, j, a)"
+apply (unfold recfunAC16_def)
+apply (rule transrec2_mono)
+apply (auto ); 
+done
+
+(* ********************************************************************** *)
+(* case of limit ordinal                                                  *)
+(* ********************************************************************** *)
+
+
+lemma lemma3_1:
+    "[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) --> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+        \<forall>i j. i\<le>j --> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a;
+        V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |]   
+     ==> V = W"
+apply (erule asm_rl allE impE)+
+apply (drule subsetD, assumption)
+apply blast 
+done
+
+
+lemma lemma3:
+    "[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) --> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y);
+        \<forall>i j. i\<le>j --> F(i) \<subseteq> F(j); i<x; j<x; z<a;   
+        V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |]   
+     ==> V = W"
+apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+)
+apply (erule lemma3_1 [symmetric], assumption+)
+apply (erule lemma3_1, assumption+)
+done
+
+lemma lemma4:
+     "[| \<forall>y<x. F(y) \<subseteq> X &   
+                (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) -->   
+                 (\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y)); 
+         x < a |]   
+      ==> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) -->   
+                       (\<exists>! Y. Y \<in> F(y) & h(z) \<subseteq> Y)"
+apply (intro oallI impI)
+apply (drule ospec, assumption)
+apply clarify
+apply (blast elim!: oallE ) 
+done
+
+lemma lemma5:
+     "[| \<forall>y<x. F(y) \<subseteq> X &   
+               (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) -->   
+                (\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y));  
+         x < a; Limit(x); \<forall>i j. i\<le>j --> F(i) \<subseteq> F(j) |]   
+      ==> (\<Union>x<x. F(x)) \<subseteq> X &   
+          (\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x)   
+                --> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) & h(xa) \<subseteq> Y))"
+apply (rule conjI)
+apply (rule subsetI)
+apply (erule OUN_E)
+apply (drule ospec, assumption)
+apply fast
+apply (drule lemma4, assumption)
+apply (rule oallI)
+apply (rule impI)
+apply (erule disjE)
+apply (frule ospec, erule Limit_has_succ, assumption) 
+apply (drule_tac A = "a" and x = "xa" in ospec, assumption)
+apply (erule impE, rule le_refl [THEN disjI1], erule lt_Ord) 
+apply (blast intro: lemma3 Limit_has_succ) 
+apply (blast intro: lemma3) 
+done
+
+(* ********************************************************************** *)
+(* case of successor ordinal                                              *)
+(* ********************************************************************** *)
+
+(*
+  First quite complicated proof of the fact used in the recursive construction
+  of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage
+  gamma the set (s - Union(...) - k_gamma) is equipollent to s
+  (Rubin & Rubin page 15).
+*)
+
+(* ********************************************************************** *)
+(* dbl_Diff_eqpoll_Card                                                   *)
+(* ********************************************************************** *)
+
+
+lemma dbl_Diff_eqpoll_Card:
+      "[| A\<approx>a; Card(a); ~Finite(a); B\<prec>a; C\<prec>a |] ==> A - B - C\<approx>a"
+by (blast intro: Diff_lesspoll_eqpoll_Card) 
+
+(* ********************************************************************** *)
+(* Case of finite ordinals                                                *)
+(* ********************************************************************** *)
+
+
+lemma Finite_lesspoll_infinite_Ord: 
+      "[| Finite(X); ~Finite(a); Ord(a) |] ==> X\<prec>a"
+apply (unfold lesspoll_def)
+apply (rule conjI)
+apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption)
+apply (unfold Finite_def)
+apply (blast intro: leI [THEN le_imp_subset, THEN subset_imp_lepoll]
+                    ltI eqpoll_imp_lepoll lepoll_trans) 
+apply (blast intro: eqpoll_sym [THEN eqpoll_trans])
+done
+
+lemma Union_lesspoll:
+     "[| \<forall>x \<in> X. x lepoll n & x \<subseteq> T; well_ord(T, R); X lepoll b;   
+         b<a; ~Finite(a); Card(a); n \<in> nat |]   
+      ==> Union(X)\<prec>a"
+apply (case_tac "Finite (X)")
+apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord 
+                    lepoll_nat_imp_Finite Finite_Union)
+apply (drule lepoll_imp_ex_le_eqpoll) 
+apply (erule lt_Ord) 
+apply (elim exE conjE)
+apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)
+apply (erule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1],
+                         THEN exE])
+apply (frule bij_is_surj [THEN surj_image_eq])
+apply (drule image_fun [OF bij_is_fun subset_refl])
+apply (drule sym [THEN trans], assumption)
+apply (blast intro: lt_Ord UN_lepoll lt_Card_imp_lesspoll
+                    lt_trans1 lesspoll_trans1)
+done
+
+(* ********************************************************************** *)
+(* recfunAC16_lepoll_index                                                *)
+(* ********************************************************************** *)
+
+lemma Un_sing_eq_cons: "A Un {a} = cons(a, A)"
+by fast
+
+lemma Un_lepoll_succ: "A lepoll B ==> A Un {a} lepoll succ(B)"
+apply (simp add: Un_sing_eq_cons succ_def)
+apply (blast elim!: mem_irrefl intro: cons_lepoll_cong)
+done
+
+lemma Diff_UN_succ_empty: "Ord(a) ==> F(a) - (\<Union>b<succ(a). F(b)) = 0"
+by (fast intro!: le_refl)
+
+lemma Diff_UN_succ_subset: "Ord(a) ==> F(a) Un X - (\<Union>b<succ(a). F(b)) \<subseteq> X"
+by blast
+
+lemma recfunAC16_Diff_lepoll_1:
+     "Ord(x) 
+      ==> recfunAC16(f, g, x, a) - (\<Union>i<x. recfunAC16(f, g, i, a)) lepoll 1"
+apply (erule Ord_cases)
+  apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll])
+(*Limit case*)
+prefer 2 apply (simp add: recfunAC16_Limit Diff_cancel 
+                          empty_subsetI [THEN subset_imp_lepoll])
+(*succ case*)
+apply (simp add: recfunAC16_succ
+                 Diff_UN_succ_empty [of _ "%j. recfunAC16(f,g,j,a)"]
+                 empty_subsetI [THEN subset_imp_lepoll])
+apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll]
+                   singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans)
+done
+
+lemma in_Least_Diff:
+     "[| z \<in> F(x); Ord(x) |]   
+      ==> z \<in> F(LEAST i. z \<in> F(i)) - (\<Union>j<(LEAST i. z \<in> F(i)). F(j))"
+by (fast elim: less_LeastE elim!: LeastI);
+
+lemma Least_eq_imp_ex:
+     "[| (LEAST i. w \<in> F(i)) = (LEAST i. z \<in> F(i));   
+         w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i)) |]   
+      ==> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))"
+apply (elim OUN_E)
+apply (drule in_Least_Diff, erule lt_Ord) 
+apply (frule in_Least_Diff, erule lt_Ord) 
+apply (rule oexI, force) 
+apply (blast intro: lt_Ord Least_le [THEN lt_trans1]) 
+done
+
+
+lemma two_in_lepoll_1: "[| A lepoll 1; a \<in> A; b \<in> A |] ==> a=b"
+by (fast dest!: lepoll_1_is_sing)
+
+
+lemma UN_lepoll_index:
+     "[| \<forall>i<a. F(i)-(\<Union>j<i. F(j)) lepoll 1; Limit(a) |]   
+      ==> (\<Union>x<a. F(x)) lepoll a"
+apply (rule lepoll_def [THEN def_imp_iff [THEN iffD2]])
+apply (rule_tac x = "\<lambda>z \<in> (\<Union>x<a. F (x)). LEAST i. z \<in> F (i) " in exI)
+apply (unfold inj_def)
+apply (rule CollectI)
+apply (rule lam_type)
+apply (erule OUN_E)
+apply (erule Least_in_Ord)
+apply (erule ltD)
+apply (erule lt_Ord2)
+apply (intro ballI)
+apply (simp (no_asm_simp))
+apply (rule impI)
+apply (drule Least_eq_imp_ex, assumption+)
+apply (fast elim!: two_in_lepoll_1)
+done
+
+
+lemma recfunAC16_lepoll_index: "Ord(y) ==> recfunAC16(f, h, y, a) lepoll y"
+apply (erule trans_induct3)
+(*0 case*)
+apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl)
+(*succ case*)
+apply (simp (no_asm_simp) add: recfunAC16_succ)
+apply (blast dest!: succI1 [THEN rev_bspec] 
+             intro: subset_succI [THEN subset_imp_lepoll] Un_lepoll_succ  
+                    lepoll_trans)
+apply (simp (no_asm_simp) add: recfunAC16_Limit)
+apply (blast intro: lt_Ord [THEN recfunAC16_Diff_lepoll_1] UN_lepoll_index)
+done
+
 
-WO2_AC16 = AC_Equiv + recfunAC16 + AC16_lemmas + Cardinal_aux
+lemma Union_recfunAC16_lesspoll:
+     "[| recfunAC16(f,g,y,a) \<subseteq> {X \<in> Pow(A). X\<approx>n};   
+         A\<approx>a;  y<a;  ~Finite(a);  Card(a);  n \<in> nat |]   
+      ==> Union(recfunAC16(f,g,y,a))\<prec>a"
+apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE])
+apply (rule_tac T="A" in Union_lesspoll, simp_all)
+apply (blast intro!: eqpoll_imp_lepoll)
+apply (blast intro: bij_is_inj Card_is_Ord [THEN well_ord_Memrel]
+                    well_ord_rvimage)
+apply (erule lt_Ord [THEN recfunAC16_lepoll_index])
+done
+
+lemma dbl_Diff_eqpoll:
+     "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
+	 Card(a); ~ Finite(a); A\<approx>a;   
+	 k \<in> nat;  y<a;   
+	 h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}) |]   
+      ==> A - Union(recfunAC16(f, h, y, a)) - h`y\<approx>a"
+apply (rule dbl_Diff_eqpoll_Card, simp_all)
+apply (simp add: Union_recfunAC16_lesspoll)
+apply (rule Finite_lesspoll_infinite_Ord) 
+apply (rule Finite_def [THEN def_imp_iff, THEN iffD2]) 
+apply (blast dest: ltD bij_is_fun [THEN apply_type], assumption);  
+apply (blast intro: Card_is_Ord) 
+done;
+
+(* back to the proof *)
+
+lemmas disj_Un_eqpoll_nat_sum = 
+    eqpoll_trans [THEN eqpoll_trans, 
+                  OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum,
+                  standard];
+
+
+lemma Un_in_Collect: "[| x \<in> Pow(A - B - h`i); x\<approx>m;   
+        h \<in> bij(a, {x \<in> Pow(A) . x\<approx>k}); i<a; k \<in> nat; m \<in> nat |]   
+        ==> h ` i Un x \<in> {x \<in> Pow(A) . x\<approx>k #+ m}"
+by (blast intro: disj_Un_eqpoll_nat_sum
+          dest:  ltD bij_is_fun [THEN apply_type])
+
+
+(* ********************************************************************** *)
+(* Lemmas simplifying assumptions                                         *)
+(* ********************************************************************** *)
+
+lemma lemma6:
+     "[| \<forall>y<succ(j). F(y)<=X & (\<forall>x<a. x<y | P(x,y) --> Q(x,y)); succ(j)<a |]
+      ==> F(j)<=X & (\<forall>x<a. x<j | P(x,j) --> Q(x,j))"
+by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]); 
+
+
+lemma lemma7:
+     "[| \<forall>x<a. x<j | P(x,j) --> Q(x,j);  succ(j)<a |]   
+      ==> P(j,j) --> (\<forall>x<a. x\<le>j | P(x,j) --> Q(x,j))"
+by (fast elim!: leE)
+
+(* ********************************************************************** *)
+(* Lemmas needed to prove ex_next_set, which means that for any successor *)
+(* ordinal there is a set satisfying certain properties                   *)
+(* ********************************************************************** *)
+
+lemma ex_subset_eqpoll:
+     "[| A\<approx>a; ~ Finite(a); Ord(a); m \<in> nat |] ==> \<exists>X \<in> Pow(A). X\<approx>m"
+apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE]) 
+ apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll])
+  apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat)
+apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) 
+apply (fast elim!: eqpoll_sym)
+done
+
+lemma subset_Un_disjoint: "[| A \<subseteq> B Un C; A Int C = 0 |] ==> A \<subseteq> B"
+by blast
+
+
+lemma Int_empty:
+     "[| X \<in> Pow(A - Union(B) -C); T \<in> B; F \<subseteq> T |] ==> F Int X = 0"
+by blast
+
+
+(* ********************************************************************** *)
+(* equipollent subset (and finite) is the whole set                       *)
+(* ********************************************************************** *)
+
+
+lemma subset_imp_eq_lemma:
+     "m \<in> nat ==> \<forall>A B. A \<subseteq> B & m lepoll A & B lepoll m --> A=B"
+apply (induct_tac "m")
+apply (fast dest!: lepoll_0_is_0)
+apply (intro allI impI)
+apply (elim conjE)
+apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], assumption)
+apply (frule subsetD [THEN Diff_sing_lepoll], assumption+)
+apply (frule lepoll_Diff_sing)
+apply (erule allE impE)+
+apply (rule conjI)
+prefer 2 apply fast
+apply fast
+apply (blast elim: equalityE) 
+done
+
+
+lemma subset_imp_eq: "[| A \<subseteq> B; m lepoll A; B lepoll m; m \<in> nat |] ==> A=B"
+by (blast dest!: subset_imp_eq_lemma)
+
+
+lemma bij_imp_arg_eq:
+     "[| f \<in> bij(a, {Y \<in> X. Y\<approx>succ(k)}); k \<in> nat; f`b \<subseteq> f`y; b<a; y<a |] 
+      ==> b=y"
+apply (drule subset_imp_eq)
+apply (erule_tac [3] nat_succI)
+apply (unfold bij_def inj_def)
+apply (blast intro: eqpoll_sym eqpoll_imp_lepoll
+             dest:  ltD apply_type)+
+done
+
+
+lemma ex_next_set:
+     "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
+         Card(a); ~ Finite(a); A\<approx>a;   
+         k \<in> nat; m \<in> nat; y<a;   
+         h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
+         ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]   
+      ==> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X &   
+                (\<forall>b<a. h`b \<subseteq> X -->   
+                (\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
+apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], 
+       assumption+)
+apply (erule Card_is_Ord, assumption)
+apply (frule Un_in_Collect, (erule asm_rl nat_succI)+) 
+apply (erule CollectE)
+apply (rule rev_bexI, simp)
+apply (rule conjI, blast) 
+apply (intro ballI impI oallI notI)
+apply (drule subset_Un_disjoint, rule Int_empty, assumption+)
+apply (blast dest: bij_imp_arg_eq) 
+done
+
+(* ********************************************************************** *)
+(* Lemma ex_next_Ord states that for any successor                        *)
+(* ordinal there is a number of the set satisfying certain properties     *)
+(* ********************************************************************** *)
+
+lemma ex_next_Ord:
+     "[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
+	 Card(a); ~ Finite(a); A\<approx>a;   
+	 k \<in> nat; m \<in> nat; y<a;   
+	 h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
+	 f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});   
+	 ~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |]   
+      ==> \<exists>c<a. h`y \<subseteq> f`c &   
+	        (\<forall>b<a. h`b \<subseteq> f`c -->   
+		(\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))"
+apply (drule ex_next_set, assumption+)
+apply (erule bexE)
+apply (rule oexI)
+apply (subst right_inverse_bij, blast, assumption+)
+apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI
+                    Card_is_Ord)
+done
+
+
+(* ********************************************************************** *)
+(* Lemma simplifying assumptions                                          *)
+(* ********************************************************************** *)
+
+lemma lemma8:
+     "[| \<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa))   
+         --> (\<exists>! Y. Y \<in> F(j) & P(x, Y)); F(j) \<subseteq> X;   
+         L \<in> X; P(j, L) & (\<forall>x<a. P(x, L) --> (\<forall>xa \<in> F(j). ~P(x, xa))) |]   
+      ==> F(j) Un {L} \<subseteq> X &   
+          (\<forall>x<a. x\<le>j | (\<exists>xa \<in> (F(j) Un {L}). P(x, xa)) -->   
+                 (\<exists>! Y. Y \<in> (F(j) Un {L}) & P(x, Y)))"
+apply (rule conjI)
+apply (fast intro!: singleton_subsetI)
+apply (rule oallI)
+apply (blast elim!: leE oallE)
+done
+
+(* ********************************************************************** *)
+(* The main part of the proof: inductive proof of the property of T_gamma *)
+(* lemma main_induct                                                      *)
+(* ********************************************************************** *)
+
+lemma main_induct:
+     "[| b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)});   
+         h \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k)});   
+         ~Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat |]   
+      ==> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} &   
+          (\<forall>x<a. x < b | (\<exists>Y \<in> recfunAC16(f, h, b, a). h ` x \<subseteq> Y) -->   
+          (\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) & h ` x \<subseteq> Y))"
+apply (erule lt_induct)
+apply (frule lt_Ord)
+apply (erule Ord_cases)
+(* case 0 *)
+apply (simp add: recfunAC16_0)
+(* case Limit *)
+prefer 2  apply (simp add: recfunAC16_Limit)
+          apply (rule lemma5, assumption+)
+          apply (blast dest!: recfunAC16_mono)
+(* case succ *)
+apply clarify 
+apply (erule lemma6 [THEN conjE], assumption)
+apply (simp (no_asm_simp) split del: split_if add: recfunAC16_succ)
+apply (rule conjI [THEN split_if [THEN iffD2]])
+ apply (simp, erule lemma7, assumption)
+apply (rule impI)
+apply (rule ex_next_Ord [THEN oexE], 
+       assumption+, rule le_refl [THEN lt_trans], assumption+) 
+apply (erule lemma8, assumption)
+ apply (rule bij_is_fun [THEN apply_type], assumption)
+ apply (erule Least_le [THEN lt_trans2, THEN ltD])
+  apply (erule lt_Ord) 
+ apply (erule succ_leI)
+apply (erule LeastI) 
+apply (erule lt_Ord) 
+done
+
+(* ********************************************************************** *)
+(* Lemma to simplify the inductive proof                                  *)
+(*   - the desired property is a consequence of the inductive assumption  *)
+(* ********************************************************************** *)
+
+lemma lemma_simp_induct:
+     "[| \<forall>b. b<a --> F(b) \<subseteq> S & (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y))  
+                                   --> (\<exists>! Y. Y \<in> F(b) & f`x \<subseteq> Y));
+         f \<in> a->f``(a); Limit(a); 
+         \<forall>i j. i\<le>j --> F(i) \<subseteq> F(j) |]   
+        ==> (\<Union>j<a. F(j)) \<subseteq> S &   
+            (\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) & x \<subseteq> Y)"
+apply (rule conjI)
+apply (rule subsetI)
+apply (erule OUN_E, blast) 
+apply (rule ballI)
+apply (erule imageE)
+apply (drule ltI, erule Limit_is_Ord) 
+apply (drule Limit_has_succ, assumption) 
+apply (frule_tac x1="succ(xa)" in spec [THEN mp], assumption);
+apply (erule conjE)
+apply (drule ospec) 
+(** LEVEL 10 **)
+apply (erule leI [THEN succ_leE])
+apply (erule impE)
+apply (fast elim!: leI [THEN succ_leE, THEN lt_Ord, THEN le_refl])
+apply (drule apply_equality, assumption) 
+apply (elim conjE ex1E)
+(** LEVEL 15 **)
+apply (rule ex1I, blast) 
+apply (elim conjE OUN_E)
+apply (erule_tac i="succ(xa)" and j=aa 
+       in Ord_linear_le [OF lt_Ord lt_Ord], assumption)
+prefer 2 
+apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) 
+(** LEVEL 20 **)
+apply (drule_tac x1="aa" in spec [THEN mp], assumption)
+apply (frule succ_leE)
+apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) 
+done
+
+(* ********************************************************************** *)
+(* The target theorem                                                     *)
+(* ********************************************************************** *)
+
+theorem WO2_AC16: "[| WO2; 0<m; k \<in> nat; m \<in> nat |] ==> AC16(k #+ m,k)"
+apply (unfold AC16_def)
+apply (rule allI)
+apply (rule impI)
+apply (frule WO2_infinite_subsets_eqpoll_X, assumption+)
+apply (frule_tac n="k #+ m" in WO2_infinite_subsets_eqpoll_X, simp) 
+apply simp 
+apply (frule WO2_imp_ex_Card)
+apply (elim exE conjE)
+apply (drule eqpoll_trans [THEN eqpoll_sym, 
+                           THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]],
+       assumption)
+apply (drule eqpoll_trans [THEN eqpoll_sym, 
+                           THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], 
+       assumption+)
+apply (elim exE)
+apply (rename_tac h)
+apply (rule_tac x = "\<Union>j<a. recfunAC16 (h,f,j,a) " in exI)
+apply (rule_tac P="%z. ?Y & (\<forall>x \<in> z. ?Z (x))" 
+       in bij_is_surj [THEN surj_image_eq, THEN subst], assumption)
+apply (rule lemma_simp_induct)
+apply (blast del: conjI notI
+	     intro!: main_induct eqpoll_imp_lepoll [THEN lepoll_infinite] ) 
+apply (blast intro: bij_is_fun [THEN surj_image, THEN surj_is_fun])
+apply (erule eqpoll_imp_lepoll [THEN lepoll_infinite, 
+                                THEN infinite_Card_is_InfCard, 
+                                THEN InfCard_is_Limit], 
+       assumption+)
+apply (blast dest!: recfunAC16_mono)
+done
+
+end
--- a/src/ZF/AC/WO6_WO1.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,538 +0,0 @@
-(*  Title:      ZF/AC/WO6_WO1.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-  Proofs needed to state that formulations WO1,...,WO6 are all equivalent.
-  The only hard one is WO6 ==> WO1.
-
-  Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear"
-  by Rubin & Rubin (page 2). 
-  They refer reader to a book by Gödel to see the proof WO1 ==> WO2.
-  Fortunately order types made this proof also very easy.
-*)
-
-(* ********************************************************************** *)
-
-Goalw WO_defs "WO2 ==> WO3";
-by (Fast_tac 1);
-qed "WO2_WO3";
-
-(* ********************************************************************** *)
-
-Goalw (eqpoll_def::WO_defs) "WO3 ==> WO1";
-by (fast_tac (claset() addSEs [bij_is_inj RS well_ord_rvimage, 
-			      well_ord_Memrel RS well_ord_subset]) 1);
-qed "WO3_WO1";
-
-(* ********************************************************************** *)
-
-Goalw (eqpoll_def::WO_defs) "WO1 ==> WO2";
-by (fast_tac (claset() addSIs [Ord_ordertype, ordermap_bij]) 1);
-qed "WO1_WO2";
-
-(* ********************************************************************** *)
-
-Goal "f \\<in> A->B ==> (\\<lambda>x \\<in> A. {f`x}): A -> {{b}. b \\<in> B}";
-by (fast_tac (claset() addSIs [lam_type, apply_type]) 1);
-qed "lam_sets";
-
-Goalw [surj_def] "f \\<in> surj(A,B) ==> (\\<Union>a \\<in> A. {f`a}) = B";
-by (fast_tac (claset() addSEs [apply_type]) 1);
-qed "surj_imp_eq_";
-
-Goal "[| f \\<in> surj(A,B); Ord(A) |] ==> (\\<Union>a<A. {f`a}) = B";
-by (fast_tac (claset() addSDs [surj_imp_eq_]
-                addSIs [ltI] addSEs [ltE]) 1);
-qed "surj_imp_eq";
-
-Goalw WO_defs "WO1 ==> WO4(1)";
-by (rtac allI 1);
-by (eres_inst_tac [("x","A")] allE 1);
-by (etac exE 1);
-by (REPEAT (resolve_tac [exI, conjI] 1));
-by (etac Ord_ordertype 1);
-by (rtac conjI 1);
-by (eresolve_tac [ordermap_bij RS bij_converse_bij RS bij_is_fun RS
-                lam_sets RS domain_of_fun] 1);
-by (asm_simp_tac (simpset() addsimps [singleton_eqpoll_1 RS eqpoll_imp_lepoll,
-                  Ord_ordertype RSN (2, ordermap_bij RS bij_converse_bij RS
-                        bij_is_surj RS surj_imp_eq)]) 1);
-qed "WO1_WO4";
-
-(* ********************************************************************** *)
-
-Goalw WO_defs "[| m le n; WO4(m) |] ==> WO4(n)";
-by (blast_tac (claset() addSDs [spec] 
-                        addIs [le_imp_lepoll RSN (2, lepoll_trans)]) 1);
-qed "WO4_mono";
-
-(* ********************************************************************** *)
-
-Goalw WO_defs "[| m \\<in> nat; 1 le m; WO4(m) |] ==> WO5";
-by (Blast_tac 1);
-qed "WO4_WO5";
-
-(* ********************************************************************** *)
-
-Goalw WO_defs "WO5 ==> WO6";
-by (Blast_tac 1);
-qed "WO5_WO6";
-
-
-(* ********************************************************************** 
-    The proof of "WO6 ==> WO1".  Simplified by L C Paulson.
-
-    From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,
-    pages 2-5
-************************************************************************* *)
-
-goal OrderType.thy 
-      "!!i j k. [| k < i++j;  Ord(i);  Ord(j) |] ==>  \
-\                  k < i  |  (~ k<i & k = i ++ (k--i) & (k--i)<j)";
-by (res_inst_tac [("i","k"),("j","i")] Ord_linear2 1);
-by (dtac odiff_lt_mono2 4 THEN assume_tac 4);
-by (asm_full_simp_tac
-    (simpset() addsimps [oadd_odiff_inverse, odiff_oadd_inverse]) 4);
-by (safe_tac (claset() addSEs [lt_Ord]));
-qed "lt_oadd_odiff_disj";
-
-(*The corresponding elimination rule*)
-val lt_oadd_odiff_cases = rule_by_tactic Safe_tac
-                                         (lt_oadd_odiff_disj RS disjE);
-
-(* ********************************************************************** *)
-(* The most complicated part of the proof - lemma ii - p. 2-4             *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* some properties of relation uu(beta, gamma, delta) - p. 2              *)
-(* ********************************************************************** *)
-
-Goalw [uu_def] "domain(uu(f,b,g,d)) \\<subseteq> f`b";
-by (Blast_tac 1);
-qed "domain_uu_subset";
-
-Goal "\\<forall>b<a. f`b lepoll m ==> \
-\               \\<forall>b<a. \\<forall>g<a. \\<forall>d<a. domain(uu(f,b,g,d)) lepoll m";
-by (fast_tac (claset() addSEs
-        [domain_uu_subset RS subset_imp_lepoll RS lepoll_trans]) 1);
-qed "quant_domain_uu_lepoll_m";
-
-Goalw [uu_def] "uu(f,b,g,d) \\<subseteq> f`b * f`g";
-by (Blast_tac 1);
-qed "uu_subset1";
-
-Goalw [uu_def] "uu(f,b,g,d) \\<subseteq> f`d";
-by (Blast_tac 1);
-qed "uu_subset2";
-
-Goal "[| \\<forall>b<a. f`b lepoll m;  d<a |] ==> uu(f,b,g,d) lepoll m";
-by (fast_tac (claset()
-        addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1);
-qed "uu_lepoll_m";
-
-(* ********************************************************************** *)
-(* Two cases for lemma ii                                                 *)
-(* ********************************************************************** *)
-Goalw [lesspoll_def] 
-     "\\<forall>b<a. \\<forall>g<a. \\<forall>d<a. u(f,b,g,d) lepoll m  \
-\     ==> (\\<forall>b<a. f`b \\<noteq> 0 --> \
-\                 (\\<exists>g<a. \\<exists>d<a. u(f,b,g,d) \\<noteq> 0 & u(f,b,g,d) lesspoll m)) \
-\       | (\\<exists>b<a. f`b \\<noteq> 0 & (\\<forall>g<a. \\<forall>d<a. u(f,b,g,d) \\<noteq> 0 -->  \
-\                                       u(f,b,g,d) eqpoll m))";
-by (Asm_simp_tac 1);
-by (blast_tac (claset() delrules [equalityI]) 1);
-qed "cases";
-
-(* ********************************************************************** *)
-(* Lemmas used in both cases                                              *)
-(* ********************************************************************** *)
-Goal "Ord(a) ==> (\\<Union>b<a++a. C(b)) = (\\<Union>b<a. C(b) Un C(a++b))";
-by (fast_tac (claset() addSIs [equalityI] addIs [ltI] 
-                    addSDs [lt_oadd_disj]
-                    addSEs [lt_oadd1, oadd_lt_mono2]) 1);
-qed "UN_oadd";
-
-
-(* ********************************************************************** *)
-(* Case 1 \\<in> lemmas                                                        *)
-(* ********************************************************************** *)
-
-Goalw [vv1_def] "vv1(f,m,b) \\<subseteq> f`b";
-by (rtac (LetI RS LetI) 1);
-by (simp_tac (simpset() addsimps [domain_uu_subset]) 1);
-qed "vv1_subset";
-
-(* ********************************************************************** *)
-(* Case 1 \\<in> Union of images is the whole "y"                              *)
-(* ********************************************************************** *)
-Goalw [gg1_def]
-  "!! a f y. [| Ord(a);  m \\<in> nat |] ==>   \
-\            (\\<Union>b<a++a. gg1(f,a,m)`b) = (\\<Union>b<a. f`b)";
-by (asm_simp_tac
-    (simpset() addsimps [UN_oadd, lt_oadd1,
-                           oadd_le_self RS le_imp_not_lt, lt_Ord,
-                           odiff_oadd_inverse, ltD,
-                           vv1_subset RS Diff_partition, ww1_def]) 1);
-qed "UN_gg1_eq";
-
-Goal "domain(gg1(f,a,m)) = a++a";
-by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg1_def]) 1);
-qed "domain_gg1";
-
-(* ********************************************************************** *)
-(* every value of defined function is less than or equipollent to m       *)
-(* ********************************************************************** *)
-Goal "[| P(a, b);  Ord(a);  Ord(b);  \
-\               Least_a = (LEAST a. \\<exists>x. Ord(x) & P(a, x)) |]  \
-\               ==> P(Least_a, LEAST b. P(Least_a, b))";
-by (etac ssubst 1);
-by (res_inst_tac [("Q","%z. P(z, LEAST b. P(z, b))")] LeastI2 1);
-by (REPEAT (fast_tac (claset() addSEs [LeastI]) 1));
-qed "nested_LeastI";
-
-bind_thm ("nested_Least_instance",
-	  inst "P" 
-               "%g d. domain(uu(f,b,g,d)) \\<noteq> 0 & domain(uu(f,b,g,d)) lepoll m" 
-	  nested_LeastI);
-
-Goalw [gg1_def]
-    "!!a. [| Ord(a);  m \\<in> nat;  \
-\            \\<forall>b<a. f`b \\<noteq>0 -->                                       \
-\            (\\<exists>g<a. \\<exists>d<a. domain(uu(f,b,g,d)) \\<noteq> 0  &               \
-\                             domain(uu(f,b,g,d)) lepoll m);            \
-\            \\<forall>b<a. f`b lepoll succ(m);  b<a++a                       \
-\         |] ==> gg1(f,a,m)`b lepoll m";
-by (Asm_simp_tac 1);
-by (safe_tac (claset() addSEs [lt_oadd_odiff_cases]));
-(*Case b<a   \\<in> show vv1(f,m,b) lepoll m *)
-by (asm_simp_tac (simpset() addsimps [vv1_def, Let_def]) 1);
-by (fast_tac (claset() addIs [nested_Least_instance RS conjunct2]
-                addSEs [lt_Ord]
-                addSIs [empty_lepollI]) 1);
-(*Case a le b \\<in> show ww1(f,m,b--a) lepoll m *)
-by (asm_simp_tac (simpset() addsimps [ww1_def]) 1);
-by (excluded_middle_tac "f`(b--a) = 0" 1);
-by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2);
-by (rtac Diff_lepoll 1);
-by (Blast_tac 1);
-by (rtac vv1_subset 1);
-by (dtac (ospec RS mp) 1);
-by (REPEAT (eresolve_tac [asm_rl, oexE, conjE] 1));
-by (asm_simp_tac (simpset()
-        addsimps [vv1_def, Let_def, lt_Ord, 
-                  nested_Least_instance RS conjunct1]) 1);
-qed "gg1_lepoll_m";
-
-(* ********************************************************************** *)
-(* Case 2 \\<in> lemmas                                                        *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(* Case 2 \\<in> vv2_subset                                                    *)
-(* ********************************************************************** *)
-
-Goalw [uu_def] "[| b<a;  g<a;  f`b\\<noteq>0;  f`g\\<noteq>0;        \
-\                  y*y \\<subseteq> y;  (\\<Union>b<a. f`b)=y          \
-\               |] ==> \\<exists>d<a. uu(f,b,g,d) \\<noteq> 0";
-by (fast_tac (claset() addSIs [not_emptyI] 
-                    addSDs [SigmaI RSN (2, subsetD)]
-                    addSEs [not_emptyE]) 1);
-qed "ex_d_uu_not_empty";
-
-Goal "[| b<a; g<a; f`b\\<noteq>0; f`g\\<noteq>0;  \
-\        y*y \\<subseteq> y; (\\<Union>b<a. f`b)=y |]  \
-\     ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) \\<noteq> 0)) \\<noteq> 0";
-by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSEs [LeastI, lt_Ord]) 1);
-qed "uu_not_empty";
-
-Goal "[| r \\<subseteq> A*B; r\\<noteq>0 |] ==> domain(r)\\<noteq>0";
-by (Blast_tac 1);
-qed "not_empty_rel_imp_domain";
-
-Goal "[| b<a; g<a; f`b\\<noteq>0; f`g\\<noteq>0;  \
-\        y*y \\<subseteq> y; (\\<Union>b<a. f`b)=y |]  \
-\     ==> (LEAST d. uu(f,b,g,d) \\<noteq> 0) < a";
-by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1
-        THEN REPEAT (assume_tac 1));
-by (resolve_tac [Least_le RS lt_trans1] 1
-        THEN (REPEAT (ares_tac [lt_Ord] 1)));
-qed "Least_uu_not_empty_lt_a";
-
-Goal "[| B \\<subseteq> A; a\\<notin>B |] ==> B \\<subseteq> A-{a}";
-by (Blast_tac 1);
-qed "subset_Diff_sing";
-
-(*Could this be proved more directly?*)
-Goal "[| A lepoll m; m lepoll B; B \\<subseteq> A; m \\<in> nat |] ==> A=B";
-by (etac natE 1);
-by (fast_tac (claset() addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
-by (safe_tac (claset() addSIs [equalityI])); 
-by (rtac ccontr 1);
-by (etac (subset_Diff_sing RS subset_imp_lepoll 
-	  RSN (2, Diff_sing_lepoll RSN (3, lepoll_trans RS lepoll_trans)) RS 
-	  succ_lepoll_natE) 1
-        THEN REPEAT (assume_tac 1));
-qed "supset_lepoll_imp_eq";
-
-Goal "[| \\<forall>g<a. \\<forall>d<a. domain(uu(f, b, g, d))\\<noteq>0 -->               \
-\         domain(uu(f, b, g, d)) eqpoll succ(m);                        \
-\         \\<forall>b<a. f`b lepoll succ(m);  y*y \\<subseteq> y;                       \
-\         (\\<Union>b<a. f`b)=y;  b<a;  g<a;  d<a;                            \
-\         f`b\\<noteq>0;  f`g\\<noteq>0;  m \\<in> nat;  s \\<in> f`b                               \
-\      |] ==> uu(f, b, g, LEAST d. uu(f,b,g,d)\\<noteq>0) \\<in> f`b -> f`g";
-by (dres_inst_tac [("x2","g")] (ospec RS ospec RS mp) 1);
-by (rtac ([uu_subset1, uu_not_empty] MRS not_empty_rel_imp_domain) 3);
-by (rtac Least_uu_not_empty_lt_a 2 THEN TRYALL assume_tac);
-by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS 
-        (Least_uu_not_empty_lt_a RSN (2, uu_lepoll_m) RSN (2, 
-        uu_subset1 RSN (4, rel_is_fun)))] 1
-        THEN TRYALL assume_tac);
-by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1);
-by (REPEAT_FIRST assume_tac);
-by (REPEAT (fast_tac (claset() addSIs [domain_uu_subset]) 1));
-qed "uu_Least_is_fun";
-
-Goalw [vv2_def]
-    "!!a. [| \\<forall>g<a. \\<forall>d<a. domain(uu(f, b, g, d))\\<noteq>0 -->            \
-\            domain(uu(f, b, g, d)) eqpoll succ(m);                     \
-\            \\<forall>b<a. f`b lepoll succ(m); y*y \\<subseteq> y;                     \
-\            (\\<Union>b<a. f`b)=y;  b<a;  g<a;  m \\<in> nat;  s \\<in> f`b                \
-\          |] ==> vv2(f,b,g,s) \\<subseteq> f`g";
-by (split_tac [split_if] 1);
-by Safe_tac;
-by (etac (uu_Least_is_fun RS apply_type) 1);
-by (REPEAT_SOME (fast_tac (claset() addSIs [not_emptyI, singleton_subsetI])));
-qed "vv2_subset";
-
-(* ********************************************************************** *)
-(* Case 2 \\<in> Union of images is the whole "y"                              *)
-(* ********************************************************************** *)
-Goalw [gg2_def]
-    "!!a. [| \\<forall>g<a. \\<forall>d<a. domain(uu(f,b,g,d)) \\<noteq> 0 -->             \
-\            domain(uu(f,b,g,d)) eqpoll succ(m);                        \
-\            \\<forall>b<a. f`b lepoll succ(m); y*y \\<subseteq> y;                       \
-\            (\\<Union>b<a. f`b)=y;  Ord(a);  m \\<in> nat;  s \\<in> f`b;  b<a              \
-\         |] ==> (\\<Union>g<a++a. gg2(f,a,b,s) ` g) = y";
-by (dtac sym 1);
-by (asm_simp_tac
-    (simpset() addsimps [UN_oadd, lt_oadd1,
-                           oadd_le_self RS le_imp_not_lt, lt_Ord,
-                           odiff_oadd_inverse, ww2_def,
-                           vv2_subset RS Diff_partition]) 1);
-qed "UN_gg2_eq";
-
-Goal "domain(gg2(f,a,b,s)) = a++a";
-by (simp_tac (simpset() addsimps [lam_funtype RS domain_of_fun, gg2_def]) 1);
-qed "domain_gg2";
-
-(* ********************************************************************** *)
-(* every value of defined function is less than or equipollent to m       *)
-(* ********************************************************************** *)
-
-Goalw [vv2_def] "[| m \\<in> nat; m\\<noteq>0 |] ==> vv2(f,b,g,s) lepoll m";
-by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 1);
-by (fast_tac (claset()
-        addSDs [le_imp_subset RS subset_imp_lepoll RS lepoll_0_is_0]
-        addSIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
-                not_lt_imp_le RS le_imp_subset RS subset_imp_lepoll,
-                nat_into_Ord, nat_1I]) 1);
-qed "vv2_lepoll";
-
-Goalw [ww2_def]
-    "[| \\<forall>b<a. f`b lepoll succ(m);  g<a;  m \\<in> nat;  vv2(f,b,g,d) \\<subseteq> f`g |] \
-\    ==> ww2(f,b,g,d) lepoll m";
-by (excluded_middle_tac "f`g = 0" 1);
-by (asm_simp_tac (simpset() addsimps [empty_lepollI]) 2);
-by (dtac ospec 1 THEN (assume_tac 1));
-by (rtac Diff_lepoll 1 THEN (TRYALL assume_tac));
-by (asm_simp_tac (simpset() addsimps [vv2_def, not_emptyI]) 1);
-qed "ww2_lepoll";
-
-Goalw [gg2_def]
-    "!!a. [| \\<forall>g<a. \\<forall>d<a. domain(uu(f,b,g,d)) \\<noteq> 0 -->             \
-\            domain(uu(f,b,g,d)) eqpoll succ(m);                        \
-\            \\<forall>b<a. f`b lepoll succ(m);  y*y \\<subseteq> y;                    \
-\            (\\<Union>b<a. f`b)=y;  b<a;  s \\<in> f`b;  m \\<in> nat;  m\\<noteq> 0;  g<a++a     \
-\         |] ==> gg2(f,a,b,s) ` g lepoll m";
-by (Asm_simp_tac 1);
-by (safe_tac (claset() addSEs [lt_oadd_odiff_cases, lt_Ord2]));
-by (asm_simp_tac (simpset() addsimps [vv2_lepoll]) 1);
-by (asm_simp_tac (simpset() addsimps [ww2_lepoll, vv2_subset]) 1);
-qed "gg2_lepoll_m";
-
-(* ********************************************************************** *)
-(* lemma ii                                                               *)
-(* ********************************************************************** *)
-Goalw [NN_def] "[| succ(m) \\<in> NN(y); y*y \\<subseteq> y; m \\<in> nat; m\\<noteq>0 |] ==> m \\<in> NN(y)";
-by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1));
-by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1
-    THEN (assume_tac 1));
-(* case 1 *)
-by (asm_full_simp_tac (simpset() addsimps [lesspoll_succ_iff]) 1);
-by (res_inst_tac [("x","a++a")] exI 1);
-by (fast_tac (claset() addSIs [Ord_oadd, domain_gg1, UN_gg1_eq, 
-                                  gg1_lepoll_m]) 1);
-(* case 2 *)
-by (REPEAT (eresolve_tac [oexE, conjE] 1));
-by (res_inst_tac [("A","f`?B")] not_emptyE 1 THEN (assume_tac 1));
-by (rtac CollectI 1);
-by (etac succ_natD 1);
-by (res_inst_tac [("x","a++a")] exI 1);
-by (res_inst_tac [("x","gg2(f,a,b,x)")] exI 1);
-(*Calling fast_tac might get rid of the res_inst_tac calls, but it
-  is just too slow.*)
-by (asm_simp_tac (simpset() addsimps 
-                  [Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1);
-qed "lemma_ii";
-
-
-(* ********************************************************************** *)
-(* lemma iv - p. 4 \\<in>                                                      *)
-(* For every set x there is a set y such that   x Un (y * y) \\<subseteq> y         *)
-(* ********************************************************************** *)
-
-(* the quantifier \\<forall>looks inelegant but makes the proofs shorter  *)
-(* (used only in the following two lemmas)                          *)
-
-Goal "\\<forall>n \\<in> nat. rec(n, x, %k r. r Un r*r) \\<subseteq>  \
-\                    rec(succ(n), x, %k r. r Un r*r)";
-by (fast_tac (claset() addIs [rec_succ RS ssubst]) 1);
-qed "z_n_subset_z_succ_n";
-
-Goal "[| \\<forall>n \\<in> nat. f(n)<=f(succ(n)); n le m; n \\<in> nat; m \\<in> nat |]  \
-\     ==> f(n)<=f(m)";
-by (eres_inst_tac [("P","n le m")] rev_mp 1);
-by (res_inst_tac [("P","%z. n le z --> f(n) \\<subseteq> f(z)")] nat_induct 1);
-by (REPEAT (fast_tac le_cs 1));
-qed "le_subsets";
-
-Goal "[| n le m; m \\<in> nat |] ==>  \
-\       rec(n, x, %k r. r Un r*r) \\<subseteq> rec(m, x, %k r. r Un r*r)";
-by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 
-    THEN (TRYALL assume_tac));
-by (eresolve_tac [Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD] 1
-    THEN (assume_tac 1));
-qed "le_imp_rec_subset";
-
-Goal "\\<exists>y. x Un y*y \\<subseteq> y";
-by (res_inst_tac [("x","\\<Union>n \\<in> nat. rec(n, x, %k r. r Un r*r)")] exI 1);
-by Safe_tac;
-by (rtac (nat_0I RS UN_I) 1);
-by (Asm_simp_tac 1);
-by (res_inst_tac [("a","succ(n Un na)")] UN_I 1);
-by (eresolve_tac [Un_nat_type RS nat_succI] 1 THEN (assume_tac 1));
-by (fast_tac (ZF_cs addIs [le_imp_rec_subset RS subsetD]
-                addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type]
-                addSEs [nat_into_Ord] addss (simpset())) 1);
-qed "lemma_iv";
-
-(* ********************************************************************** *)
-(* Rubin & Rubin wrote \\<in>                                                  *)
-(* "It follows from (ii) and mathematical induction that if y*y \\<subseteq> y then *)
-(* y can be well-ordered"                                                 *)
-
-(* In fact we have to prove \\<in>                                             *)
-(*      * WO6 ==> NN(y) \\<noteq> 0                                              *)
-(*      * reverse induction which lets us infer that 1 \\<in> NN(y)            *)
-(*      * 1 \\<in> NN(y) ==> y can be well-ordered                             *)
-(* ********************************************************************** *)
-
-(* ********************************************************************** *)
-(*      WO6 ==> NN(y) \\<noteq> 0                                                *)
-(* ********************************************************************** *)
-
-Goalw [WO6_def, NN_def] "WO6 ==> NN(y) \\<noteq> 0";
-by (fast_tac ZF_cs 1);  (*SLOW if current claset is used*)
-qed "WO6_imp_NN_not_empty";
-
-(* ********************************************************************** *)
-(*      1 \\<in> NN(y) ==> y can be well-ordered                               *)
-(* ********************************************************************** *)
-
-Goal "[| (\\<Union>b<a. f`b)=y; x \\<in> y; \\<forall>b<a. f`b lepoll 1; Ord(a) |]  \
-\     ==> \\<exists>c<a. f`c = {x}";
-by (fast_tac (claset() addSEs [lepoll_1_is_sing]) 1);
-val lemma1 = result();
-
-Goal "[| (\\<Union>b<a. f`b)=y; x \\<in> y; \\<forall>b<a. f`b lepoll 1; Ord(a) |]  \
-\     ==> f` (LEAST i. f`i = {x}) = {x}";
-by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSEs [lt_Ord] addIs [LeastI]) 1);
-val lemma2 = result();
-
-Goalw [NN_def] "1 \\<in> NN(y) ==> \\<exists>a f. Ord(a) & f \\<in> inj(y, a)";
-by (etac CollectE 1);
-by (REPEAT (eresolve_tac [exE, conjE] 1));
-by (res_inst_tac [("x","a")] exI 1);
-by (res_inst_tac [("x","\\<lambda>x \\<in> y. LEAST i. f`i = {x}")] exI 1);
-by (rtac conjI 1 THEN (assume_tac 1));
-by (res_inst_tac [("d","%i. THE x. x \\<in> f`i")] lam_injective 1);
-by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (claset() addSEs [Least_le RS lt_trans1 RS ltD, lt_Ord]) 1);
-by (resolve_tac [lemma2 RS ssubst] 1 THEN REPEAT (assume_tac 1));
-by (Blast_tac 1);
-qed "NN_imp_ex_inj";
-
-Goal "[| y*y \\<subseteq> y; 1 \\<in> NN(y) |] ==> \\<exists>r. well_ord(y, r)";
-by (dtac NN_imp_ex_inj 1);
-by (fast_tac (claset() addSEs [well_ord_Memrel RSN (2,  well_ord_rvimage)]) 1);
-qed "y_well_ord";
-
-(* ********************************************************************** *)
-(*      reverse induction which lets us infer that 1 \\<in> NN(y)              *)
-(* ********************************************************************** *)
-
-val [prem1, prem2] = goal thy
-        "[| n \\<in> nat; !!m. [| m \\<in> nat; m\\<noteq>0; P(succ(m)) |] ==> P(m) |]  \
-\       ==> n\\<noteq>0 --> P(n) --> P(1)";
-by (rtac (prem1 RS nat_induct) 1);
-by (Blast_tac 1);
-by (excluded_middle_tac "x=0" 1);
-by (Blast_tac 2);
-by (fast_tac (claset() addSIs [prem2]) 1);
-qed "rev_induct_lemma";
-
-val prems = 
-Goal    "[| P(n); n \\<in> nat; n\\<noteq>0;  \
-\           !!m. [| m \\<in> nat; m\\<noteq>0; P(succ(m)) |] ==> P(m) |]  \
-\        ==> P(1)";
-by (resolve_tac [rev_induct_lemma RS impE] 1);
-by (etac impE 4 THEN (assume_tac 5));
-by (REPEAT (ares_tac prems 1));
-qed "rev_induct";
-
-Goalw [NN_def] "n \\<in> NN(y) ==> n \\<in> nat";
-by (etac CollectD1 1);
-qed "NN_into_nat";
-
-Goal "[| n \\<in> NN(y); y*y \\<subseteq> y; n\\<noteq>0 |] ==> 1 \\<in> NN(y)";
-by (rtac rev_induct 1 THEN REPEAT (ares_tac [NN_into_nat] 1));
-by (rtac lemma_ii 1 THEN REPEAT (assume_tac 1));
-val lemma3 = result();
-
-(* ********************************************************************** *)
-(* Main theorem "WO6 ==> WO1"                                             *)
-(* ********************************************************************** *)
-
-(* another helpful lemma *)
-Goalw [NN_def] "0 \\<in> NN(y) ==> y=0";
-by (fast_tac (claset() addSIs [equalityI] 
-                       addSDs [lepoll_0_is_0] addEs [subst]) 1);
-qed "NN_y_0";
-
-Goalw [WO1_def] "WO6 ==> WO1";
-by (rtac allI 1);
-by (excluded_middle_tac "A=0" 1);
-by (fast_tac (claset() addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2);
-by (res_inst_tac [("x1","A")] (lemma_iv RS revcut_rl) 1);
-by (etac exE 1);
-by (dtac WO6_imp_NN_not_empty 1);
-by (eresolve_tac [Un_subset_iff RS iffD1 RS conjE] 1);
-by (eres_inst_tac [("A","NN(y)")] not_emptyE 1);
-by (ftac y_well_ord 1);
-by (fast_tac (claset() addEs [well_ord_subset]) 2);
-by (fast_tac (claset() addSIs [lemma3] addSDs [NN_y_0] addSEs [not_emptyE]) 1);
-qed "WO6_imp_WO1";
-
--- a/src/ZF/AC/WO6_WO1.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/AC/WO6_WO1.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -2,49 +2,532 @@
     ID:         $Id$
     Author:     Krzysztof Grabczewski
 
-The proof of "WO6 ==> WO1".
+Proofs needed to state that formulations WO1,...,WO6 are all equivalent.
+The only hard one is WO6 ==> WO1.
 
-From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,
-pages 2-5
+Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear"
+by Rubin & Rubin (page 2). 
+They refer reader to a book by Gödel to see the proof WO1 ==> WO2.
+Fortunately order types made this proof also very easy.
 *)
 
-WO6_WO1 = "rel_is_fun" + AC_Equiv + Let +
+theory WO6_WO1 = Cardinal_aux:
 
-consts
+constdefs
 (* Auxiliary definitions used in proof *)
-  NN            :: i => i
-  uu            :: [i, i, i, i] => i
-  vv1, ww1, gg1 :: [i, i, i] => i
-  vv2, ww2, gg2 :: [i, i, i, i] => i
+  NN  :: "i => i"
+    "NN(y) == {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a  &  
+                        (\<Union>b<a. f`b) = y & (\<forall>b<a. f`b \<lesssim> m)}"
+  
+  uu  :: "[i, i, i, i] => i"
+    "uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta"
+
+  (** Definitions for case 1 **)
+  vv1 :: "[i, i, i] => i"
+     "vv1(f,m,b) ==                                                
+           let g = LEAST g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 & 
+                                 domain(uu(f,b,g,d)) \<lesssim> m));      
+               d = LEAST d. domain(uu(f,b,g,d)) \<noteq> 0 &                  
+                            domain(uu(f,b,g,d)) \<lesssim> m         
+           in  if f`b \<noteq> 0 then domain(uu(f,b,g,d)) else 0"
+
+  ww1 :: "[i, i, i] => i"
+     "ww1(f,m,b) == f`b - vv1(f,m,b)"
+
+  gg1 :: "[i, i, i] => i"
+     "gg1(f,a,m) == \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
+
+  (** Definitions for case 2 **)
+  vv2 :: "[i, i, i, i] => i"
+     "vv2(f,b,g,s) ==   
+           if f`g \<noteq> 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \<noteq> 0)`s} else 0"
+
+  ww2 :: "[i, i, i, i] => i"
+     "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)"
+
+  gg2 :: "[i, i, i, i] => i"
+     "gg2(f,a,b,s) ==
+	      \<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
+
+
+lemma WO2_WO3: "WO2 ==> WO3"
+by (unfold WO2_def WO3_def, fast)
+
+(* ********************************************************************** *)
+
+lemma WO3_WO1: "WO3 ==> WO1"
+apply (unfold eqpoll_def WO1_def WO3_def)
+apply (intro allI)
+apply (drule_tac x=A in spec) 
+apply (blast intro: bij_is_inj well_ord_rvimage 
+                    well_ord_Memrel [THEN well_ord_subset])
+done
+
+(* ********************************************************************** *)
+
+lemma WO1_WO2: "WO1 ==> WO2"
+apply (unfold eqpoll_def WO1_def WO2_def)
+apply (blast intro!: Ord_ordertype ordermap_bij)
+done
+
+(* ********************************************************************** *)
+
+lemma lam_sets: "f \<in> A->B ==> (\<lambda>x \<in> A. {f`x}): A -> {{b}. b \<in> B}"
+by (fast intro!: lam_type apply_type)
+
+lemma surj_imp_eq_: "f \<in> surj(A,B) ==> (\<Union>a \<in> A. {f`a}) = B"
+apply (unfold surj_def)
+apply (fast elim!: apply_type)
+done
+
+lemma surj_imp_eq: "[| f \<in> surj(A,B); Ord(A) |] ==> (\<Union>a<A. {f`a}) = B"
+by (fast dest!: surj_imp_eq_ intro!: ltI elim!: ltE)
+
+lemma WO1_WO4: "WO1 ==> WO4(1)"
+apply (unfold WO1_def WO4_def)
+apply (rule allI)
+apply (erule_tac x = "A" in allE)
+apply (erule exE)
+apply (intro exI conjI)
+apply (erule Ord_ordertype)
+apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN lam_sets, THEN domain_of_fun])
+apply (simp_all add: singleton_eqpoll_1 eqpoll_imp_lepoll Ord_ordertype
+     ordermap_bij [THEN bij_converse_bij, THEN bij_is_surj, THEN surj_imp_eq])
+done
+
+(* ********************************************************************** *)
+
+lemma WO4_mono: "[| m\<le>n; WO4(m) |] ==> WO4(n)"
+apply (unfold WO4_def)
+apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll])
+done
+
+(* ********************************************************************** *)
+
+lemma WO4_WO5: "[| m \<in> nat; 1\<le>m; WO4(m) |] ==> WO5"
+by (unfold WO4_def WO5_def, blast)
+
+(* ********************************************************************** *)
+
+lemma WO5_WO6: "WO5 ==> WO6"
+by (unfold WO4_def WO5_def WO6_def, blast)
+
+
+(* ********************************************************************** 
+    The proof of "WO6 ==> WO1".  Simplified by L C Paulson.
+
+    From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,
+    pages 2-5
+************************************************************************* *)
+
+lemma lt_oadd_odiff_disj:
+      "[| k < i++j;  Ord(i);  Ord(j) |] 
+       ==> k < i  |  (~ k<i & k = i ++ (k--i) & (k--i)<j)"
+apply (rule_tac i = "k" and j = "i" in Ord_linear2)
+prefer 4 
+   apply (drule odiff_lt_mono2, assumption)
+   apply (simp add: oadd_odiff_inverse odiff_oadd_inverse)
+apply (auto elim!: lt_Ord)
+done
+
+
+(* ********************************************************************** *)
+(* The most complicated part of the proof - lemma ii - p. 2-4             *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(* some properties of relation uu(beta, gamma, delta) - p. 2              *)
+(* ********************************************************************** *)
 
-defs
+lemma domain_uu_subset: "domain(uu(f,b,g,d)) \<subseteq> f`b"
+by (unfold uu_def, blast)
+
+lemma quant_domain_uu_lepoll_m:
+     "\<forall>b<a. f`b \<lesssim> m ==> \<forall>b<a. \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<lesssim> m"
+by (blast intro: domain_uu_subset [THEN subset_imp_lepoll] lepoll_trans)
+
+lemma uu_subset1: "uu(f,b,g,d) \<subseteq> f`b * f`g"
+by (unfold uu_def, blast)
+
+lemma uu_subset2: "uu(f,b,g,d) \<subseteq> f`d"
+by (unfold uu_def, blast)
+
+lemma uu_lepoll_m: "[| \<forall>b<a. f`b \<lesssim> m;  d<a |] ==> uu(f,b,g,d) \<lesssim> m"
+by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans)
+
+(* ********************************************************************** *)
+(* Two cases for lemma ii                                                 *)
+(* ********************************************************************** *)
+lemma cases: 
+     "\<forall>b<a. \<forall>g<a. \<forall>d<a. u(f,b,g,d) \<lesssim> m   
+      ==> (\<forall>b<a. f`b \<noteq> 0 -->  
+                  (\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 & u(f,b,g,d) \<prec> m))  
+        | (\<exists>b<a. f`b \<noteq> 0 & (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 -->   
+                                        u(f,b,g,d) \<approx> m))"
+apply (unfold lesspoll_def)
+apply (blast del: equalityI)
+done
+
+(* ********************************************************************** *)
+(* Lemmas used in both cases                                              *)
+(* ********************************************************************** *)
+lemma UN_oadd: "Ord(a) ==> (\<Union>b<a++a. C(b)) = (\<Union>b<a. C(b) Un C(a++b))"
+by (blast intro: ltI lt_oadd1 oadd_lt_mono2 dest!: lt_oadd_disj)
+
+
+(* ********************************************************************** *)
+(* Case 1: lemmas                                                        *)
+(* ********************************************************************** *)
+
+lemma vv1_subset: "vv1(f,m,b) \<subseteq> f`b"
+by (simp add: vv1_def Let_def domain_uu_subset)
+
+(* ********************************************************************** *)
+(* Case 1: Union of images is the whole "y"                              *)
+(* ********************************************************************** *)
+lemma UN_gg1_eq: 
+  "[| Ord(a);  m \<in> nat |] ==> (\<Union>b<a++a. gg1(f,a,m)`b) = (\<Union>b<a. f`b)"
+by (simp add: gg1_def UN_oadd lt_oadd1 oadd_le_self [THEN le_imp_not_lt] 
+              lt_Ord odiff_oadd_inverse ltD vv1_subset [THEN Diff_partition]
+              ww1_def)
+
+lemma domain_gg1: "domain(gg1(f,a,m)) = a++a"
+by (simp add: lam_funtype [THEN domain_of_fun] gg1_def)
 
-  NN_def  "NN(y) == {m \\<in> nat. \\<exists>a. \\<exists>f. Ord(a) & domain(f)=a  &  
-                            (\\<Union>b<a. f`b) = y & (\\<forall>b<a. f`b lepoll m)}"
+(* ********************************************************************** *)
+(* every value of defined function is less than or equipollent to m       *)
+(* ********************************************************************** *)
+lemma nested_LeastI:
+     "[| P(a, b);  Ord(a);  Ord(b);   
+         Least_a = (LEAST a. \<exists>x. Ord(x) & P(a, x)) |]   
+      ==> P(Least_a, LEAST b. P(Least_a, b))"
+apply (erule ssubst)
+apply (rule_tac Q = "%z. P (z, LEAST b. P (z, b))" in LeastI2)
+apply (fast elim!: LeastI)+
+done
+
+lemmas nested_Least_instance = 
+       nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \<noteq> 0 & 
+                                domain(uu(f,b,g,d)) \<lesssim> m", 
+                      standard]    (*generalizes the Variables!*)
 
-  uu_def  "uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta"
+lemma gg1_lepoll_m: 
+     "[| Ord(a);  m \<in> nat;   
+         \<forall>b<a. f`b \<noteq>0 -->                                        
+             (\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0  &                
+                          domain(uu(f,b,g,d)) \<lesssim> m);             
+         \<forall>b<a. f`b \<lesssim> succ(m);  b<a++a |] 
+      ==> gg1(f,a,m)`b \<lesssim> m"
+apply (unfold gg1_def, simp)
+apply (safe dest!: lt_oadd_odiff_disj)
+(*Case b<a   \<in> show vv1(f,m,b) \<lesssim> m *)
+apply (simp add: vv1_def Let_def)
+apply (fast intro: nested_Least_instance [THEN conjunct2]
+             elim!: lt_Ord intro!: empty_lepollI)
+(*Case a\<le>b \<in> show ww1(f,m,b--a) \<lesssim> m *)
+apply (simp add: ww1_def)
+apply (case_tac "f` (b--a) = 0", simp add: empty_lepollI)
+apply (rule Diff_lepoll, blast)
+apply (rule vv1_subset)
+apply (drule ospec [THEN mp], assumption+)
+apply (elim oexE conjE)
+apply (simp add: vv1_def Let_def lt_Ord nested_Least_instance [THEN conjunct1])
+done
+
+(* ********************************************************************** *)
+(* Case 2: lemmas                                                        *)
+(* ********************************************************************** *)
 
-  (*Definitions for case 1*)
+(* ********************************************************************** *)
+(* Case 2: vv2_subset                                                    *)
+(* ********************************************************************** *)
+
+lemma ex_d_uu_not_empty:
+     "[| b<a;  g<a;  f`b\<noteq>0;  f`g\<noteq>0;         
+         y*y \<subseteq> y;  (\<Union>b<a. f`b)=y |] 
+      ==> \<exists>d<a. uu(f,b,g,d) \<noteq> 0"
+by (unfold uu_def, blast) 
+
+lemma uu_not_empty:
+     "[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0;  y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]   
+      ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) \<noteq> 0)) \<noteq> 0"
+apply (drule ex_d_uu_not_empty, assumption+)
+apply (fast elim!: LeastI lt_Ord)
+done
+
+lemma not_empty_rel_imp_domain: "[| r \<subseteq> A*B; r\<noteq>0 |] ==> domain(r)\<noteq>0"
+by blast
+
+lemma Least_uu_not_empty_lt_a:
+     "[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]   
+      ==> (LEAST d. uu(f,b,g,d) \<noteq> 0) < a"
+apply (erule ex_d_uu_not_empty [THEN oexE], assumption+)
+apply (blast intro: Least_le [THEN lt_trans1] lt_Ord)
+done
+
+lemma subset_Diff_sing: "[| B \<subseteq> A; a\<notin>B |] ==> B \<subseteq> A-{a}"
+by blast
 
-  vv1_def "vv1(f,m,b) ==                                                
-           let g = LEAST g. (\\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \\<noteq> 0 & 
-                                   domain(uu(f,b,g,d)) lepoll m));      
-               d = LEAST d. domain(uu(f,b,g,d)) \\<noteq> 0 &                  
-                           domain(uu(f,b,g,d)) lepoll m         
-           in  if f`b \\<noteq> 0 then domain(uu(f,b,g,d)) else 0"
+(*Could this be proved more directly?*)
+lemma supset_lepoll_imp_eq:
+     "[| A \<lesssim> m; m \<lesssim> B; B \<subseteq> A; m \<in> nat |] ==> A=B"
+apply (erule natE)
+apply (fast dest!: lepoll_0_is_0 intro!: equalityI)
+apply (safe intro!: equalityI)
+apply (rule ccontr)
+apply (rule succ_lepoll_natE)
+ apply (erule lepoll_trans)  
+ apply (rule lepoll_trans)  
+  apply (erule subset_Diff_sing [THEN subset_imp_lepoll], assumption)
+ apply (rule Diff_sing_lepoll, assumption+) 
+done
+
+lemma uu_Least_is_fun:
+     "[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 -->                
+          domain(uu(f, b, g, d)) \<approx> succ(m);                         
+          \<forall>b<a. f`b \<lesssim> succ(m);  y*y \<subseteq> y;                        
+          (\<Union>b<a. f`b)=y;  b<a;  g<a;  d<a;                             
+          f`b\<noteq>0;  f`g\<noteq>0;  m \<in> nat;  s \<in> f`b |] 
+      ==> uu(f, b, g, LEAST d. uu(f,b,g,d)\<noteq>0) \<in> f`b -> f`g"
+apply (drule_tac x2=g in ospec [THEN ospec, THEN mp])
+   apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty])
+        apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+)
+apply (rule rel_is_fun)
+    apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) 
+   apply (erule uu_lepoll_m) 
+   apply (rule Least_uu_not_empty_lt_a, assumption+) 
+apply (rule uu_subset1)
+apply (rule supset_lepoll_imp_eq [OF _ eqpoll_sym [THEN eqpoll_imp_lepoll]])
+apply (fast intro!: domain_uu_subset)+
+done
+
+lemma vv2_subset: 
+     "[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 -->             
+		       domain(uu(f, b, g, d)) \<approx> succ(m);
+	 \<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;
+	 (\<Union>b<a. f`b)=y;  b<a;  g<a;  m \<in> nat;  s \<in> f`b |] 
+      ==> vv2(f,b,g,s) \<subseteq> f`g"
+apply (simp add: vv2_def)
+apply (blast intro: uu_Least_is_fun [THEN apply_type])
+done
+
+(* ********************************************************************** *)
+(* Case 2: Union of images is the whole "y"                              *)
+(* ********************************************************************** *)
+lemma UN_gg2_eq: 
+     "[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 -->              
+               domain(uu(f,b,g,d)) \<approx> succ(m);                         
+         \<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;                        
+         (\<Union>b<a. f`b)=y;  Ord(a);  m \<in> nat;  s \<in> f`b;  b<a |] 
+      ==> (\<Union>g<a++a. gg2(f,a,b,s) ` g) = y"
+apply (unfold gg2_def)
+apply (drule sym)
+apply (simp add: UN_oadd lt_oadd1 oadd_le_self [THEN le_imp_not_lt] 
+                 lt_Ord odiff_oadd_inverse ww2_def 
+                 vv2_subset [THEN Diff_partition])
+done
+
+lemma domain_gg2: "domain(gg2(f,a,b,s)) = a++a"
+by (simp add: lam_funtype [THEN domain_of_fun] gg2_def)
+
 
-  ww1_def "ww1(f,m,b) == f`b - vv1(f,m,b)"
+(* ********************************************************************** *)
+(* every value of defined function is less than or equipollent to m       *)
+(* ********************************************************************** *)
+
+lemma vv2_lepoll: "[| m \<in> nat; m\<noteq>0 |] ==> vv2(f,b,g,s) \<lesssim> m"
+apply (unfold vv2_def)
+apply (simp add: empty_lepollI)
+apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0] 
+       intro!: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans]
+               not_lt_imp_le [THEN le_imp_subset, THEN subset_imp_lepoll]
+               nat_into_Ord nat_1I)
+done
+
+lemma ww2_lepoll: 
+    "[| \<forall>b<a. f`b \<lesssim> succ(m);  g<a;  m \<in> nat;  vv2(f,b,g,d) \<subseteq> f`g |]  
+     ==> ww2(f,b,g,d) \<lesssim> m"
+apply (unfold ww2_def)
+apply (case_tac "f`g = 0")
+apply (simp add: empty_lepollI)
+apply (drule ospec, assumption)
+apply (rule Diff_lepoll, assumption+)
+apply (simp add: vv2_def not_emptyI)
+done
+
+lemma gg2_lepoll_m: 
+     "[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 -->              
+                      domain(uu(f,b,g,d)) \<approx> succ(m);                         
+         \<forall>b<a. f`b \<lesssim> succ(m);  y*y \<subseteq> y;                     
+         (\<Union>b<a. f`b)=y;  b<a;  s \<in> f`b;  m \<in> nat;  m\<noteq> 0;  g<a++a |] 
+      ==> gg2(f,a,b,s) ` g \<lesssim> m"
+apply (unfold gg2_def, simp)
+apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj)
+ apply (simp add: vv2_lepoll)
+apply (simp add: ww2_lepoll vv2_subset)
+done
+
+(* ********************************************************************** *)
+(* lemma ii                                                               *)
+(* ********************************************************************** *)
+lemma lemma_ii: "[| succ(m) \<in> NN(y); y*y \<subseteq> y; m \<in> nat; m\<noteq>0 |] ==> m \<in> NN(y)"
+apply (unfold NN_def)
+apply (elim CollectE exE conjE) 
+apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption)
+(* case 1 *)
+apply (simp add: lesspoll_succ_iff)
+apply (rule_tac x = "a++a" in exI)
+apply (fast intro!: Ord_oadd domain_gg1 UN_gg1_eq gg1_lepoll_m)
+(* case 2 *)
+apply (elim oexE conjE)
+apply (rule_tac A = "f`?B" in not_emptyE, assumption)
+apply (rule CollectI)
+apply (erule succ_natD)
+apply (rule_tac x = "a++a" in exI)
+apply (rule_tac x = "gg2 (f,a,b,x) " in exI)
+(*Calling fast_tac might get rid of the res_inst_tac calls, but it
+  is just too slow.*)
+apply (simp add: Ord_oadd domain_gg2 UN_gg2_eq gg2_lepoll_m)
+done
+
+
+(* ********************************************************************** *)
+(* lemma iv - p. 4:                                                       *)
+(* For every set x there is a set y such that   x Un (y * y) \<subseteq> y         *)
+(* ********************************************************************** *)
+
+
+(* The leading \<forall>-quantifier looks odd but makes the proofs shorter 
+   (used only in the following two lemmas)                          *)
 
-  gg1_def "gg1(f,a,m) == \\<lambda>b \\<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
-  
-  (*Definitions for case 2*)
+lemma z_n_subset_z_succ_n:
+     "\<forall>n \<in> nat. rec(n, x, %k r. r Un r*r) \<subseteq> rec(succ(n), x, %k r. r Un r*r)"
+by (fast intro: rec_succ [THEN ssubst])
+
+lemma le_subsets:
+     "[| \<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat |]   
+      ==> f(n)<=f(m)"
+apply (erule_tac P = "n\<le>m" in rev_mp)
+apply (rule_tac P = "%z. n\<le>z --> f (n) \<subseteq> f (z) " in nat_induct) 
+apply (auto simp add: le_iff) 
+done
+
+lemma le_imp_rec_subset:
+     "[| n\<le>m; m \<in> nat |] 
+      ==> rec(n, x, %k r. r Un r*r) \<subseteq> rec(m, x, %k r. r Un r*r)"
+apply (rule z_n_subset_z_succ_n [THEN le_subsets])
+apply (blast intro: lt_nat_in_nat)+
+done
+
+lemma lemma_iv: "\<exists>y. x Un y*y \<subseteq> y"
+apply (rule_tac x = "\<Union>n \<in> nat. rec (n, x, %k r. r Un r*r) " in exI)
+apply safe
+apply (rule nat_0I [THEN UN_I], simp)
+apply (rule_tac a = "succ (n Un na) " in UN_I)
+apply (erule Un_nat_type [THEN nat_succI], assumption)
+apply (auto intro: le_imp_rec_subset [THEN subsetD] 
+            intro!: Un_upper1_le Un_upper2_le Un_nat_type 
+            elim!: nat_into_Ord)
+done
 
-  vv2_def "vv2(f,b,g,s) ==   
-           if f`g \\<noteq> 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \\<noteq> 0)`s} else 0"
+(* ********************************************************************** *)
+(* Rubin & Rubin wrote,                                                   *)
+(* "It follows from (ii) and mathematical induction that if y*y \<subseteq> y then *)
+(* y can be well-ordered"                                                 *)
+
+(* In fact we have to prove                                               *)
+(*      * WO6 ==> NN(y) \<noteq> 0                                              *)
+(*      * reverse induction which lets us infer that 1 \<in> NN(y)            *)
+(*      * 1 \<in> NN(y) ==> y can be well-ordered                             *)
+(* ********************************************************************** *)
+
+(* ********************************************************************** *)
+(*      WO6 ==> NN(y) \<noteq> 0                                                *)
+(* ********************************************************************** *)
+
+lemma WO6_imp_NN_not_empty: "WO6 ==> NN(y) \<noteq> 0"
+apply (unfold WO6_def NN_def, clarify) 
+apply blast
+done
+
+(* ********************************************************************** *)
+(*      1 \<in> NN(y) ==> y can be well-ordered                               *)
+(* ********************************************************************** *)
+
+lemma lemma1:
+     "[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |] ==> \<exists>c<a. f`c = {x}"
+by (fast elim!: lepoll_1_is_sing)
+
+lemma lemma2:
+     "[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |]   
+      ==> f` (LEAST i. f`i = {x}) = {x}"
+apply (drule lemma1, assumption+)
+apply (fast elim!: lt_Ord intro: LeastI)
+done
 
-  ww2_def "ww2(f,b,g,s) == f`g - vv2(f,b,g,s)"
+lemma NN_imp_ex_inj: "1 \<in> NN(y) ==> \<exists>a f. Ord(a) & f \<in> inj(y, a)"
+apply (unfold NN_def)
+apply (elim CollectE exE conjE)
+apply (rule_tac x = "a" in exI)
+apply (rule_tac x = "\<lambda>x \<in> y. LEAST i. f`i = {x}" in exI)
+apply (rule conjI, assumption)
+apply (rule_tac d = "%i. THE x. x \<in> f`i" in lam_injective)
+apply (drule lemma1, assumption+)
+apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord)
+apply (rule lemma2 [THEN ssubst], assumption+, blast)
+done
+
+lemma y_well_ord: "[| y*y \<subseteq> y; 1 \<in> NN(y) |] ==> \<exists>r. well_ord(y, r)"
+apply (drule NN_imp_ex_inj)
+apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel])
+done
+
+(* ********************************************************************** *)
+(*      reverse induction which lets us infer that 1 \<in> NN(y)              *)
+(* ********************************************************************** *)
+
+lemma rev_induct_lemma [rule_format]: 
+     "[| n \<in> nat; !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]   
+      ==> n\<noteq>0 --> P(n) --> P(1)"
+by (erule nat_induct, blast+)
+
+lemma rev_induct:
+     "[| n \<in> nat;  P(n);  n\<noteq>0;   
+         !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]   
+      ==> P(1)"
+by (rule rev_induct_lemma, blast+)
 
-  gg2_def "gg2(f,a,b,s) ==
-	      \\<lambda>g \\<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
-  
+lemma NN_into_nat: "n \<in> NN(y) ==> n \<in> nat"
+by (simp add: NN_def)
+
+lemma lemma3: "[| n \<in> NN(y); y*y \<subseteq> y; n\<noteq>0 |] ==> 1 \<in> NN(y)"
+apply (rule rev_induct [OF NN_into_nat], assumption+)
+apply (rule lemma_ii, assumption+)
+done
+
+(* ********************************************************************** *)
+(* Main theorem "WO6 ==> WO1"                                             *)
+(* ********************************************************************** *)
+
+(* another helpful lemma *)
+lemma NN_y_0: "0 \<in> NN(y) ==> y=0"
+apply (unfold NN_def) 
+apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst)
+done
+
+lemma WO6_imp_WO1: "WO6 ==> WO1"
+apply (unfold WO1_def)
+apply (rule allI)
+apply (case_tac "A=0")
+apply (fast intro!: well_ord_Memrel nat_0I [THEN nat_into_Ord])
+apply (rule_tac x1 = "A" in lemma_iv [THEN revcut_rl])
+apply (erule exE)
+apply (drule WO6_imp_NN_not_empty)
+apply (erule Un_subset_iff [THEN iffD1, THEN conjE])
+apply (erule_tac A = "NN (y) " in not_emptyE)
+apply (frule y_well_ord)
+ apply (fast intro!: lemma3 dest!: NN_y_0 elim!: not_emptyE)
+apply (fast elim: well_ord_subset)
+done
+
 end
--- a/src/ZF/AC/WO_AC.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,20 +0,0 @@
-(*  Title:      ZF/AC/WO_AC.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Lemmas used in the proofs like WO? ==> AC?
-*)
-
-Goal "[| well_ord(Union(A),r); 0\\<notin>A; B \\<in> A |]  \
-\               ==> (THE b. first(b,B,r)) \\<in> B";
-by (fast_tac (claset() addSEs [well_ord_imp_ex1_first RS theI RS
-                (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
-qed "first_in_B";
-
-Goal "[| well_ord(Union(A), R); 0\\<notin>A |] ==> \\<exists>f. f:(\\<Pi>X \\<in> A. X)";
-by (fast_tac (claset() addSEs [first_in_B] addSIs [lam_type]) 1);
-qed "ex_choice_fun";
-
-Goal "well_ord(A, R) ==> \\<exists>f. f:(\\<Pi>X \\<in> Pow(A)-{0}. X)";
-by (fast_tac (claset() addSEs [well_ord_subset RS ex_choice_fun]) 1);
-qed "ex_choice_fun_Pow";
--- a/src/ZF/AC/WO_AC.thy	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-(*Dummy theory to document dependencies *)
-
-WO_AC = Order
--- a/src/ZF/AC/recfunAC16.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,65 +0,0 @@
-(*  Title:      ZF/AC/recfunAC16.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Properties of the recursive definition used in the proof of WO2 ==> AC16
-*)
-
-(* ********************************************************************** *)
-(* Basic properties of recfunAC16                                         *)
-(* ********************************************************************** *)
-
-Goalw [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0";
-by (rtac transrec2_0 1);
-qed "recfunAC16_0";
-
-Goalw [recfunAC16_def]
-     "recfunAC16(f,fa,succ(i),a) =  \
-\     (if (\\<exists>y \\<in> recfunAC16(f,fa,i,a). fa ` i \\<subseteq> y) then recfunAC16(f,fa,i,a) \
-\      else recfunAC16(f,fa,i,a) Un \
-\           {f ` (LEAST j. fa ` i \\<subseteq> f ` j &  \
-\            (\\<forall>b<a. (fa`b \\<subseteq> f`j  \
-\             --> (\\<forall>t \\<in> recfunAC16(f,fa,i,a). ~ fa`b \\<subseteq> t))))})";
-by (rtac transrec2_succ 1);
-qed "recfunAC16_succ";
-
-Goalw [recfunAC16_def] "Limit(i)  \
-\       ==> recfunAC16(f,fa,i,a) = (\\<Union>j<i. recfunAC16(f,fa,j,a))";
-by (etac transrec2_Limit 1);
-qed "recfunAC16_Limit";
-
-(* ********************************************************************** *)
-(* Monotonicity of transrec2                                              *)
-(* ********************************************************************** *)
-
-val [prem1, prem2] = goal thy 
-    "[| !!g r. r \\<subseteq> B(g,r);  Ord(i) |]  \
-\       ==> j<i --> transrec2(j, 0, B) \\<subseteq> transrec2(i, 0, B)";
-by (resolve_tac [prem2 RS trans_induct] 1);
-by (rtac Ord_cases 1 THEN (REPEAT (assume_tac 1)));
-by (Fast_tac 1);
-by (Asm_simp_tac 1);
-by (fast_tac (FOL_cs addSIs [succI1, prem1]
-        addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1);
-by (fast_tac (claset() addIs [OUN_I, ltI]
-        addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl,
-                transrec2_Limit RS ssubst]) 1);
-qed "transrec2_mono_lemma";
-
-val [prem1, prem2] = goal thy "[| !!g r. r \\<subseteq> B(g,r); j le i |]  \
-\       ==> transrec2(j, 0, B) \\<subseteq> transrec2(i, 0, B)";
-by (resolve_tac [prem2 RS leE] 1);
-by (resolve_tac [transrec2_mono_lemma RS impE] 1);
-by (TRYALL (fast_tac (claset() addSIs [prem1, prem2, lt_Ord2])));
-qed "transrec2_mono";
-
-(* ********************************************************************** *)
-(* Monotonicity of recfunAC16                                             *)
-(* ********************************************************************** *)
-
-Goalw [recfunAC16_def]
-        "!!i. i le j ==> recfunAC16(f, g, i, a) \\<subseteq> recfunAC16(f, g, j, a)";
-by (rtac transrec2_mono 1);
-by (REPEAT (fast_tac (claset() addIs [split_if RS iffD2]) 1));
-qed "recfunAC16_mono";
-
--- a/src/ZF/AC/recfunAC16.thy	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-(*  Title:      ZF/AC/recfunAC16.thy
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-A recursive definition used in the proof of WO2 ==> AC16
-*)
-
-recfunAC16 = Cardinal + Epsilon +
-
-constdefs
-  recfunAC16 :: [i, i, i, i] => i
-
-    "recfunAC16(f,fa,i,a) == 
-         transrec2(i, 0, 
-              %g r. if(\\<exists>y \\<in> r. fa`g \\<subseteq> y, r, 
-                       r Un {f`(LEAST i. fa`g \\<subseteq> f`i & 
-                       (\\<forall>b<a. (fa`b \\<subseteq> f`i --> (\\<forall>t \\<in> r. ~ fa`b \\<subseteq> t))))}))"
-
-end
--- a/src/ZF/AC/rel_is_fun.ML	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,42 +0,0 @@
-(*  Title:      ZF/AC/rel_is_fun.ML
-    ID:         $Id$
-    Author:     Krzysztof Grabczewski
-
-Lemmas needed to state when a finite relation is a function.
-
-The criteria are cardinalities of the relation and its domain.
-Used in WO6WO1.ML
-*)
-
-(*Using AC we could trivially prove, for all u, domain(u) lepoll u*)
-goalw Cardinal.thy [lepoll_def]
-     "!!m. [| m \\<in> nat; u lepoll m |] ==> domain(u) lepoll m";
-by (etac exE 1);
-by (res_inst_tac [("x",
-        "\\<lambda>x \\<in> domain(u). LEAST i. \\<exists>y. <x,y> \\<in> u & f`<x,y> = i")] exI 1);
-by (res_inst_tac [("d","%y. fst(converse(f)`y)")] lam_injective 1);
-by (fast_tac (claset() addIs [LeastI2, nat_into_Ord RS Ord_in_Ord,
-                        inj_is_fun RS apply_type]) 1);
-by (etac domainE 1);
-by (forward_tac [inj_is_fun RS apply_type] 1 THEN (atac 1));
-by (rtac LeastI2 1);
-by (auto_tac (claset() addSEs [nat_into_Ord RS Ord_in_Ord],
-	      simpset()));
-qed "lepoll_m_imp_domain_lepoll_m";
-
-goalw Cardinal.thy [function_def]
-    "!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m \\<in> nat |] ==> \
-\         function(r)";
-by Safe_tac;
-by (resolve_tac [excluded_middle RS disjE] 1 THEN (atac 2));
-by (fast_tac (claset() addSEs [lepoll_trans RS succ_lepoll_natE, 
-                        Diff_sing_lepoll RSN (2, lepoll_m_imp_domain_lepoll_m)]
-                addEs [not_sym RSN (2, domain_Diff_eq) RS subst]) 1);
-qed "rel_domain_ex1";
-
-goal Cardinal.thy
-    "!!r. [| succ(m) lepoll domain(r);  r lepoll succ(m);  m \\<in> nat;  \
-\            r \\<subseteq> A*B; A=domain(r) |] ==> r \\<in> A->B";
-by (hyp_subst_tac 1);
-by (asm_simp_tac (simpset() addsimps [Pi_iff, rel_domain_ex1]) 1);
-qed "rel_is_fun";
--- a/src/ZF/AC/rel_is_fun.thy	Wed Jan 16 15:04:37 2002 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-rel_is_fun = Cardinal
--- a/src/ZF/CardinalArith.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/CardinalArith.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -102,4 +102,60 @@
      "(!!x. x:A ==> Card(K(x))) ==> Card(UN x<A. K(x))"
 by (simp add: OUnion_def Card_0) 
 
+lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat"
+apply (unfold lesspoll_def)
+apply (rule conjI)
+apply (erule OrdmemD [THEN subset_imp_lepoll], rule Ord_nat)
+apply (rule notI)
+apply (erule eqpollE)
+apply (rule succ_lepoll_natE)
+apply (blast intro: nat_succI [THEN OrdmemD, THEN subset_imp_lepoll] 
+                    lepoll_trans, assumption); 
+done
+
+lemma in_Card_imp_lesspoll: "[| Card(K); b \<in> K |] ==> b \<prec> K"
+apply (unfold lesspoll_def)
+apply (simp add: Card_iff_initial)
+apply (fast intro!: le_imp_lepoll ltI leI)
+done
+
+lemma succ_lepoll_imp_not_empty: "succ(x) \<lesssim> y ==> y \<noteq> 0"
+by (fast dest!: lepoll_0_is_0)
+
+lemma eqpoll_succ_imp_not_empty: "x \<approx> succ(n) ==> x \<noteq> 0"
+by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0])
+
+lemma Finite_Fin_lemma [rule_format]:
+     "n \<in> nat ==> \<forall>A. (A\<approx>n & A \<subseteq> X) --> A \<in> Fin(X)"
+apply (induct_tac "n")
+apply (rule allI)
+apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
+apply (rule allI)
+apply (rule impI)
+apply (erule conjE)
+apply (rule eqpoll_succ_imp_not_empty [THEN not_emptyE], (assumption))
+apply (frule Diff_sing_eqpoll, (assumption))
+apply (erule allE)
+apply (erule impE, fast)
+apply (drule subsetD, (assumption))
+apply (drule Fin.consI, (assumption))
+apply (simp add: cons_Diff)
+done
+
+lemma Finite_Fin: "[| Finite(A); A \<subseteq> X |] ==> A \<in> Fin(X)"
+by (unfold Finite_def, blast intro: Finite_Fin_lemma) 
+
+lemma lesspoll_lemma: 
+        "[| ~ A \<prec> B; C \<prec> B |] ==> A - C \<noteq> 0"
+apply (unfold lesspoll_def)
+apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll]
+            intro!: eqpollI elim: notE 
+            elim!: eqpollE lepoll_trans)
+done
+
+lemma eqpoll_imp_Finite_iff: "A \<approx> B ==> Finite(A) <-> Finite(B)"
+apply (unfold Finite_def) 
+apply (blast intro: eqpoll_trans eqpoll_sym) 
+done
+
 end
--- a/src/ZF/IsaMakefile	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/IsaMakefile	Wed Jan 16 17:52:06 2002 +0100
@@ -58,16 +58,11 @@
 ZF-AC: ZF $(LOG)/ZF-AC.gz
 
 $(LOG)/ZF-AC.gz: $(OUT)/ZF \
-  AC/AC15_WO6.ML AC/AC15_WO6.thy AC/AC16_WO4.ML AC/AC16_WO4.thy \
-  AC/AC16_lemmas.ML AC/AC16_lemmas.thy AC/AC17_AC1.ML AC/AC17_AC1.thy \
-  AC/AC18_AC19.ML AC/AC18_AC19.thy AC/AC1_WO2.ML \
-  AC/AC1_WO2.thy AC/AC7_AC9.ML AC/AC_Equiv.ML \
-  AC/AC_Equiv.thy AC/Cardinal_aux.ML AC/Cardinal_aux.thy AC/DC.ML \
-  AC/DC.thy AC/DC_lemmas.ML AC/DC_lemmas.thy AC/HH.ML AC/HH.thy \
-  AC/Hartog.ML AC/Hartog.thy AC/ROOT.ML AC/WO1_AC.ML AC/WO1_AC.thy \
-  AC/WO1_WO7.ML AC/WO1_WO7.thy AC/WO2_AC16.ML \
-  AC/WO2_AC16.thy AC/WO6_WO1.ML AC/WO6_WO1.thy AC/WO_AC.ML AC/WO_AC.thy \
-  AC/recfunAC16.ML AC/recfunAC16.thy AC/rel_is_fun.ML AC/rel_is_fun.thy
+  AC/ROOT.ML  AC/AC15_WO6.thy AC/AC16_WO4.thy \
+  AC/AC16_lemmas.thy AC/AC17_AC1.thy AC/AC18_AC19.thy AC/AC7_AC9.thy \
+  AC/AC_Equiv.thy AC/Cardinal_aux.thy \
+  AC/DC.thy AC/HH.thy AC/Hartog.thy AC/WO1_AC.thy AC/WO1_WO7.thy \
+  AC/WO2_AC16.thy AC/WO6_WO1.thy 
 	@$(ISATOOL) usedir $(OUT)/ZF AC
 
 
--- a/src/ZF/Main.thy	Wed Jan 16 15:04:37 2002 +0100
+++ b/src/ZF/Main.thy	Wed Jan 16 17:52:06 2002 +0100
@@ -57,4 +57,15 @@
           (Limit(K) --> f(K) = (UN j<K. f(j)))"
 by (simp add: transrec2_Limit)
 
+(* belongs to theory CardinalArith *)
+
+lemma InfCard_square_eqpoll: "InfCard(K) \<Longrightarrow> K \<times> K \<approx> K"
+apply (rule well_ord_InfCard_square_eq);  
+ apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]); 
+apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]); 
+done
+
+lemma Inf_Card_is_InfCard: "[| ~Finite(i); Card(i) |] ==> InfCard(i)"
+by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord])
+
 end