src/ZF/AC/WO1_AC.ML
author paulson
Fri, 03 Jan 1997 15:01:55 +0100
changeset 2469 b50b8c0eec01
parent 1461 6bcb44e4d6e5
child 2496 40efb87985b5
permissions -rw-r--r--
Implicit simpsets and clasets for FOL and ZF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     1
(*  Title:      ZF/AC/WO1_AC.ML
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
     3
    Author:     Krzysztof Grabczewski
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     4
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     5
The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     6
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     7
The latter proof is referred to as clear by the Rubins.
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     8
However it seems to be quite complicated.
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     9
The formal proof presented below is a mechanisation of the proof 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    10
by Lawrence C. Paulson which is the following:
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    11
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    12
Assume WO1.  Let s be a set of infinite sets.
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    13
 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    14
Suppose x:s.  Then x is equipollent to |x| (by WO1), an infinite cardinal; 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    15
call it K.  Since K = K |+| K = |K+K| (by InfCard_cdouble_eq) there is an 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    16
isomorphism h: bij(K+K, x).  (Here + means disjoint sum.)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    17
 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    18
So there is a partition of x into 2-element sets, namely
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    19
 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    20
        {{h(Inl(i)), h(Inr(i))} . i:K}
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    21
 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    22
So for all x:s the desired partition exists.  By AC1 (which follows from WO1) 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    23
there exists a function f that chooses a partition for each x:s.  Therefore we 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    24
have AC10(2).
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    25
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    26
*)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    27
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    28
open WO1_AC;
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    29
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    30
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    31
(* WO1 ==> AC1                                                            *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    32
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    33
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    34
goalw thy [AC1_def, WO1_def] "!!Z. WO1 ==> AC1";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    35
by (fast_tac (!claset addSEs [ex_choice_fun]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    36
qed "WO1_AC1";
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    37
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    38
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    39
(* WO1 ==> AC10(n) (n >= 1)                                               *)
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    40
(* ********************************************************************** *)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    41
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    42
goalw thy [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    43
\               ==> EX f. ALL B:A. P(f`B,B)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    44
by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    45
by (etac exE 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    46
by (dtac ex_choice_fun 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    47
by (fast_tac (!claset addEs [RepFunE, sym RS equals0D]) 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    48
by (etac exE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    49
by (res_inst_tac [("x","lam x:A. f`{C:D(x). P(C,x)}")] exI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    50
by (Asm_full_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    51
by (fast_tac (!claset addSDs [RepFunI RSN (2, apply_type)]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    52
                addSEs [CollectD2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    53
val lemma1 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    54
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    55
goalw thy [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll  B";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    56
by (rtac eqpoll_trans 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    57
by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 2);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    58
by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    59
by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    60
by (resolve_tac [cadd_def RS def_imp_eq RS subst] 1);
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    61
by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    62
                        InfCard_cdouble_eq RS ssubst] 1);
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    63
by (rtac eqpoll_refl 2);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    64
by (rtac notI 1);
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    65
by (etac notE 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    66
by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    67
        THEN (assume_tac 2));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    68
by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    69
val lemma2_1 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    70
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    71
goal thy "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    72
by (fast_tac (!claset addSIs [InlI, InrI]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    73
                addSEs [RepFunE, bij_is_fun RS apply_type]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    74
val lemma2_2 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    75
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    76
goal thy "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    77
by (rtac inj_equality 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    78
by (TRYALL (fast_tac (!claset addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    79
val lemma = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    80
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    81
goalw thy AC_aux_defs
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    82
        "!!f. f : bij(D+D, B) ==>  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    83
\               pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    84
by (fast_tac (!claset addSEs [RepFunE, not_emptyE] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    85
        addDs [bij_is_inj RS lemma, Inl_iff RS iffD1,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    86
                Inr_iff RS iffD1, Inl_Inr_iff RS iffD1 RS FalseE,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    87
                Inr_Inl_iff RS iffD1 RS FalseE]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    88
        addSIs [InlI, InrI]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    89
val lemma2_3 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    90
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    91
val [major, minor] = goalw thy AC_aux_defs 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    92
        "[| f : bij(D+D, B); 1 le n |] ==>  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    93
\       sets_of_size_between({{f`Inl(i), f`Inr(i)}. i:D}, 2, succ(n))";
1208
bc3093616ba4 Ran expandshort and corrected spelling of Grabczewski
paulson
parents: 1196
diff changeset
    94
by (rewtac succ_def);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    95
by (fast_tac (!claset addSIs [cons_lepoll_cong, minor, lepoll_refl, InlI, InrI] 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    96
        addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    97
                le_imp_subset RS subset_imp_lepoll]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    98
        addDs [major RS bij_is_inj RS lemma, Inl_Inr_iff RS iffD1 RS FalseE]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
    99
        addSEs [RepFunE, not_emptyE, mem_irrefl]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   100
val lemma2_4 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   101
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   102
goalw thy [bij_def, surj_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   103
        "!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   104
by (fast_tac (!claset addSEs [inj_is_fun RS apply_type, CollectE, sumE]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   105
        addSIs [InlI, InrI, equalityI]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   106
val lemma2_5 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   107
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   108
goal thy "!!A. [| WO1; ~Finite(B); 1 le n  |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   109
\       ==> EX C:Pow(Pow(B)). pairwise_disjoint(C) &  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   110
\               sets_of_size_between(C, 2, succ(n)) &  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   111
\               Union(C)=B";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   112
by (eresolve_tac [lemma2_1 RS (eqpoll_def RS def_imp_iff RS iffD1 RS exE)] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   113
        THEN (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   114
by (fast_tac (FOL_cs addSIs [bexI]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1208
diff changeset
   115
                addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   116
val lemma2 = result();
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   117
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   118
goalw thy AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   119
by (fast_tac (!claset addSIs [lemma1] addSEs [lemma2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
   120
qed "WO1_AC10";