src/ZF/AC/WO1_AC.ML
changeset 1196 d43c1f7a53fe
child 1208 bc3093616ba4
equal deleted inserted replaced
1195:686e3eb613b9 1196:d43c1f7a53fe
       
     1 (*  Title: 	ZF/AC/WO1_AC.ML
       
     2     ID:         $Id$
       
     3     Author: 	Krzysztof Gr`abczewski
       
     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: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: 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:K}
       
    21  
       
    22 So for all x:s the desired partition exists.  By AC1 (which follows from WO1) 
       
    23 there exists a function f that chooses a partition for each x:s.  Therefore we 
       
    24 have AC10(2).
       
    25 
       
    26 *)
       
    27 
       
    28 open WO1_AC;
       
    29 
       
    30 (* ********************************************************************** *)
       
    31 (* WO1 ==> AC1								  *)
       
    32 (* ********************************************************************** *)
       
    33 
       
    34 goalw thy [AC1_def, WO1_def] "!!Z. WO1 ==> AC1";
       
    35 by (fast_tac (AC_cs addSEs [ex_choice_fun]) 1);
       
    36 qed "WO1_AC1";
       
    37 
       
    38 (* ********************************************************************** *)
       
    39 (* WO1 ==> AC10(n) (n >= 1)						  *)
       
    40 (* ********************************************************************** *)
       
    41 
       
    42 goalw thy [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
       
    43 \		==> EX f. ALL B:A. P(f`B,B)";
       
    44 by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1);
       
    45 by (eresolve_tac [exE] 1);
       
    46 by (dresolve_tac [ex_choice_fun] 1);
       
    47 by (fast_tac (AC_cs addEs [RepFunE, sym RS equals0D]) 1);
       
    48 by (eresolve_tac [exE] 1);
       
    49 by (res_inst_tac [("x","lam x:A. f`{C:D(x). P(C,x)}")] exI 1);
       
    50 by (asm_full_simp_tac AC_ss 1);
       
    51 by (fast_tac (AC_cs addSDs [RepFunI RSN (2, apply_type)]
       
    52 		addSEs [CollectD2]) 1);
       
    53 val lemma1 = result();
       
    54 
       
    55 goalw thy [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
       
    56 by (resolve_tac [eqpoll_trans] 1);
       
    57 by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 2);
       
    58 by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
       
    59 by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1);
       
    60 by (resolve_tac [cadd_def RS def_imp_eq RS subst] 1);
       
    61 by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS 
       
    62 			InfCard_cdouble_eq RS ssubst] 1);
       
    63 by (resolve_tac [eqpoll_refl] 2);
       
    64 by (resolve_tac [notI] 1);
       
    65 by (eresolve_tac [notE] 1);
       
    66 by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1
       
    67 	THEN (assume_tac 2));
       
    68 by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1);
       
    69 val lemma2_1 = result();
       
    70 
       
    71 goal thy "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
       
    72 by (fast_tac (AC_cs addSIs [InlI, InrI]
       
    73 		addSEs [RepFunE, bij_is_fun RS apply_type]) 1);
       
    74 val lemma2_2 = result();
       
    75 
       
    76 goal thy "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
       
    77 by (resolve_tac [inj_equality] 1);
       
    78 by (TRYALL (fast_tac (AC_cs addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
       
    79 val lemma = result();
       
    80 
       
    81 goalw thy AC_aux_defs
       
    82 	"!!f. f : bij(D+D, B) ==>  \
       
    83 \		pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})";
       
    84 by (fast_tac (AC_cs addSEs [RepFunE, not_emptyE] 
       
    85 	addDs [bij_is_inj RS lemma, Inl_iff RS iffD1,
       
    86 		Inr_iff RS iffD1, Inl_Inr_iff RS iffD1 RS FalseE,
       
    87 		Inr_Inl_iff RS iffD1 RS FalseE]
       
    88 	addSIs [InlI, InrI]) 1);
       
    89 val lemma2_3 = result();
       
    90 
       
    91 val [major, minor] = goalw thy AC_aux_defs 
       
    92 	"[| f : bij(D+D, B); 1 le n |] ==>  \
       
    93 \	sets_of_size_between({{f`Inl(i), f`Inr(i)}. i:D}, 2, succ(n))";
       
    94 by (rewrite_goals_tac [succ_def]);
       
    95 by (fast_tac (AC_cs addSIs [cons_lepoll_cong, minor, lepoll_refl, InlI, InrI] 
       
    96 	addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
       
    97 		le_imp_subset RS subset_imp_lepoll]
       
    98 	addDs [major RS bij_is_inj RS lemma, Inl_Inr_iff RS iffD1 RS FalseE]
       
    99 	addSEs [RepFunE, not_emptyE, mem_irrefl]) 1);
       
   100 val lemma2_4 = result();
       
   101 
       
   102 goalw thy [bij_def, surj_def]
       
   103 	"!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B";
       
   104 by (fast_tac (AC_cs addSEs [inj_is_fun RS apply_type, CollectE, sumE]
       
   105 	addSIs [InlI, InrI, equalityI]) 1);
       
   106 val lemma2_5 = result();
       
   107 
       
   108 goal thy "!!A. [| WO1; ~Finite(B); 1 le n  |]  \
       
   109 \	==> EX C:Pow(Pow(B)). pairwise_disjoint(C) &  \
       
   110 \		sets_of_size_between(C, 2, succ(n)) &  \
       
   111 \		Union(C)=B";
       
   112 by (eresolve_tac [lemma2_1 RS (eqpoll_def RS def_imp_iff RS iffD1 RS exE)] 1
       
   113 	THEN (assume_tac 1));
       
   114 by (fast_tac (FOL_cs addSIs [bexI]
       
   115 		addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1);
       
   116 val lemma2 = result();
       
   117 
       
   118 goalw thy AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
       
   119 by (fast_tac (AC_cs addSIs [lemma1] addSEs [lemma2]) 1);
       
   120 qed "WO1_AC10";