author | nipkow |
Tue, 12 Jan 1999 15:48:59 +0100 | |
changeset 6099 | d4866f6ff2f9 |
parent 6068 | 2d8f3e1f1151 |
child 6153 | bff90585cce5 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Cardinal_AC.ML |
484 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
484 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Cardinal arithmetic WITH the Axiom of Choice |
|
517 | 7 |
|
8 |
These results help justify infinite-branching datatypes |
|
484 | 9 |
*) |
10 |
||
11 |
open Cardinal_AC; |
|
12 |
||
13 |
(*** Strengthened versions of existing theorems about cardinals ***) |
|
14 |
||
5067 | 15 |
Goal "|A| eqpoll A"; |
484 | 16 |
by (resolve_tac [AC_well_ord RS exE] 1); |
1461 | 17 |
by (etac well_ord_cardinal_eqpoll 1); |
760 | 18 |
qed "cardinal_eqpoll"; |
484 | 19 |
|
20 |
val cardinal_idem = cardinal_eqpoll RS cardinal_cong; |
|
21 |
||
5137 | 22 |
Goal "|X| = |Y| ==> X eqpoll Y"; |
484 | 23 |
by (resolve_tac [AC_well_ord RS exE] 1); |
24 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
1461 | 25 |
by (rtac well_ord_cardinal_eqE 1); |
484 | 26 |
by (REPEAT_SOME assume_tac); |
760 | 27 |
qed "cardinal_eqE"; |
484 | 28 |
|
5067 | 29 |
Goal "|X| = |Y| <-> X eqpoll Y"; |
4091 | 30 |
by (blast_tac (claset() addIs [cardinal_cong, cardinal_eqE]) 1); |
1609 | 31 |
qed "cardinal_eqpoll_iff"; |
32 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
33 |
Goal "[| |A|=|B|; |C|=|D|; A Int C = 0; B Int D = 0 |] ==> \ |
1609 | 34 |
\ |A Un C| = |B Un D|"; |
4091 | 35 |
by (asm_full_simp_tac (simpset() addsimps [cardinal_eqpoll_iff, |
2033 | 36 |
eqpoll_disjoint_Un]) 1); |
1609 | 37 |
qed "cardinal_disjoint_Un"; |
38 |
||
5137 | 39 |
Goal "A lepoll B ==> |A| le |B|"; |
484 | 40 |
by (resolve_tac [AC_well_ord RS exE] 1); |
1461 | 41 |
by (etac well_ord_lepoll_imp_Card_le 1); |
484 | 42 |
by (assume_tac 1); |
766 | 43 |
qed "lepoll_imp_Card_le"; |
484 | 44 |
|
5067 | 45 |
Goal "(i |+| j) |+| k = i |+| (j |+| k)"; |
484 | 46 |
by (resolve_tac [AC_well_ord RS exE] 1); |
47 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
48 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
1461 | 49 |
by (rtac well_ord_cadd_assoc 1); |
484 | 50 |
by (REPEAT_SOME assume_tac); |
760 | 51 |
qed "cadd_assoc"; |
484 | 52 |
|
5067 | 53 |
Goal "(i |*| j) |*| k = i |*| (j |*| k)"; |
484 | 54 |
by (resolve_tac [AC_well_ord RS exE] 1); |
55 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
56 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
1461 | 57 |
by (rtac well_ord_cmult_assoc 1); |
484 | 58 |
by (REPEAT_SOME assume_tac); |
760 | 59 |
qed "cmult_assoc"; |
484 | 60 |
|
5067 | 61 |
Goal "(i |+| j) |*| k = (i |*| k) |+| (j |*| k)"; |
847 | 62 |
by (resolve_tac [AC_well_ord RS exE] 1); |
63 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
64 |
by (resolve_tac [AC_well_ord RS exE] 1); |
|
1461 | 65 |
by (rtac well_ord_cadd_cmult_distrib 1); |
847 | 66 |
by (REPEAT_SOME assume_tac); |
67 |
qed "cadd_cmult_distrib"; |
|
68 |
||
5137 | 69 |
Goal "InfCard(|A|) ==> A*A eqpoll A"; |
484 | 70 |
by (resolve_tac [AC_well_ord RS exE] 1); |
1461 | 71 |
by (etac well_ord_InfCard_square_eq 1); |
484 | 72 |
by (assume_tac 1); |
760 | 73 |
qed "InfCard_square_eq"; |
484 | 74 |
|
75 |
||
76 |
(*** Other applications of AC ***) |
|
77 |
||
5137 | 78 |
Goal "|A| le |B| ==> A lepoll B"; |
484 | 79 |
by (resolve_tac [cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll RS |
1461 | 80 |
lepoll_trans] 1); |
484 | 81 |
by (eresolve_tac [le_imp_subset RS subset_imp_lepoll RS lepoll_trans] 1); |
82 |
by (resolve_tac [cardinal_eqpoll RS eqpoll_imp_lepoll] 1); |
|
766 | 83 |
qed "Card_le_imp_lepoll"; |
484 | 84 |
|
5137 | 85 |
Goal "Card(K) ==> |A| le K <-> A lepoll K"; |
484 | 86 |
by (eresolve_tac [Card_cardinal_eq RS subst] 1 THEN |
87 |
rtac iffI 1 THEN |
|
766 | 88 |
DEPTH_SOLVE (eresolve_tac [Card_le_imp_lepoll,lepoll_imp_Card_le] 1)); |
760 | 89 |
qed "le_Card_iff"; |
484 | 90 |
|
5137 | 91 |
Goalw [surj_def] "f: surj(X,Y) ==> EX g. g: inj(Y,X)"; |
484 | 92 |
by (etac CollectE 1); |
93 |
by (res_inst_tac [("A1", "Y"), ("B1", "%y. f-``{y}")] (AC_Pi RS exE) 1); |
|
4091 | 94 |
by (fast_tac (claset() addSEs [apply_Pair]) 1); |
95 |
by (blast_tac (claset() addDs [apply_type, Pi_memberD] |
|
3016 | 96 |
addIs [apply_equality, Pi_type, f_imp_injective]) 1); |
760 | 97 |
qed "surj_implies_inj"; |
484 | 98 |
|
99 |
(*Kunen's Lemma 10.20*) |
|
5137 | 100 |
Goal "f: surj(X,Y) ==> |Y| le |X|"; |
1461 | 101 |
by (rtac lepoll_imp_Card_le 1); |
484 | 102 |
by (eresolve_tac [surj_implies_inj RS exE] 1); |
103 |
by (rewtac lepoll_def); |
|
1461 | 104 |
by (etac exI 1); |
760 | 105 |
qed "surj_implies_cardinal_le"; |
484 | 106 |
|
107 |
(*Kunen's Lemma 10.21*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
108 |
Goal "[| InfCard(K); ALL i:K. |X(i)| le K |] ==> |UN i:K. X(i)| le K"; |
4091 | 109 |
by (asm_full_simp_tac (simpset() addsimps [InfCard_is_Card, le_Card_iff]) 1); |
1461 | 110 |
by (rtac lepoll_trans 1); |
484 | 111 |
by (resolve_tac [InfCard_square_eq RS eqpoll_imp_lepoll] 2); |
4091 | 112 |
by (asm_simp_tac (simpset() addsimps [InfCard_is_Card, Card_cardinal_eq]) 2); |
1461 | 113 |
by (rewtac lepoll_def); |
484 | 114 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); |
115 |
by (etac (AC_ball_Pi RS exE) 1); |
|
1461 | 116 |
by (rtac exI 1); |
484 | 117 |
(*Lemma needed in both subgoals, for a fixed z*) |
118 |
by (subgoal_tac |
|
119 |
"ALL z: (UN i:K. X(i)). z: X(LEAST i. z:X(i)) & (LEAST i. z:X(i)) : K" 1); |
|
4091 | 120 |
by (fast_tac (claset() addSIs [Least_le RS lt_trans1 RS ltD, ltI] |
3016 | 121 |
addSEs [LeastI, Ord_in_Ord]) 2); |
484 | 122 |
by (res_inst_tac [("c", "%z. <LEAST i. z:X(i), f ` (LEAST i. z:X(i)) ` z>"), |
1461 | 123 |
("d", "%<i,j>. converse(f`i) ` j")] |
124 |
lam_injective 1); |
|
484 | 125 |
(*Instantiate the lemma proved above*) |
126 |
by (ALLGOALS ball_tac); |
|
4091 | 127 |
by (blast_tac (claset() addIs [inj_is_fun RS apply_type] |
3016 | 128 |
addDs [apply_type]) 1); |
1461 | 129 |
by (dtac apply_type 1); |
130 |
by (etac conjunct2 1); |
|
4091 | 131 |
by (asm_simp_tac (simpset() addsimps [left_inverse]) 1); |
760 | 132 |
qed "cardinal_UN_le"; |
484 | 133 |
|
488 | 134 |
(*The same again, using csucc*) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
135 |
Goal "[| InfCard(K); ALL i:K. |X(i)| < csucc(K) |] ==> \ |
484 | 136 |
\ |UN i:K. X(i)| < csucc(K)"; |
137 |
by (asm_full_simp_tac |
|
4091 | 138 |
(simpset() addsimps [Card_lt_csucc_iff, cardinal_UN_le, |
1461 | 139 |
InfCard_is_Card, Card_cardinal]) 1); |
760 | 140 |
qed "cardinal_UN_lt_csucc"; |
488 | 141 |
|
785 | 142 |
(*The same again, for a union of ordinals. In use, j(i) is a bit like rank(i), |
143 |
the least ordinal j such that i:Vfrom(A,j). *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
144 |
Goal "[| InfCard(K); ALL i:K. j(i) < csucc(K) |] ==> \ |
488 | 145 |
\ (UN i:K. j(i)) < csucc(K)"; |
146 |
by (resolve_tac [cardinal_UN_lt_csucc RS Card_lt_imp_lt] 1); |
|
147 |
by (assume_tac 1); |
|
4091 | 148 |
by (blast_tac (claset() addIs [Ord_cardinal_le RS lt_trans1] addEs [ltE]) 1); |
149 |
by (blast_tac (claset() addSIs [Ord_UN] addEs [ltE]) 1); |
|
488 | 150 |
by (eresolve_tac [InfCard_is_Card RS Card_is_Ord RS Card_csucc] 1); |
760 | 151 |
qed "cardinal_UN_Ord_lt_csucc"; |
488 | 152 |
|
517 | 153 |
|
785 | 154 |
(** Main result for infinite-branching datatypes. As above, but the index |
155 |
set need not be a cardinal. Surprisingly complicated proof! |
|
156 |
**) |
|
157 |
||
517 | 158 |
(*Saves checking Ord(j) below*) |
159 |
goal Ordinal.thy "!!i j. [| i <= j; j<k; Ord(i) |] ==> i<k"; |
|
160 |
by (resolve_tac [subset_imp_le RS lt_trans1] 1); |
|
161 |
by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); |
|
760 | 162 |
qed "lt_subset_trans"; |
517 | 163 |
|
785 | 164 |
(*Work backwards along the injection from W into K, given that W~=0.*) |
6068 | 165 |
Goal "[| f: inj(A,B); a:A |] ==> \ |
166 |
\ (UN x:A. C(x)) <= (UN y:B. C(if y: range(f) then converse(f)`y else a))"; |
|
1461 | 167 |
by (rtac UN_least 1); |
785 | 168 |
by (res_inst_tac [("x1", "f`x")] (UN_upper RSN (2,subset_trans)) 1); |
169 |
by (eresolve_tac [inj_is_fun RS apply_type] 2 THEN assume_tac 2); |
|
170 |
by (asm_simp_tac |
|
4091 | 171 |
(simpset() addsimps [inj_is_fun RS apply_rangeI, left_inverse]) 1); |
785 | 172 |
val inj_UN_subset = result(); |
173 |
||
174 |
(*Simpler to require |W|=K; we'd have a bijection; but the theorem would |
|
175 |
be weaker.*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
176 |
Goal "[| InfCard(K); |W| le K; ALL w:W. j(w) < csucc(K) |] ==> \ |
517 | 177 |
\ (UN w:W. j(w)) < csucc(K)"; |
178 |
by (excluded_middle_tac "W=0" 1); |
|
1461 | 179 |
by (asm_simp_tac (*solve the easy 0 case*) |
4091 | 180 |
(simpset() addsimps [UN_0, InfCard_is_Card, Card_is_Ord RS Card_csucc, |
1461 | 181 |
Card_is_Ord, Ord_0_lt_csucc]) 2); |
516 | 182 |
by (asm_full_simp_tac |
4091 | 183 |
(simpset() addsimps [InfCard_is_Card, le_Card_iff, lepoll_def]) 1); |
184 |
by (safe_tac (claset() addSIs [equalityI])); |
|
785 | 185 |
by (swap_res_tac [[inj_UN_subset, cardinal_UN_Ord_lt_csucc] |
1461 | 186 |
MRS lt_subset_trans] 1); |
785 | 187 |
by (REPEAT (assume_tac 1)); |
4091 | 188 |
by (blast_tac (claset() addSIs [Ord_UN] addEs [ltE]) 2); |
5137 | 189 |
by (asm_simp_tac (simpset() addsimps [inj_converse_fun RS apply_type]) 1); |
760 | 190 |
qed "le_UN_Ord_lt_csucc"; |
516 | 191 |