| author | kleing | 
| Tue, 21 Nov 2000 13:48:47 +0100 | |
| changeset 10501 | 98fe9e987a17 | 
| parent 9305 | 3dfae8f90dcf | 
| child 11317 | 7f9e4c389318 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/AC/AC10_AC15.ML  | 
| 1123 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Krzysztof Grabczewski  | 
| 1123 | 4  | 
|
5  | 
The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.  | 
|
6  | 
We need the following:  | 
|
7  | 
||
8  | 
WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6  | 
|
9  | 
||
10  | 
In order to add the formulations AC13 and AC14 we need:  | 
|
11  | 
||
12  | 
AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15  | 
|
13  | 
||
14  | 
or  | 
|
15  | 
||
| 1196 | 16  | 
AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m le n)  | 
| 1123 | 17  | 
|
| 1196 | 18  | 
So we don't have to prove all implications of both cases.  | 
19  | 
Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as  | 
|
| 1123 | 20  | 
Rubin & Rubin do.  | 
21  | 
*)  | 
|
22  | 
||
23  | 
(* ********************************************************************** *)  | 
|
| 1461 | 24  | 
(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *)  | 
25  | 
(* or AC15 *)  | 
|
26  | 
(* - cons_times_nat_not_Finite *)  | 
|
27  | 
(* - ex_fun_AC13_AC15 *)  | 
|
| 1123 | 28  | 
(* ********************************************************************** *)  | 
29  | 
||
| 5137 | 30  | 
Goalw [lepoll_def] "A~=0 ==> B lepoll A*B";  | 
| 1204 | 31  | 
by (etac not_emptyE 1);  | 
| 1123 | 32  | 
by (res_inst_tac [("x","lam z:B. <x,z>")] exI 1);
 | 
| 4091 | 33  | 
by (fast_tac (claset() addSIs [snd_conv, lam_injective]) 1);  | 
| 3731 | 34  | 
qed "lepoll_Sigma";  | 
| 1123 | 35  | 
|
| 5137 | 36  | 
Goal "0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)";
 | 
| 1204 | 37  | 
by (rtac ballI 1);  | 
38  | 
by (etac RepFunE 1);  | 
|
| 1123 | 39  | 
by (hyp_subst_tac 1);  | 
| 1204 | 40  | 
by (rtac notI 1);  | 
| 1123 | 41  | 
by (dresolve_tac [subset_consI RS subset_imp_lepoll RS lepoll_Finite] 1);  | 
42  | 
by (resolve_tac [lepoll_Sigma RS lepoll_Finite RS (nat_not_Finite RS notE)] 1  | 
|
| 1461 | 43  | 
THEN (assume_tac 2));  | 
| 2469 | 44  | 
by (Fast_tac 1);  | 
| 3731 | 45  | 
qed "cons_times_nat_not_Finite";  | 
| 1123 | 46  | 
|
| 5137 | 47  | 
Goal "[| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A";  | 
| 2469 | 48  | 
by (Fast_tac 1);  | 
| 1123 | 49  | 
val lemma1 = result();  | 
50  | 
||
| 5067 | 51  | 
Goalw [pairwise_disjoint_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
52  | 
"[| pairwise_disjoint(A); B:A; C:A; a:B; a:C |] ==> B=C";  | 
| 1204 | 53  | 
by (dtac IntI 1 THEN (assume_tac 1));  | 
| 1123 | 54  | 
by (dres_inst_tac [("A","B Int C")] not_emptyI 1);
 | 
| 2469 | 55  | 
by (Fast_tac 1);  | 
| 1123 | 56  | 
val lemma2 = result();  | 
57  | 
||
| 5067 | 58  | 
Goalw [sets_of_size_between_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
59  | 
        "ALL B:{cons(0, x*nat). x:A}. pairwise_disjoint(f`B) &  \
 | 
| 1461 | 60  | 
\ sets_of_size_between(f`B, 2, n) & Union(f`B)=B \  | 
61  | 
\ ==> ALL B:A. EX! u. u:f`cons(0, B*nat) & u <= cons(0, B*nat) & \  | 
|
62  | 
\ 0:u & 2 lepoll u & u lepoll n";  | 
|
| 1204 | 63  | 
by (rtac ballI 1);  | 
64  | 
by (etac ballE 1);  | 
|
| 2469 | 65  | 
by (Fast_tac 2);  | 
| 1204 | 66  | 
by (REPEAT (etac conjE 1));  | 
| 1123 | 67  | 
by (dresolve_tac [consI1 RSN (2, lemma1)] 1);  | 
| 1204 | 68  | 
by (etac bexE 1);  | 
69  | 
by (rtac ex1I 1);  | 
|
| 2469 | 70  | 
by (Fast_tac 1);  | 
| 1204 | 71  | 
by (REPEAT (etac conjE 1));  | 
72  | 
by (rtac lemma2 1 THEN (REPEAT (assume_tac 1)));  | 
|
| 1123 | 73  | 
val lemma3 = result();  | 
74  | 
||
| 5137 | 75  | 
Goalw [lepoll_def] "[| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i";
 | 
| 1204 | 76  | 
by (etac exE 1);  | 
| 1123 | 77  | 
by (res_inst_tac  | 
| 1461 | 78  | 
        [("x", "lam x:RepFun(A, P). LEAST j. EX a:A. x=P(a) & f`a=j")] exI 1);
 | 
| 1123 | 79  | 
by (res_inst_tac [("d", "%y. P(converse(f)`y)")] lam_injective 1);
 | 
| 1204 | 80  | 
by (etac RepFunE 1);  | 
| 1196 | 81  | 
by (forward_tac [inj_is_fun RS apply_type] 1 THEN (assume_tac 1));  | 
| 4091 | 82  | 
by (fast_tac (claset() addIs [LeastI2]  | 
| 1461 | 83  | 
addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);  | 
| 1204 | 84  | 
by (etac RepFunE 1);  | 
85  | 
by (rtac LeastI2 1);  | 
|
| 2469 | 86  | 
by (Fast_tac 1);  | 
| 4091 | 87  | 
by (fast_tac (claset() addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);  | 
88  | 
by (fast_tac (claset() addEs [sym, left_inverse RS ssubst]) 1);  | 
|
| 1123 | 89  | 
val lemma4 = result();  | 
90  | 
||
| 5137 | 91  | 
Goal "[| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B); \  | 
| 1461 | 92  | 
\ u(B) lepoll succ(n) |] \  | 
93  | 
\       ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 &  \
 | 
|
94  | 
\               (lam x:A. {fst(x). x:u(x)-{0}})`B <= B &  \
 | 
|
95  | 
\               (lam x:A. {fst(x). x:u(x)-{0}})`B lepoll n";
 | 
|
| 2469 | 96  | 
by (Asm_simp_tac 1);  | 
| 1204 | 97  | 
by (rtac conjI 1);  | 
| 1123 | 98  | 
by (fast_tac (empty_cs addSDs [RepFun_eq_0_iff RS iffD1]  | 
| 1461 | 99  | 
addDs [lepoll_Diff_sing]  | 
100  | 
addEs [lepoll_trans RS succ_lepoll_natE, ssubst]  | 
|
101  | 
addSIs [notI, lepoll_refl, nat_0I]) 1);  | 
|
| 1204 | 102  | 
by (rtac conjI 1);  | 
| 4091 | 103  | 
by (fast_tac (claset() addSIs [fst_type] addSEs [consE]) 1);  | 
104  | 
by (fast_tac (claset() addSEs [equalityE,  | 
|
| 1461 | 105  | 
Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1);  | 
| 1123 | 106  | 
val lemma5 = result();  | 
107  | 
||
| 5137 | 108  | 
Goal "[| EX f. ALL B:{cons(0, x*nat). x:A}.  \
 | 
| 1461 | 109  | 
\ pairwise_disjoint(f`B) & \  | 
110  | 
\ sets_of_size_between(f`B, 2, succ(n)) & \  | 
|
111  | 
\ Union(f`B)=B; n:nat |] \  | 
|
112  | 
\ ==> EX f. ALL B:A. f`B ~= 0 & f`B <= B & f`B lepoll n";  | 
|
| 1200 | 113  | 
by (fast_tac (empty_cs addSDs [lemma3, theI] addDs [bspec]  | 
| 1461 | 114  | 
addSEs [exE, conjE]  | 
115  | 
addIs [exI, ballI, lemma5]) 1);  | 
|
| 3731 | 116  | 
qed "ex_fun_AC13_AC15";  | 
| 1123 | 117  | 
|
118  | 
(* ********************************************************************** *)  | 
|
| 1461 | 119  | 
(* The target proofs *)  | 
| 1123 | 120  | 
(* ********************************************************************** *)  | 
121  | 
||
122  | 
(* ********************************************************************** *)  | 
|
| 1461 | 123  | 
(* AC10(n) ==> AC11 *)  | 
| 1123 | 124  | 
(* ********************************************************************** *)  | 
125  | 
||
| 5137 | 126  | 
Goalw AC_defs "[| n:nat; 1 le n; AC10(n) |] ==> AC11";  | 
| 1204 | 127  | 
by (rtac bexI 1 THEN (assume_tac 2));  | 
| 2469 | 128  | 
by (Fast_tac 1);  | 
| 1196 | 129  | 
qed "AC10_AC11";  | 
| 1123 | 130  | 
|
131  | 
(* ********************************************************************** *)  | 
|
| 1461 | 132  | 
(* AC11 ==> AC12 *)  | 
| 1123 | 133  | 
(* ********************************************************************** *)  | 
134  | 
||
| 5137 | 135  | 
Goalw AC_defs "AC11 ==> AC12";  | 
| 1123 | 136  | 
by (fast_tac (FOL_cs addSEs [bexE] addIs [bexI]) 1);  | 
| 1196 | 137  | 
qed "AC11_AC12";  | 
| 1123 | 138  | 
|
139  | 
(* ********************************************************************** *)  | 
|
| 1461 | 140  | 
(* AC12 ==> AC15 *)  | 
| 1123 | 141  | 
(* ********************************************************************** *)  | 
142  | 
||
| 5137 | 143  | 
Goalw AC_defs "AC12 ==> AC15";  | 
| 4152 | 144  | 
by Safe_tac;  | 
| 1204 | 145  | 
by (etac allE 1);  | 
146  | 
by (etac impE 1);  | 
|
147  | 
by (etac cons_times_nat_not_Finite 1);  | 
|
| 4091 | 148  | 
by (fast_tac (claset() addSIs [ex_fun_AC13_AC15]) 1);  | 
| 1196 | 149  | 
qed "AC12_AC15";  | 
| 1123 | 150  | 
|
151  | 
(* ********************************************************************** *)  | 
|
| 1461 | 152  | 
(* AC15 ==> WO6 *)  | 
| 1123 | 153  | 
(* ********************************************************************** *)  | 
154  | 
||
155  | 
(* in a separate file *)  | 
|
156  | 
||
157  | 
(* ********************************************************************** *)  | 
|
| 1461 | 158  | 
(* The proof needed in the first case, not in the second *)  | 
| 1123 | 159  | 
(* ********************************************************************** *)  | 
160  | 
||
161  | 
(* ********************************************************************** *)  | 
|
| 1461 | 162  | 
(* AC10(n) ==> AC13(n-1) if 2 le n *)  | 
163  | 
(* *)  | 
|
164  | 
(* Because of the change to the formal definition of AC10(n) we prove *)  | 
|
165  | 
(* the following obviously equivalent theorem : *)  | 
|
166  | 
(* AC10(n) implies AC13(n) for (1 le n) *)  | 
|
| 1123 | 167  | 
(* ********************************************************************** *)  | 
168  | 
||
| 5137 | 169  | 
Goalw AC_defs "[| n:nat; 1 le n; AC10(n) |] ==> AC13(n)";  | 
| 4152 | 170  | 
by Safe_tac;  | 
| 1123 | 171  | 
by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE),  | 
| 1461 | 172  | 
ex_fun_AC13_AC15]) 1);  | 
| 1200 | 173  | 
qed "AC10_AC13";  | 
| 1123 | 174  | 
|
175  | 
(* ********************************************************************** *)  | 
|
| 1461 | 176  | 
(* The proofs needed in the second case, not in the first *)  | 
| 1123 | 177  | 
(* ********************************************************************** *)  | 
178  | 
||
179  | 
(* ********************************************************************** *)  | 
|
| 1461 | 180  | 
(* AC1 ==> AC13(1) *)  | 
| 1123 | 181  | 
(* ********************************************************************** *)  | 
182  | 
||
| 5137 | 183  | 
Goalw AC_defs "AC1 ==> AC13(1)";  | 
| 1204 | 184  | 
by (rtac allI 1);  | 
185  | 
by (etac allE 1);  | 
|
186  | 
by (rtac impI 1);  | 
|
| 1123 | 187  | 
by (mp_tac 1);  | 
| 1204 | 188  | 
by (etac exE 1);  | 
| 1123 | 189  | 
by (res_inst_tac [("x","lam x:A. {f`x}")] exI 1);
 | 
| 9305 | 190  | 
by (asm_simp_tac (simpset() addsimps  | 
191  | 
[singleton_eqpoll_1 RS eqpoll_imp_lepoll,  | 
|
192  | 
singletonI RS not_emptyI]) 1);  | 
|
| 1196 | 193  | 
qed "AC1_AC13";  | 
| 1123 | 194  | 
|
195  | 
(* ********************************************************************** *)  | 
|
| 1461 | 196  | 
(* AC13(m) ==> AC13(n) for m <= n *)  | 
| 1123 | 197  | 
(* ********************************************************************** *)  | 
198  | 
||
| 6070 | 199  | 
Goalw AC_defs "[| m le n; AC13(m) |] ==> AC13(n)";  | 
200  | 
by (dtac le_imp_lepoll 1);  | 
|
| 4091 | 201  | 
by (fast_tac (claset() addSEs [lepoll_trans]) 1);  | 
| 1196 | 202  | 
qed "AC13_mono";  | 
| 1123 | 203  | 
|
204  | 
(* ********************************************************************** *)  | 
|
| 1461 | 205  | 
(* The proofs necessary for both cases *)  | 
| 1123 | 206  | 
(* ********************************************************************** *)  | 
207  | 
||
208  | 
(* ********************************************************************** *)  | 
|
| 1461 | 209  | 
(* AC13(n) ==> AC14 if 1 <= n *)  | 
| 1123 | 210  | 
(* ********************************************************************** *)  | 
211  | 
||
| 5137 | 212  | 
Goalw AC_defs "[| n:nat; 1 le n; AC13(n) |] ==> AC14";  | 
| 1123 | 213  | 
by (fast_tac (FOL_cs addIs [bexI]) 1);  | 
| 1196 | 214  | 
qed "AC13_AC14";  | 
| 1123 | 215  | 
|
216  | 
(* ********************************************************************** *)  | 
|
| 1461 | 217  | 
(* AC14 ==> AC15 *)  | 
| 1123 | 218  | 
(* ********************************************************************** *)  | 
219  | 
||
| 5137 | 220  | 
Goalw AC_defs "AC14 ==> AC15";  | 
| 2469 | 221  | 
by (Fast_tac 1);  | 
| 1196 | 222  | 
qed "AC14_AC15";  | 
| 1123 | 223  | 
|
224  | 
(* ********************************************************************** *)  | 
|
| 1461 | 225  | 
(* The redundant proofs; however cited by Rubin & Rubin *)  | 
| 1123 | 226  | 
(* ********************************************************************** *)  | 
227  | 
||
228  | 
(* ********************************************************************** *)  | 
|
| 1461 | 229  | 
(* AC13(1) ==> AC1 *)  | 
| 1123 | 230  | 
(* ********************************************************************** *)  | 
231  | 
||
| 5137 | 232  | 
Goal "[| A~=0; A lepoll 1 |] ==> EX a. A={a}";
 | 
| 4091 | 233  | 
by (fast_tac (claset() addSEs [not_emptyE, lepoll_1_is_sing]) 1);  | 
| 3731 | 234  | 
qed "lemma_aux";  | 
| 1123 | 235  | 
|
| 5137 | 236  | 
Goal "ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1 \  | 
| 9305 | 237  | 
\     ==> (lam x:A. THE y. f(x)={y}) : (PROD X:A. X)";
 | 
| 1204 | 238  | 
by (rtac lam_type 1);  | 
239  | 
by (dtac bspec 1 THEN (assume_tac 1));  | 
|
240  | 
by (REPEAT (etac conjE 1));  | 
|
| 1196 | 241  | 
by (eresolve_tac [lemma_aux RS exE] 1 THEN (assume_tac 1));  | 
| 4091 | 242  | 
by (asm_full_simp_tac (simpset() addsimps [the_element]) 1);  | 
| 1123 | 243  | 
val lemma = result();  | 
244  | 
||
| 5137 | 245  | 
Goalw AC_defs "AC13(1) ==> AC1";  | 
| 4091 | 246  | 
by (fast_tac (claset() addSEs [lemma]) 1);  | 
| 1200 | 247  | 
qed "AC13_AC1";  | 
| 1123 | 248  | 
|
249  | 
(* ********************************************************************** *)  | 
|
| 1461 | 250  | 
(* AC11 ==> AC14 *)  | 
| 1123 | 251  | 
(* ********************************************************************** *)  | 
252  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
253  | 
Goalw [AC11_def, AC14_def] "AC11 ==> AC14";  | 
| 4091 | 254  | 
by (fast_tac (claset() addSIs [AC10_AC13]) 1);  | 
| 1196 | 255  | 
qed "AC11_AC14";  |