author | wenzelm |
Mon, 22 Oct 2001 17:58:11 +0200 | |
changeset 11879 | 1a386a1e002c |
parent 11317 | 7f9e4c389318 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/AC/HH.ML |
1123 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Krzysztof Grabczewski |
1123 | 4 |
|
5 |
Some properties of the recursive definition of HH used in the proofs of |
|
6 |
AC17 ==> AC1 |
|
7 |
AC1 ==> WO2 |
|
8 |
AC15 ==> WO6 |
|
9 |
*) |
|
10 |
||
11 |
(* ********************************************************************** *) |
|
1461 | 12 |
(* Lemmas useful in each of the three proofs *) |
1123 | 13 |
(* ********************************************************************** *) |
14 |
||
5068 | 15 |
Goal "HH(f,x,a) = \ |
11317 | 16 |
\ (let z = x - (\\<Union>b \\<in> a. HH(f,x,b)) \ |
17 |
\ in if(f`z \\<in> Pow(z)-{0}, f`z, {x}))"; |
|
1123 | 18 |
by (resolve_tac [HH_def RS def_transrec RS trans] 1); |
2469 | 19 |
by (Simp_tac 1); |
3731 | 20 |
qed "HH_def_satisfies_eq"; |
1123 | 21 |
|
11317 | 22 |
Goal "HH(f,x,a) \\<in> Pow(x)-{0} | HH(f,x,a)={x}"; |
1123 | 23 |
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); |
5137 | 24 |
by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); |
2469 | 25 |
by (Fast_tac 1); |
3731 | 26 |
qed "HH_values"; |
1123 | 27 |
|
11317 | 28 |
Goal "B \\<subseteq> A ==> X-(\\<Union>a \\<in> A. P(a)) = X-(\\<Union>a \\<in> A-B. P(a))-(\\<Union>b \\<in> B. P(b))"; |
2496 | 29 |
by (Fast_tac 1); |
3731 | 30 |
qed "subset_imp_Diff_eq"; |
1123 | 31 |
|
11317 | 32 |
Goal "[| c \\<in> a-b; b<a |] ==> c=b | b<c & c<a"; |
1207 | 33 |
by (etac ltE 1); |
3731 | 34 |
by (dtac Ord_linear 1); |
4091 | 35 |
by (fast_tac (claset() addSIs [ltI] addIs [Ord_in_Ord]) 2); |
36 |
by (fast_tac (claset() addEs [Ord_in_Ord]) 1); |
|
3731 | 37 |
qed "Ord_DiffE"; |
1123 | 38 |
|
11317 | 39 |
val prems = goal thy "(!!y. y \\<in> A ==> P(y) = {x}) ==> x - (\\<Union>y \\<in> A. P(y)) = x"; |
4091 | 40 |
by (asm_full_simp_tac (simpset() addsimps prems) 1); |
41 |
by (fast_tac (claset() addSDs [prem] addSEs [mem_irrefl]) 1); |
|
3731 | 42 |
qed "Diff_UN_eq_self"; |
1123 | 43 |
|
11317 | 44 |
Goal "x - (\\<Union>b \\<in> a. HH(f,x,b)) = x - (\\<Union>b \\<in> a1. HH(f,x,b)) \ |
1461 | 45 |
\ ==> HH(f,x,a) = HH(f,x,a1)"; |
1123 | 46 |
by (resolve_tac [HH_def_satisfies_eq RS |
1461 | 47 |
(HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1); |
1207 | 48 |
by (etac subst_context 1); |
3731 | 49 |
qed "HH_eq"; |
1123 | 50 |
|
5137 | 51 |
Goal "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}"; |
1123 | 52 |
by (res_inst_tac [("P","b<a")] impE 1 THEN REPEAT (assume_tac 2)); |
53 |
by (eresolve_tac [lt_Ord2 RS trans_induct] 1); |
|
1207 | 54 |
by (rtac impI 1); |
1123 | 55 |
by (resolve_tac [HH_eq RS trans] 1 THEN (assume_tac 2)); |
56 |
by (resolve_tac [leI RS le_imp_subset RS subset_imp_Diff_eq RS ssubst] 1 |
|
1461 | 57 |
THEN (assume_tac 1)); |
1123 | 58 |
by (res_inst_tac [("t","%z. z-?X")] subst_context 1); |
1207 | 59 |
by (rtac Diff_UN_eq_self 1); |
60 |
by (dtac Ord_DiffE 1 THEN (assume_tac 1)); |
|
4091 | 61 |
by (fast_tac (claset() addEs [ltE]) 1); |
3731 | 62 |
qed "HH_is_x_gt_too"; |
1123 | 63 |
|
11317 | 64 |
Goal "[| HH(f,x,a) \\<in> Pow(x)-{0}; b<a |] ==> HH(f,x,b) \\<in> Pow(x)-{0}"; |
1123 | 65 |
by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1)); |
1207 | 66 |
by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1)); |
67 |
by (dtac subst 1 THEN (assume_tac 1)); |
|
4091 | 68 |
by (fast_tac (claset() addSEs [mem_irrefl]) 1); |
3731 | 69 |
qed "HH_subset_x_lt_too"; |
1123 | 70 |
|
11317 | 71 |
Goal "HH(f,x,a) \\<in> Pow(x)-{0} \ |
72 |
\ ==> HH(f,x,a) \\<in> Pow(x - (\\<Union>b \\<in> a. HH(f,x,b)))-{0}"; |
|
1123 | 73 |
by (dresolve_tac [HH_def_satisfies_eq RS subst] 1); |
74 |
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); |
|
4091 | 75 |
by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); |
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5068
diff
changeset
|
76 |
by (dresolve_tac [split_if RS iffD1] 1); |
5137 | 77 |
by (Simp_tac 1); |
1200 | 78 |
by (fast_tac (subset_cs addSEs [mem_irrefl]) 1); |
3731 | 79 |
qed "HH_subset_x_imp_subset_Diff_UN"; |
1123 | 80 |
|
11317 | 81 |
Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v): Pow(x)-{0}; v \\<in> w |] ==> P"; |
82 |
by (forw_inst_tac [("P","%y. y \\<in> Pow(x)-{0}")] subst 1 THEN (assume_tac 1)); |
|
1123 | 83 |
by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1); |
1207 | 84 |
by (dtac subst_elem 1 THEN (assume_tac 1)); |
4091 | 85 |
by (fast_tac (claset() addSIs [singleton_iff RS iffD2, equals0I]) 1); |
3731 | 86 |
qed "HH_eq_arg_lt"; |
1123 | 87 |
|
5137 | 88 |
Goal "[| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0}; \ |
1461 | 89 |
\ Ord(v); Ord(w) |] ==> v=w"; |
1123 | 90 |
by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac); |
91 |
by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2 |
|
1461 | 92 |
THEN REPEAT (assume_tac 2)); |
1207 | 93 |
by (dtac subst_elem 1 THEN (assume_tac 1)); |
1123 | 94 |
by (fast_tac (FOL_cs addDs [ltD] addSEs [HH_eq_arg_lt]) 1); |
3731 | 95 |
qed "HH_eq_imp_arg_eq"; |
1123 | 96 |
|
5068 | 97 |
Goalw [lepoll_def, inj_def] |
11317 | 98 |
"[| HH(f, x, i) \\<in> Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"; |
99 |
by (res_inst_tac [("x","\\<lambda>j \\<in> i. HH(f, x, j)")] exI 1); |
|
2469 | 100 |
by (Asm_simp_tac 1); |
1123 | 101 |
by (fast_tac (FOL_cs addSEs [HH_eq_imp_arg_eq, Ord_in_Ord, HH_subset_x_lt_too] |
1461 | 102 |
addSIs [lam_type, ballI, ltI] addIs [bexI]) 1); |
3731 | 103 |
qed "HH_subset_x_imp_lepoll"; |
1123 | 104 |
|
5068 | 105 |
Goal "HH(f, x, Hartog(Pow(x)-{0})) = {x}"; |
1123 | 106 |
by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 2)); |
107 |
by (fast_tac (FOL_cs addSDs [HH_subset_x_imp_lepoll] |
|
1461 | 108 |
addSIs [Ord_Hartog] addSEs [HartogE]) 1); |
3731 | 109 |
qed "HH_Hartog_is_x"; |
1123 | 110 |
|
5068 | 111 |
Goal "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}"; |
4091 | 112 |
by (fast_tac (claset() addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1); |
3731 | 113 |
qed "HH_Least_eq_x"; |
1123 | 114 |
|
11317 | 115 |
Goal "a \\<in> (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \\<in> Pow(x)-{0}"; |
1123 | 116 |
by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1)); |
1207 | 117 |
by (rtac less_LeastE 1); |
1123 | 118 |
by (eresolve_tac [Ord_Least RSN (2, ltI)] 2); |
119 |
by (assume_tac 1); |
|
3731 | 120 |
qed "less_Least_subset_x"; |
1123 | 121 |
|
122 |
(* ********************************************************************** *) |
|
1461 | 123 |
(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) |
1123 | 124 |
(* ********************************************************************** *) |
125 |
||
5068 | 126 |
Goalw [inj_def] |
11317 | 127 |
"(\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ |
128 |
\ \\<in> inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"; |
|
2469 | 129 |
by (Asm_full_simp_tac 1); |
4091 | 130 |
by (fast_tac (claset() addSIs [lam_type] addDs [less_Least_subset_x] |
1461 | 131 |
addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1); |
3731 | 132 |
qed "lam_Least_HH_inj_Pow"; |
1123 | 133 |
|
11317 | 134 |
Goal "\\<forall>a \\<in> (LEAST i. HH(f,x,i)={x}). \\<exists>z \\<in> x. HH(f,x,a) = {z} \ |
135 |
\ ==> (\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ |
|
136 |
\ \\<in> inj(LEAST i. HH(f,x,i)={x}, {{y}. y \\<in> x})"; |
|
1123 | 137 |
by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1); |
2469 | 138 |
by (Asm_full_simp_tac 1); |
3731 | 139 |
qed "lam_Least_HH_inj"; |
1123 | 140 |
|
5068 | 141 |
Goalw [surj_def] |
11317 | 142 |
"[| x - (\\<Union>a \\<in> A. F(a)) = 0; \ |
143 |
\ \\<forall>a \\<in> A. \\<exists>z \\<in> x. F(a) = {z} |] \ |
|
144 |
\ ==> (\\<lambda>a \\<in> A. F(a)) \\<in> surj(A, {{y}. y \\<in> x})"; |
|
4091 | 145 |
by (asm_full_simp_tac (simpset() addsimps [lam_type, Diff_eq_0_iff]) 1); |
3731 | 146 |
by Safe_tac; |
2469 | 147 |
by (set_mp_tac 1); |
4091 | 148 |
by (deepen_tac (claset() addSIs [bexI] addSEs [equalityE]) 4 1); |
3731 | 149 |
qed "lam_surj_sing"; |
1123 | 150 |
|
11317 | 151 |
Goal "y \\<in> Pow(x)-{0} ==> x \\<noteq> 0"; |
5241 | 152 |
by Auto_tac; |
3731 | 153 |
qed "not_emptyI2"; |
1123 | 154 |
|
11317 | 155 |
Goal "f`(x - (\\<Union>j \\<in> i. HH(f,x,j))): Pow(x - (\\<Union>j \\<in> i. HH(f,x,j)))-{0} \ |
156 |
\ ==> HH(f, x, i) \\<in> Pow(x) - {0}"; |
|
1123 | 157 |
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); |
4091 | 158 |
by (asm_full_simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI, |
1461 | 159 |
not_emptyI2 RS if_P]) 1); |
2469 | 160 |
by (Fast_tac 1); |
3731 | 161 |
qed "f_subset_imp_HH_subset"; |
1123 | 162 |
|
11317 | 163 |
val [prem] = goal thy "(!!z. z \\<in> Pow(x)-{0} ==> f`z \\<in> Pow(z)-{0}) ==> \ |
164 |
\ x - (\\<Union>j \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"; |
|
165 |
by (excluded_middle_tac "?P \\<in> {0}" 1); |
|
2469 | 166 |
by (Fast_tac 2); |
1123 | 167 |
by (dresolve_tac [Diff_subset RS PowI RS DiffI RS prem RS |
1461 | 168 |
f_subset_imp_HH_subset] 1); |
4091 | 169 |
by (fast_tac (claset() addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)] |
1461 | 170 |
addSEs [mem_irrefl]) 1); |
3731 | 171 |
qed "f_subsets_imp_UN_HH_eq_x"; |
1123 | 172 |
|
11317 | 173 |
Goal "HH(f,x,i)=f`(x - (\\<Union>j \\<in> i. HH(f,x,j))) | HH(f,x,i)={x}"; |
1123 | 174 |
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); |
5137 | 175 |
by (simp_tac (simpset() addsimps [Let_def, Diff_subset RS PowI]) 1); |
3731 | 176 |
qed "HH_values2"; |
1123 | 177 |
|
11317 | 178 |
Goal "HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\\<Union>j \\<in> i. HH(f,x,j)))"; |
1123 | 179 |
by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1)); |
4091 | 180 |
by (fast_tac (claset() addSEs [equalityE, mem_irrefl] |
1461 | 181 |
addSDs [singleton_subsetD]) 1); |
3731 | 182 |
qed "HH_subset_imp_eq"; |
1123 | 183 |
|
11317 | 184 |
Goal "[| f \\<in> (Pow(x)-{0}) -> {{z}. z \\<in> x}; \ |
185 |
\ a \\<in> (LEAST i. HH(f,x,i)={x}) |] ==> \\<exists>z \\<in> x. HH(f,x,a) = {z}"; |
|
1207 | 186 |
by (dtac less_Least_subset_x 1); |
7499 | 187 |
by (ftac HH_subset_imp_eq 1); |
1207 | 188 |
by (dtac apply_type 1); |
1123 | 189 |
by (resolve_tac [Diff_subset RS PowI RS DiffI] 1); |
2493 | 190 |
by (fast_tac |
4091 | 191 |
(claset() addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1); |
192 |
by (fast_tac (claset() addss (simpset())) 1); |
|
3731 | 193 |
qed "f_sing_imp_HH_sing"; |
1123 | 194 |
|
5068 | 195 |
Goalw [bij_def] |
11317 | 196 |
"[| x - (\\<Union>j \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0; \ |
197 |
\ f \\<in> (Pow(x)-{0}) -> {{z}. z \\<in> x} |] \ |
|
198 |
\ ==> (\\<lambda>a \\<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ |
|
199 |
\ \\<in> bij(LEAST i. HH(f,x,i)={x}, {{y}. y \\<in> x})"; |
|
4091 | 200 |
by (fast_tac (claset() addSIs [lam_Least_HH_inj, lam_surj_sing, |
2493 | 201 |
f_sing_imp_HH_sing]) 1); |
3731 | 202 |
qed "f_sing_lam_bij"; |
1123 | 203 |
|
11317 | 204 |
Goal "f \\<in> (\\<Pi>X \\<in> Pow(x)-{0}. F(X)) \ |
205 |
\ ==> (\\<lambda>X \\<in> Pow(x)-{0}. {f`X}) \\<in> (\\<Pi>X \\<in> Pow(x)-{0}. {{z}. z \\<in> F(X)})"; |
|
1123 | 206 |
by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2] |
2493 | 207 |
addDs [apply_type]) 1); |
3731 | 208 |
qed "lam_singI"; |
1123 | 209 |
|
2493 | 210 |
val bij_Least_HH_x = |
211 |
(lam_singI RSN (2, [f_sing_lam_bij, lam_sing_bij RS bij_converse_bij] |
|
212 |
MRS comp_bij)) |> standard; |