src/ZF/AC/AC7_AC9.ML
author clasohm
Tue, 30 Jan 1996 13:42:57 +0100
changeset 1461 6bcb44e4d6e5
parent 1207 3f460842e919
child 2469 b50b8c0eec01
permissions -rw-r--r--
expanded tabs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
     1
(*  Title:      ZF/AC/AC7-AC9.ML
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
     3
    Author:     Krzysztof Grabczewski
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     4
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     5
The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     6
instances of AC.
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     7
*)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     8
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
     9
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    10
(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1                  *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    11
(*  - Sigma_fun_space_not0                                                *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    12
(*  - all_eqpoll_imp_pair_eqpoll                                          *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    13
(*  - Sigma_fun_space_eqpoll                                              *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    14
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    15
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    16
goal ZF.thy "!!A. [| C~:A; B:A |] ==> B~=C";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    17
by (fast_tac ZF_cs 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    18
val mem_not_eq_not_mem = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    19
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    20
goal thy "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    21
by (fast_tac (ZF_cs addSDs [Sigma_empty_iff RS iffD1]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    22
                addDs [fun_space_emptyD, mem_not_eq_not_mem]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    23
                addEs [equals0D]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    24
                addSIs [equals0I,UnionI]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    25
val Sigma_fun_space_not0 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    26
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    27
goal thy "!!A C. (ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    28
by (REPEAT (rtac ballI 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    29
by (resolve_tac [bspec RS (bspec RSN (2, eqpoll_sym RSN (2, eqpoll_trans)))] 1
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    30
        THEN REPEAT (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    31
val all_eqpoll_imp_pair_eqpoll = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    32
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    33
goal thy "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A   \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    34
\       |] ==> P(b)=R(b)";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    35
by (dtac bspec 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    36
by (asm_full_simp_tac ZF_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    37
val if_eqE1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    38
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    39
goal thy "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a))  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    40
\       ==> ALL a:A. a~=b --> Q(a)=S(a)";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    41
by (rtac ballI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    42
by (rtac impI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    43
by (dtac bspec 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    44
by (asm_full_simp_tac ZF_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    45
val if_eqE2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    46
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    47
goal thy "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    48
by (fast_tac (ZF_cs addDs [subsetD]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    49
                addSIs [lamI]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    50
                addEs [equalityE, lamE]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    51
val lam_eqE = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    52
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    53
goalw thy [inj_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    54
        "!!A. C:A ==> (lam g:(nat->Union(A))*C.  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    55
\               (lam n:nat. if(n=0, snd(g), fst(g)`(n #- 1))))  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    56
\               : inj((nat->Union(A))*C, (nat->Union(A)) ) ";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    57
by (rtac CollectI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    58
by (fast_tac (ZF_cs addSIs [lam_type,RepFunI,if_type,snd_type,apply_type,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    59
                                fst_type,diff_type,nat_succI,nat_0I]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    60
by (REPEAT (resolve_tac [ballI, impI] 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    61
by (asm_full_simp_tac ZF_ss 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    62
by (REPEAT (etac SigmaE 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    63
by (REPEAT (hyp_subst_tac 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    64
by (asm_full_simp_tac ZF_ss 1);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    65
by (rtac conjI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    66
by (dresolve_tac [nat_0I RSN (2, lam_eqE)] 2);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    67
by (asm_full_simp_tac AC_ss 2);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    68
by (rtac fun_extension 1 THEN  REPEAT (assume_tac 1));
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    69
by (dresolve_tac [nat_succI RSN (2, lam_eqE)] 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    70
by (asm_full_simp_tac (AC_ss addsimps [succ_not_0 RS if_not_P]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    71
by (fast_tac (AC_cs addSEs [diff_succ_succ RS (diff_0 RSN (2, trans)) RS subst]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    72
                addSIs [nat_0I]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    73
val lemma = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    74
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    75
goal thy "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    76
by (rtac eqpollI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    77
by (fast_tac (ZF_cs addSEs [prod_lepoll_self, not_sym RS not_emptyE,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    78
                subst_elem] addEs [swap]) 2);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
    79
by (rewtac lepoll_def);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    80
by (fast_tac (ZF_cs addSIs [lemma]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    81
val Sigma_fun_space_eqpoll = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    82
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    83
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    84
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    85
(* AC6 ==> AC7                                                            *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    86
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    87
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    88
goalw thy AC_defs "!!Z. AC6 ==> AC7";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    89
by (fast_tac ZF_cs 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
    90
qed "AC6_AC7";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    91
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    92
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    93
(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8                          *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    94
(* The case of the empty family of sets added in order to complete        *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
    95
(* the proof.                                                             *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    96
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    97
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    98
goal thy "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
    99
by (fast_tac (ZF_cs addSIs [lam_type, snd_type, apply_type]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   100
val lemma1_1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   101
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   102
goal thy "!!A. y: (PROD B:{Y*C. C:A}. B)  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   103
\               ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   104
by (fast_tac (ZF_cs addSIs [lam_type, apply_type]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   105
val lemma1_2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   106
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   107
goal thy "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   108
\               ==> (PROD B:A. B) ~= 0";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   109
by (fast_tac (ZF_cs addSIs [equals0I,lemma1_1, lemma1_2]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   110
                addSEs [equals0D]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   111
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   112
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   113
goal thy "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   114
by (fast_tac (ZF_cs addEs [RepFunE,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   115
                Sigma_fun_space_not0 RS not_sym RS notE]) 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   116
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   117
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   118
goalw thy AC_defs "!!Z. AC7 ==> AC6";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   119
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   120
by (rtac impI 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   121
by (excluded_middle_tac "A=0" 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   122
by (fast_tac (ZF_cs addSIs [not_emptyI, empty_fun]) 2);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   123
by (rtac lemma1 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   124
by (etac allE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   125
by (etac impE 1 THEN (assume_tac 2));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   126
by (fast_tac (AC_cs addSEs [RepFunE]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   127
        addSIs [lemma2, all_eqpoll_imp_pair_eqpoll,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   128
                Sigma_fun_space_eqpoll]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   129
qed "AC7_AC6";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   130
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   131
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   132
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   133
(* AC1 ==> AC8                                                            *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   134
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   135
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   136
goalw thy [eqpoll_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   137
        "!!A. ALL B:A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   138
\       ==> 0 ~: { bij(fst(B),snd(B)). B:A }";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   139
by (rtac notI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   140
by (etac RepFunE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   141
by (dtac bspec 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   142
by (REPEAT (eresolve_tac [exE,conjE] 1));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   143
by (hyp_subst_tac 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   144
by (asm_full_simp_tac AC_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   145
by (fast_tac (AC_cs addSEs [sym RS equals0D]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   146
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   147
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   148
goal thy "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   149
\               ==> (lam x:A. f`p(x))`D : p(D)";
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   150
by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   151
by (fast_tac (AC_cs addSEs [apply_type]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   152
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   153
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   154
goalw thy AC_defs "!!Z. AC1 ==> AC8";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   155
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   156
by (etac allE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   157
by (rtac impI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   158
by (etac impE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   159
by (etac lemma1 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   160
by (fast_tac (AC_cs addSEs [lemma2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   161
qed "AC1_AC8";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   162
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   163
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   164
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   165
(* AC8 ==> AC9                                                            *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   166
(*  - this proof replaces the following two from Rubin & Rubin:           *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   167
(*    AC8 ==> AC1 and AC1 ==> AC9                                         *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   168
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   169
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   170
goal thy "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==>  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   171
\               ALL B:A*A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   172
by (fast_tac ZF_cs 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   173
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   174
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   175
goal thy "!!f. f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   176
by (asm_full_simp_tac AC_ss 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   177
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   178
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   179
goalw thy AC_defs "!!Z. AC8 ==> AC9";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   180
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   181
by (rtac impI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   182
by (etac allE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   183
by (etac impE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   184
by (etac lemma1 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   185
by (fast_tac (AC_cs addSEs [lemma2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   186
qed "AC8_AC9";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   187
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   188
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   189
(* ********************************************************************** *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   190
(* AC9 ==> AC1                                                            *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   191
(* The idea of this proof comes from "Equivalents of the Axiom of Choice" *)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   192
(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to        *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   193
(* (x * y) Un {0} when y is a set of total functions acting from nat to   *)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   194
(* Union(A) -- therefore we have used the set (y * nat) instead of y.     *)
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   195
(* ********************************************************************** *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   196
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   197
(* Rules nedded to prove lemma1 *)
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   198
val snd_lepoll_SigmaI = prod_lepoll_self RS 
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   199
        ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans));
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   200
val lemma1_cs = FOL_cs addSEs [UnE, RepFunE]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   201
                addSIs [all_eqpoll_imp_pair_eqpoll, ballI,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   202
                        nat_cons_eqpoll RS eqpoll_trans]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   203
                addEs [Sigma_fun_space_not0 RS not_emptyE]
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   204
                addIs [snd_lepoll_SigmaI, eqpoll_refl RSN 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   205
                        (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll];
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   206
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   207
goal thy "!!A. [| 0~:A; A~=0 |]  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   208
\       ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   209
\               Un {cons(0,((nat->Union(A))*B)*nat). B:A}).  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   210
\       ALL B2: ({((nat->Union(A))*B)*nat. B:A}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   211
\               Un {cons(0,((nat->Union(A))*B)*nat). B:A}).  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   212
\       B1 eqpoll B2";
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   213
by (fast_tac lemma1_cs 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   214
val lemma1 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   215
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   216
goal thy "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}.  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   217
\       ALL B2:{(F*B)*N. B:A}  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   218
\       Un {cons(0,(F*B)*N). B:A}. f`<B1,B2> : bij(B1, B2)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   219
\       ==> (lam B:A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) :  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1207
diff changeset
   220
\               (PROD X:A. X)";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   221
by (rtac lam_type 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   222
by (rtac snd_type 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   223
by (rtac fst_type 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   224
by (resolve_tac [consI1 RSN (2, apply_type)] 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   225
by (fast_tac (ZF_cs addSIs [fun_weaken_type, bij_is_fun]) 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   226
val lemma2 = result();
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   227
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   228
goalw thy AC_defs "!!Z. AC9 ==> AC1";
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   229
by (rtac allI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   230
by (rtac impI 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   231
by (etac allE 1);
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   232
by (excluded_middle_tac "A=0" 1);
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   233
by (fast_tac (FOL_cs addSIs [empty_fun]) 2);
1207
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   234
by (etac impE 1);
3f460842e919 Ran expandshort and changed spelling of Grabczewski
lcp
parents: 1196
diff changeset
   235
by (rtac lemma1 1 THEN (REPEAT (assume_tac 1)));
1123
5dfdc1464966 Krzysztof Grabczewski's (nearly) complete AC proofs
lcp
parents:
diff changeset
   236
by (fast_tac (AC_cs addSEs [lemma2]) 1);
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents: 1123
diff changeset
   237
qed "AC9_AC1";