| 12776 |      1 | (*  Title:      ZF/AC/WO1_AC.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Krzysztof Grabczewski
 | 
|  |      4 | 
 | 
|  |      5 | The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
 | 
|  |      6 | 
 | 
|  |      7 | The latter proof is referred to as clear by the Rubins.
 | 
|  |      8 | However it seems to be quite complicated.
 | 
|  |      9 | The formal proof presented below is a mechanisation of the proof 
 | 
|  |     10 | by Lawrence C. Paulson which is the following:
 | 
|  |     11 | 
 | 
|  |     12 | Assume WO1.  Let s be a set of infinite sets.
 | 
|  |     13 |  
 | 
|  |     14 | Suppose x \<in> s.  Then x is equipollent to |x| (by WO1), an infinite cardinal
 | 
|  |     15 | call it K.  Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an 
 | 
|  |     16 | isomorphism h \<in> bij(K+K, x).  (Here + means disjoint sum.)
 | 
|  |     17 |  
 | 
|  |     18 | So there is a partition of x into 2-element sets, namely
 | 
|  |     19 |  
 | 
|  |     20 |         {{h(Inl(i)), h(Inr(i))} . i \<in> K}
 | 
|  |     21 |  
 | 
|  |     22 | So for all x \<in> s the desired partition exists.  By AC1 (which follows from WO1) 
 | 
|  |     23 | there exists a function f that chooses a partition for each x \<in> s.  Therefore we 
 | 
|  |     24 | have AC10(2).
 | 
|  |     25 | 
 | 
|  |     26 | *)
 | 
|  |     27 | 
 | 
| 27678 |     28 | theory WO1_AC
 | 
|  |     29 | imports AC_Equiv
 | 
|  |     30 | begin
 | 
| 12776 |     31 | 
 | 
|  |     32 | (* ********************************************************************** *)
 | 
|  |     33 | (* WO1 ==> AC1                                                            *)
 | 
|  |     34 | (* ********************************************************************** *)
 | 
|  |     35 | 
 | 
|  |     36 | theorem WO1_AC1: "WO1 ==> AC1"
 | 
|  |     37 | by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun)
 | 
|  |     38 | 
 | 
|  |     39 | (* ********************************************************************** *)
 | 
|  |     40 | (* WO1 ==> AC10(n) (n >= 1)                                               *)
 | 
|  |     41 | (* ********************************************************************** *)
 | 
|  |     42 | 
 | 
|  |     43 | 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)"
 | 
|  |     44 | apply (unfold WO1_def)
 | 
|  |     45 | apply (erule_tac x = "Union ({{C \<in> D (B) . P (C,B) }. B \<in> A}) " in allE)
 | 
|  |     46 | apply (erule exE, drule ex_choice_fun, fast)
 | 
|  |     47 | apply (erule exE)
 | 
|  |     48 | apply (rule_tac x = "\<lambda>x \<in> A. f`{C \<in> D (x) . P (C,x) }" in exI)
 | 
|  |     49 | apply (simp, blast dest!: apply_type [OF _ RepFunI])
 | 
|  |     50 | done
 | 
| 1196 |     51 | 
 | 
| 12776 |     52 | lemma lemma2_1: "[| ~Finite(B); WO1 |] ==> |B| + |B| \<approx>  B"
 | 
|  |     53 | apply (unfold WO1_def)
 | 
|  |     54 | apply (rule eqpoll_trans)
 | 
|  |     55 | prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll)
 | 
|  |     56 | apply (rule eqpoll_sym [THEN eqpoll_trans])
 | 
|  |     57 | apply (fast elim!: well_ord_cardinal_eqpoll)
 | 
|  |     58 | apply (drule spec [of _ B]) 
 | 
|  |     59 | apply (clarify dest!: eqpoll_imp_Finite_iff [OF well_ord_cardinal_eqpoll]) 
 | 
|  |     60 | apply (simp add: cadd_def [symmetric] 
 | 
|  |     61 |             eqpoll_refl InfCard_cdouble_eq Card_cardinal Inf_Card_is_InfCard) 
 | 
|  |     62 | done
 | 
|  |     63 | 
 | 
|  |     64 | lemma lemma2_2:
 | 
|  |     65 |      "f \<in> bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \<in> D} \<in> Pow(Pow(B))"
 | 
|  |     66 | by (fast elim!: bij_is_fun [THEN apply_type])
 | 
|  |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | lemma lemma2_3: 
 | 
|  |     70 |         "f \<in> bij(D+D, B) ==> pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \<in> D})"
 | 
|  |     71 | apply (unfold pairwise_disjoint_def)
 | 
|  |     72 | apply (blast dest: bij_is_inj [THEN inj_apply_equality])
 | 
|  |     73 | done
 | 
|  |     74 | 
 | 
|  |     75 | lemma lemma2_4:
 | 
|  |     76 |      "[| f \<in> bij(D+D, B); 1\<le>n |] 
 | 
|  |     77 |       ==> sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \<in> D}, 2, succ(n))"
 | 
|  |     78 | apply (simp (no_asm_simp) add: sets_of_size_between_def succ_def)
 | 
|  |     79 | apply (blast intro!: cons_lepoll_cong 
 | 
|  |     80 |             intro: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll]  
 | 
|  |     81 |                    le_imp_subset [THEN subset_imp_lepoll]  lepoll_trans 
 | 
|  |     82 |             dest: bij_is_inj [THEN inj_apply_equality] elim!: mem_irrefl)
 | 
|  |     83 | done
 | 
|  |     84 | 
 | 
|  |     85 | lemma lemma2_5: 
 | 
|  |     86 |      "f \<in> bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i \<in> D})=B"
 | 
|  |     87 | apply (unfold bij_def surj_def)
 | 
|  |     88 | apply (fast elim!: inj_is_fun [THEN apply_type])
 | 
|  |     89 | done
 | 
|  |     90 | 
 | 
|  |     91 | lemma lemma2:
 | 
|  |     92 |      "[| WO1; ~Finite(B); 1\<le>n  |]   
 | 
|  |     93 |       ==> \<exists>C \<in> Pow(Pow(B)). pairwise_disjoint(C) &   
 | 
|  |     94 |                 sets_of_size_between(C, 2, succ(n)) &   
 | 
|  |     95 |                 Union(C)=B"
 | 
|  |     96 | apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], 
 | 
|  |     97 |        assumption)
 | 
|  |     98 | apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5)
 | 
|  |     99 | done
 | 
|  |    100 | 
 | 
|  |    101 | theorem WO1_AC10: "[| WO1; 1\<le>n |] ==> AC10(n)"
 | 
|  |    102 | apply (unfold AC10_def)
 | 
|  |    103 | apply (fast intro!: lemma1 elim!: lemma2)
 | 
|  |    104 | done
 | 
|  |    105 | 
 | 
|  |    106 | end
 | 
|  |    107 | 
 |