1461
|
1 |
(* Title: ZF/AC/AC_Equiv.ML
|
991
|
2 |
ID: $Id$
|
1461
|
3 |
Author: Krzysztof Grabczewski
|
991
|
4 |
|
|
5 |
*)
|
|
6 |
|
|
7 |
|
|
8 |
open AC_Equiv;
|
|
9 |
|
1071
|
10 |
val WO_defs = [WO1_def, WO2_def, WO3_def, WO4_def, WO5_def, WO6_def, WO8_def];
|
|
11 |
|
|
12 |
val AC_defs = [AC0_def, AC1_def, AC2_def, AC3_def, AC4_def, AC5_def,
|
|
13 |
AC6_def, AC7_def, AC8_def, AC9_def, AC10_def, AC11_def,
|
|
14 |
AC12_def, AC13_def, AC14_def, AC15_def, AC16_def,
|
|
15 |
AC17_def, AC18_def, AC19_def];
|
|
16 |
|
991
|
17 |
val AC_aux_defs = [pairwise_disjoint_def, sets_of_size_between_def];
|
1123
|
18 |
|
991
|
19 |
(* ******************************************** *)
|
|
20 |
|
|
21 |
(* Theorems analogous to ones presented in "ZF/Ordinal.ML" *)
|
|
22 |
|
|
23 |
(* lemma for nat_le_imp_lepoll *)
|
|
24 |
val [prem] = goalw Cardinal.thy [lepoll_def]
|
|
25 |
"m:nat ==> ALL n: nat. m le n --> m lepoll n";
|
|
26 |
by (nat_ind_tac "m" [prem] 1);
|
2469
|
27 |
by (fast_tac (!claset addSIs [le_imp_subset RS id_subset_inj]) 1);
|
991
|
28 |
by (rtac ballI 1);
|
|
29 |
by (eres_inst_tac [("n","n")] natE 1);
|
2469
|
30 |
by (asm_simp_tac (!simpset addsimps [inj_def, succI1 RS Pi_empty2]) 1);
|
|
31 |
by (fast_tac (!claset addSIs [le_imp_subset RS id_subset_inj]) 1);
|
991
|
32 |
val nat_le_imp_lepoll_lemma = result();
|
|
33 |
|
|
34 |
(* used in : AC10-AC15.ML WO1-WO6.ML WO6WO1.ML*)
|
|
35 |
val nat_le_imp_lepoll = nat_le_imp_lepoll_lemma RS bspec RS mp |> standard;
|
|
36 |
|
|
37 |
(* ********************************************************************** *)
|
|
38 |
(* lemmas concerning FOL and pure ZF theory *)
|
|
39 |
(* ********************************************************************** *)
|
|
40 |
|
1123
|
41 |
goal thy "!!X. (A->X)=0 ==> X=0";
|
2469
|
42 |
by (fast_tac (!claset addSIs [equals0I] addEs [lam_type RSN (2, equals0D)]) 1);
|
1123
|
43 |
val fun_space_emptyD = result();
|
991
|
44 |
|
|
45 |
(* used only in WO1_DC.ML *)
|
|
46 |
(*Note simpler proof*)
|
2469
|
47 |
goal upair.thy "!!A f g. [| ALL x:A. f`x=g`x; f:Df->Cf; g:Dg->Cg; \
|
991
|
48 |
\ A<=Df; A<=Dg |] ==> f``A=g``A";
|
2469
|
49 |
by (asm_simp_tac (!simpset addsimps [image_fun]) 1);
|
991
|
50 |
val images_eq = result();
|
|
51 |
|
|
52 |
(* used in : AC10-AC15.ML AC16WO4.ML WO6WO1.ML *)
|
|
53 |
(*I don't know where to put this one.*)
|
|
54 |
goal Cardinal.thy
|
|
55 |
"!!m A B. [| A lepoll succ(m); B<=A; B~=0 |] ==> A-B lepoll m";
|
1207
|
56 |
by (rtac not_emptyE 1 THEN (assume_tac 1));
|
991
|
57 |
by (forward_tac [singleton_subsetI] 1);
|
1196
|
58 |
by (forward_tac [subsetD] 1 THEN (assume_tac 1));
|
991
|
59 |
by (res_inst_tac [("A2","A")]
|
1057
|
60 |
(Diff_sing_lepoll RSN (2, subset_imp_lepoll RS lepoll_trans)) 1
|
1196
|
61 |
THEN (REPEAT (assume_tac 2)));
|
2469
|
62 |
by (Fast_tac 1);
|
991
|
63 |
val Diff_lepoll = result();
|
|
64 |
|
|
65 |
(* ********************************************************************** *)
|
|
66 |
(* lemmas concerning lepoll and eqpoll relations *)
|
|
67 |
(* ********************************************************************** *)
|
|
68 |
|
|
69 |
(* ********************************************************************** *)
|
|
70 |
(* Theorems concerning ordinals *)
|
|
71 |
(* ********************************************************************** *)
|
|
72 |
|
|
73 |
(* lemma for ordertype_Int *)
|
|
74 |
goalw Cardinal.thy [rvimage_def] "rvimage(A,id(A),r) = r Int A*A";
|
1207
|
75 |
by (rtac equalityI 1);
|
2469
|
76 |
by (Step_tac 1);
|
991
|
77 |
by (dres_inst_tac [("P","%a. <id(A)`xb,a>:r")] (id_conv RS subst) 1
|
1196
|
78 |
THEN (assume_tac 1));
|
991
|
79 |
by (dres_inst_tac [("P","%a. <a,ya>:r")] (id_conv RS subst) 1
|
1196
|
80 |
THEN (REPEAT (assume_tac 1)));
|
2469
|
81 |
by (fast_tac (!claset addIs [id_conv RS ssubst]) 1);
|
991
|
82 |
val rvimage_id = result();
|
|
83 |
|
|
84 |
(* used only in Hartog.ML *)
|
|
85 |
goal Cardinal.thy
|
1461
|
86 |
"!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)";
|
991
|
87 |
by (res_inst_tac [("P","%a. ordertype(A,a)=ordertype(A,r)")]
|
|
88 |
(rvimage_id RS subst) 1);
|
|
89 |
by (eresolve_tac [id_bij RS bij_ordertype_vimage] 1);
|
|
90 |
val ordertype_Int = result();
|
|
91 |
|
|
92 |
(* used only in AC16_lemmas.ML *)
|
|
93 |
goalw CardinalArith.thy [InfCard_def]
|
1461
|
94 |
"!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)";
|
2469
|
95 |
by (asm_simp_tac (!simpset addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
|
991
|
96 |
val Inf_Card_is_InfCard = result();
|
1123
|
97 |
|
|
98 |
goal thy "(THE z. {x}={z}) = x";
|
2469
|
99 |
by (fast_tac (!claset addSIs [the_equality]
|
1461
|
100 |
addSEs [singleton_eq_iff RS iffD1 RS sym]) 1);
|
1123
|
101 |
val the_element = result();
|
|
102 |
|
|
103 |
goal thy "(lam x:A. {x}) : bij(A, {{x}. x:A})";
|
|
104 |
by (res_inst_tac [("d","%z. THE x. z={x}")] lam_bijective 1);
|
|
105 |
by (TRYALL (eresolve_tac [RepFunI, RepFunE]));
|
2469
|
106 |
by (REPEAT (asm_full_simp_tac (!simpset addsimps [the_element]) 1));
|
1123
|
107 |
val lam_sing_bij = result();
|
|
108 |
|
|
109 |
val [major,minor] = goal thy
|
1461
|
110 |
"[| f : Pi(A,B); (!!x. x:A ==> B(x)<=C(x)) |] ==> f : Pi(A,C)";
|
2469
|
111 |
by (fast_tac (!claset addSIs [major RS Pi_type, minor RS subsetD,
|
1461
|
112 |
major RS apply_type]) 1);
|
1123
|
113 |
val Pi_weaken_type = result();
|
|
114 |
|
|
115 |
val [major, minor] = goalw thy [inj_def]
|
1461
|
116 |
"[| f:inj(A, B); (!!a. a:A ==> f`a : C) |] ==> f:inj(A,C)";
|
2469
|
117 |
by (fast_tac (!claset addSEs [minor]
|
1461
|
118 |
addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1);
|
1123
|
119 |
val inj_strengthen_type = result();
|
|
120 |
|
|
121 |
goal thy "A*B=0 <-> A=0 | B=0";
|
2469
|
122 |
by (fast_tac (!claset addSIs [equals0I] addEs [equals0D]) 1);
|
1123
|
123 |
val Sigma_empty_iff = result();
|
|
124 |
|
1196
|
125 |
goalw thy [Finite_def] "!!n. n:nat ==> Finite(n)";
|
2469
|
126 |
by (fast_tac (!claset addSIs [eqpoll_refl]) 1);
|
1196
|
127 |
val nat_into_Finite = result();
|
|
128 |
|
|
129 |
goalw thy [Finite_def] "~Finite(nat)";
|
2469
|
130 |
by (fast_tac (!claset addSDs [eqpoll_imp_lepoll]
|
1461
|
131 |
addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1);
|
1196
|
132 |
val nat_not_Finite = result();
|
|
133 |
|
|
134 |
val le_imp_lepoll = le_imp_subset RS subset_imp_lepoll;
|
|
135 |
|
|
136 |
(* ********************************************************************** *)
|
1461
|
137 |
(* Another elimination rule for EX! *)
|
1196
|
138 |
(* ********************************************************************** *)
|
|
139 |
|
|
140 |
goal thy "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y";
|
1207
|
141 |
by (etac ex1E 1);
|
1196
|
142 |
by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1);
|
2469
|
143 |
by (Fast_tac 1);
|
|
144 |
by (Fast_tac 1);
|
1196
|
145 |
val ex1_two_eq = result();
|
|
146 |
|
|
147 |
(* ********************************************************************** *)
|
1461
|
148 |
(* image of a surjection *)
|
1196
|
149 |
(* ********************************************************************** *)
|
|
150 |
|
|
151 |
goalw thy [surj_def] "!!f. f : surj(A, B) ==> f``A = B";
|
1207
|
152 |
by (etac CollectE 1);
|
1196
|
153 |
by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 THEN (assume_tac 1));
|
2469
|
154 |
by (fast_tac (!claset addSEs [RepFunE, apply_type]
|
1461
|
155 |
addSIs [RepFunI] addIs [equalityI]) 1);
|
1196
|
156 |
val surj_image_eq = result();
|
|
157 |
|
|
158 |
|
|
159 |
goal thy "!!y. succ(x) lepoll y ==> y ~= 0";
|
2469
|
160 |
by (fast_tac (!claset addSDs [lepoll_0_is_0]) 1);
|
1196
|
161 |
val succ_lepoll_imp_not_empty = result();
|
|
162 |
|
|
163 |
goal thy "!!x. x eqpoll succ(n) ==> x ~= 0";
|
2469
|
164 |
by (fast_tac (!claset addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
|
1196
|
165 |
val eqpoll_succ_imp_not_empty = result();
|
|
166 |
|