| author | wenzelm | 
| Mon, 15 Oct 2007 21:08:37 +0200 | |
| changeset 25043 | 32ed65dc1bf4 | 
| parent 16417 | 9bc16273c2d4 | 
| child 27678 | 85ea2be46c71 | 
| permissions | -rw-r--r-- | 
| 12776 | 1 | (* Title: ZF/AC/AC1_AC17.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Krzysztof Grabczewski | |
| 4 | ||
| 5 | The equivalence of AC0, AC1 and AC17 | |
| 6 | ||
| 7 | Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent | |
| 8 | to AC0 and AC1. | |
| 9 | *) | |
| 10 | ||
| 16417 | 11 | theory AC17_AC1 imports HH begin | 
| 12776 | 12 | |
| 13 | ||
| 14 | (** AC0 is equivalent to AC1. | |
| 15 | AC0 comes from Suppes, AC1 from Rubin & Rubin **) | |
| 16 | ||
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 17 | lemma AC0_AC1_lemma: "[| f:(\<Pi> X \<in> A. X); D \<subseteq> A |] ==> \<exists>g. g:(\<Pi> X \<in> D. X)" | 
| 12891 | 18 | by (fast intro!: lam_type apply_type) | 
| 12776 | 19 | |
| 20 | lemma AC0_AC1: "AC0 ==> AC1" | |
| 21 | apply (unfold AC0_def AC1_def) | |
| 22 | apply (blast intro: AC0_AC1_lemma) | |
| 23 | done | |
| 24 | ||
| 25 | lemma AC1_AC0: "AC1 ==> AC0" | |
| 26 | by (unfold AC0_def AC1_def, blast) | |
| 27 | ||
| 28 | ||
| 29 | (**** The proof of AC1 ==> AC17 ****) | |
| 30 | ||
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 31 | lemma AC1_AC17_lemma: "f \<in> (\<Pi> X \<in> Pow(A) - {0}. X) ==> f \<in> (Pow(A) - {0} -> A)"
 | 
| 12776 | 32 | apply (rule Pi_type, assumption) | 
| 33 | apply (drule apply_type, assumption, fast) | |
| 34 | done | |
| 35 | ||
| 36 | lemma AC1_AC17: "AC1 ==> AC17" | |
| 37 | apply (unfold AC1_def AC17_def) | |
| 38 | apply (rule allI) | |
| 39 | apply (rule ballI) | |
| 40 | apply (erule_tac x = "Pow (A) -{0}" in allE)
 | |
| 41 | apply (erule impE, fast) | |
| 42 | apply (erule exE) | |
| 43 | apply (rule bexI) | |
| 44 | apply (erule_tac [2] AC1_AC17_lemma) | |
| 45 | apply (rule apply_type, assumption) | |
| 46 | apply (fast dest!: AC1_AC17_lemma elim!: apply_type) | |
| 47 | done | |
| 48 | ||
| 49 | ||
| 50 | (**** The proof of AC17 ==> AC1 ****) | |
| 51 | ||
| 52 | (* *********************************************************************** *) | |
| 53 | (* more properties of HH *) | |
| 54 | (* *********************************************************************** *) | |
| 55 | ||
| 56 | lemma UN_eq_imp_well_ord: | |
| 57 |      "[| x - (\<Union>j \<in> LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x}.  
 | |
| 58 |         HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, j)) = 0;   
 | |
| 59 |         f \<in> Pow(x)-{0} -> x |]   
 | |
| 60 | ==> \<exists>r. well_ord(x,r)" | |
| 61 | apply (rule exI) | |
| 62 | apply (erule well_ord_rvimage | |
| 63 | [OF bij_Least_HH_x [THEN bij_converse_bij, THEN bij_is_inj] | |
| 64 | Ord_Least [THEN well_ord_Memrel]], assumption) | |
| 65 | done | |
| 66 | ||
| 67 | (* *********************************************************************** *) | |
| 68 | (* theorems closer to the proof *) | |
| 69 | (* *********************************************************************** *) | |
| 70 | ||
| 71 | lemma not_AC1_imp_ex: | |
| 72 |      "~AC1 ==> \<exists>A. \<forall>f \<in> Pow(A)-{0} -> A. \<exists>u \<in> Pow(A)-{0}. f`u \<notin> u"
 | |
| 73 | apply (unfold AC1_def) | |
| 74 | apply (erule swap) | |
| 75 | apply (rule allI) | |
| 76 | apply (erule swap) | |
| 77 | apply (rule_tac x = "Union (A)" in exI) | |
| 12891 | 78 | apply (blast intro: lam_type) | 
| 12776 | 79 | done | 
| 80 | ||
| 13535 | 81 | lemma AC17_AC1_aux1: | 
| 12776 | 82 |      "[| \<forall>f \<in> Pow(x) - {0} -> x. \<exists>u \<in> Pow(x) - {0}. f`u\<notin>u;   
 | 
| 83 |          \<exists>f \<in> Pow(x)-{0}->x.  
 | |
| 84 |             x - (\<Union>a \<in> (LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,i)={x}).   
 | |
| 85 |             HH(\<lambda>X \<in> Pow(x)-{0}. {f`X},x,a)) = 0 |]  
 | |
| 86 | ==> P" | |
| 87 | apply (erule bexE) | |
| 88 | apply (erule UN_eq_imp_well_ord [THEN exE], assumption) | |
| 89 | apply (erule ex_choice_fun_Pow [THEN exE]) | |
| 90 | apply (erule ballE) | |
| 91 | apply (fast intro: apply_type del: DiffE) | |
| 92 | apply (erule notE) | |
| 93 | apply (rule Pi_type, assumption) | |
| 94 | apply (blast dest: apply_type) | |
| 95 | done | |
| 96 | ||
| 13535 | 97 | lemma AC17_AC1_aux2: | 
| 12776 | 98 |       "~ (\<exists>f \<in> Pow(x)-{0}->x. x - F(f) = 0)   
 | 
| 99 |        ==> (\<lambda>f \<in> Pow(x)-{0}->x . x - F(f))   
 | |
| 100 |            \<in> (Pow(x) -{0} -> x) -> Pow(x) - {0}"
 | |
| 101 | by (fast intro!: lam_type dest!: Diff_eq_0_iff [THEN iffD1]) | |
| 102 | ||
| 13535 | 103 | lemma AC17_AC1_aux3: | 
| 12776 | 104 |      "[| f`Z \<in> Z; Z \<in> Pow(x)-{0} |] 
 | 
| 105 |       ==> (\<lambda>X \<in> Pow(x)-{0}. {f`X})`Z \<in> Pow(Z)-{0}"
 | |
| 106 | by auto | |
| 107 | ||
| 13535 | 108 | lemma AC17_AC1_aux4: | 
| 12776 | 109 | "\<exists>f \<in> F. f`((\<lambda>f \<in> F. Q(f))`f) \<in> (\<lambda>f \<in> F. Q(f))`f | 
| 110 | ==> \<exists>f \<in> F. f`Q(f) \<in> Q(f)" | |
| 111 | by simp | |
| 112 | ||
| 113 | lemma AC17_AC1: "AC17 ==> AC1" | |
| 114 | apply (unfold AC17_def) | |
| 115 | apply (rule classical) | |
| 116 | apply (erule not_AC1_imp_ex [THEN exE]) | |
| 117 | apply (case_tac | |
| 118 |        "\<exists>f \<in> Pow(x)-{0} -> x. 
 | |
| 119 |         x - (\<Union>a \<in> (LEAST i. HH (\<lambda>X \<in> Pow (x) -{0}. {f`X},x,i) ={x}) . HH (\<lambda>X \<in> Pow (x) -{0}. {f`X},x,a)) = 0")
 | |
| 13535 | 120 | apply (erule AC17_AC1_aux1, assumption) | 
| 121 | apply (drule AC17_AC1_aux2) | |
| 12776 | 122 | apply (erule allE) | 
| 123 | apply (drule bspec, assumption) | |
| 13535 | 124 | apply (drule AC17_AC1_aux4) | 
| 12776 | 125 | apply (erule bexE) | 
| 126 | apply (drule apply_type, assumption) | |
| 127 | apply (simp add: HH_Least_eq_x del: Diff_iff ) | |
| 13535 | 128 | apply (drule AC17_AC1_aux3, assumption) | 
| 12776 | 129 | apply (fast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] | 
| 130 | f_subset_imp_HH_subset elim!: mem_irrefl) | |
| 131 | done | |
| 132 | ||
| 133 | ||
| 134 | (* ********************************************************************** | |
| 135 | AC1 ==> AC2 ==> AC1 | |
| 136 | AC1 ==> AC4 ==> AC3 ==> AC1 | |
| 137 | AC4 ==> AC5 ==> AC4 | |
| 138 | AC1 <-> AC6 | |
| 139 | ************************************************************************* *) | |
| 140 | ||
| 141 | (* ********************************************************************** *) | |
| 142 | (* AC1 ==> AC2 *) | |
| 143 | (* ********************************************************************** *) | |
| 144 | ||
| 13535 | 145 | lemma AC1_AC2_aux1: | 
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 146 |      "[| f:(\<Pi> X \<in> A. X);  B \<in> A;  0\<notin>A |] ==> {f`B} \<subseteq> B Int {f`C. C \<in> A}"
 | 
| 12776 | 147 | by (fast elim!: apply_type) | 
| 1123 | 148 | |
| 13535 | 149 | lemma AC1_AC2_aux2: | 
| 12776 | 150 | "[| pairwise_disjoint(A); B \<in> A; C \<in> A; D \<in> B; D \<in> C |] ==> f`B = f`C" | 
| 151 | by (unfold pairwise_disjoint_def, fast) | |
| 152 | ||
| 153 | lemma AC1_AC2: "AC1 ==> AC2" | |
| 154 | apply (unfold AC1_def AC2_def) | |
| 155 | apply (rule allI) | |
| 156 | apply (rule impI) | |
| 157 | apply (elim asm_rl conjE allE exE impE, assumption) | |
| 158 | apply (intro exI ballI equalityI) | |
| 13535 | 159 | prefer 2 apply (rule AC1_AC2_aux1, assumption+) | 
| 160 | apply (fast elim!: AC1_AC2_aux2 elim: apply_type) | |
| 12776 | 161 | done | 
| 162 | ||
| 163 | ||
| 164 | (* ********************************************************************** *) | |
| 165 | (* AC2 ==> AC1 *) | |
| 166 | (* ********************************************************************** *) | |
| 167 | ||
| 13535 | 168 | lemma AC2_AC1_aux1: "0\<notin>A ==> 0 \<notin> {B*{B}. B \<in> A}"
 | 
| 12776 | 169 | by (fast dest!: sym [THEN Sigma_empty_iff [THEN iffD1]]) | 
| 170 | ||
| 13535 | 171 | lemma AC2_AC1_aux2: "[| X*{X} Int C = {y}; X \<in> A |]   
 | 
| 12776 | 172 |                ==> (THE y. X*{X} Int C = {y}): X*A"
 | 
| 173 | apply (rule subst_elem [of y]) | |
| 174 | apply (blast elim!: equalityE) | |
| 175 | apply (auto simp add: singleton_eq_iff) | |
| 176 | done | |
| 177 | ||
| 13535 | 178 | lemma AC2_AC1_aux3: | 
| 12776 | 179 |      "\<forall>D \<in> {E*{E}. E \<in> A}. \<exists>y. D Int C = {y}   
 | 
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 180 |       ==> (\<lambda>x \<in> A. fst(THE z. (x*{x} Int C = {z}))) \<in> (\<Pi> X \<in> A. X)"
 | 
| 12776 | 181 | apply (rule lam_type) | 
| 182 | apply (drule bspec, blast) | |
| 13535 | 183 | apply (blast intro: AC2_AC1_aux2 fst_type) | 
| 12776 | 184 | done | 
| 185 | ||
| 186 | lemma AC2_AC1: "AC2 ==> AC1" | |
| 187 | apply (unfold AC1_def AC2_def pairwise_disjoint_def) | |
| 188 | apply (intro allI impI) | |
| 189 | apply (elim allE impE) | |
| 13535 | 190 | prefer 2 apply (fast elim!: AC2_AC1_aux3) | 
| 191 | apply (blast intro!: AC2_AC1_aux1) | |
| 12776 | 192 | done | 
| 193 | ||
| 194 | ||
| 195 | (* ********************************************************************** *) | |
| 196 | (* AC1 ==> AC4 *) | |
| 197 | (* ********************************************************************** *) | |
| 198 | ||
| 199 | lemma empty_notin_images: "0 \<notin> {R``{x}. x \<in> domain(R)}"
 | |
| 200 | by blast | |
| 201 | ||
| 202 | lemma AC1_AC4: "AC1 ==> AC4" | |
| 203 | apply (unfold AC1_def AC4_def) | |
| 204 | apply (intro allI impI) | |
| 205 | apply (drule spec, drule mp [OF _ empty_notin_images]) | |
| 206 | apply (best intro!: lam_type elim!: apply_type) | |
| 207 | done | |
| 208 | ||
| 209 | ||
| 210 | (* ********************************************************************** *) | |
| 211 | (* AC4 ==> AC3 *) | |
| 212 | (* ********************************************************************** *) | |
| 213 | ||
| 13535 | 214 | lemma AC4_AC3_aux1: "f \<in> A->B ==> (\<Union>z \<in> A. {z}*f`z) \<subseteq> A*Union(B)"
 | 
| 12776 | 215 | by (fast dest!: apply_type) | 
| 216 | ||
| 13535 | 217 | lemma AC4_AC3_aux2: "domain(\<Union>z \<in> A. {z}*f(z)) = {a \<in> A. f(a)\<noteq>0}"
 | 
| 12776 | 218 | by blast | 
| 219 | ||
| 13535 | 220 | lemma AC4_AC3_aux3: "x \<in> A ==> (\<Union>z \<in> A. {z}*f(z))``{x} = f(x)"
 | 
| 12776 | 221 | by fast | 
| 222 | ||
| 223 | lemma AC4_AC3: "AC4 ==> AC3" | |
| 224 | apply (unfold AC3_def AC4_def) | |
| 225 | apply (intro allI ballI) | |
| 226 | apply (elim allE impE) | |
| 13535 | 227 | apply (erule AC4_AC3_aux1) | 
| 228 | apply (simp add: AC4_AC3_aux2 AC4_AC3_aux3 cong add: Pi_cong) | |
| 12776 | 229 | done | 
| 230 | ||
| 231 | (* ********************************************************************** *) | |
| 232 | (* AC3 ==> AC1 *) | |
| 233 | (* ********************************************************************** *) | |
| 234 | ||
| 235 | lemma AC3_AC1_lemma: | |
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 236 |      "b\<notin>A ==> (\<Pi> x \<in> {a \<in> A. id(A)`a\<noteq>b}. id(A)`x) = (\<Pi> x \<in> A. x)"
 | 
| 12776 | 237 | apply (simp add: id_def cong add: Pi_cong) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12891diff
changeset | 238 | apply (rule_tac b = A in subst_context, fast) | 
| 12776 | 239 | done | 
| 240 | ||
| 241 | lemma AC3_AC1: "AC3 ==> AC1" | |
| 242 | apply (unfold AC1_def AC3_def) | |
| 243 | apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst]) | |
| 244 | done | |
| 245 | ||
| 246 | (* ********************************************************************** *) | |
| 247 | (* AC4 ==> AC5 *) | |
| 248 | (* ********************************************************************** *) | |
| 249 | ||
| 250 | lemma AC4_AC5: "AC4 ==> AC5" | |
| 251 | apply (unfold range_def AC4_def AC5_def) | |
| 252 | apply (intro allI ballI) | |
| 253 | apply (elim allE impE) | |
| 254 | apply (erule fun_is_rel [THEN converse_type]) | |
| 255 | apply (erule exE) | |
| 256 | apply (rename_tac g) | |
| 257 | apply (rule_tac x=g in bexI) | |
| 258 | apply (blast dest: apply_equality range_type) | |
| 259 | apply (blast intro: Pi_type dest: apply_type fun_is_rel) | |
| 260 | done | |
| 261 | ||
| 262 | ||
| 263 | (* ********************************************************************** *) | |
| 264 | (* AC5 ==> AC4, Rubin & Rubin, p. 11 *) | |
| 265 | (* ********************************************************************** *) | |
| 266 | ||
| 13535 | 267 | lemma AC5_AC4_aux1: "R \<subseteq> A*B ==> (\<lambda>x \<in> R. fst(x)) \<in> R -> A" | 
| 12776 | 268 | by (fast intro!: lam_type fst_type) | 
| 269 | ||
| 13535 | 270 | lemma AC5_AC4_aux2: "R \<subseteq> A*B ==> range(\<lambda>x \<in> R. fst(x)) = domain(R)" | 
| 12776 | 271 | by (unfold lam_def, force) | 
| 272 | ||
| 13535 | 273 | lemma AC5_AC4_aux3: "[| \<exists>f \<in> A->C. P(f,domain(f)); A=B |] ==> \<exists>f \<in> B->C. P(f,B)" | 
| 12776 | 274 | apply (erule bexE) | 
| 275 | apply (frule domain_of_fun, fast) | |
| 276 | done | |
| 277 | ||
| 13535 | 278 | lemma AC5_AC4_aux4: "[| R \<subseteq> A*B; g \<in> C->R; \<forall>x \<in> C. (\<lambda>z \<in> R. fst(z))` (g`x) = x |] | 
| 14171 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 skalberg parents: 
13535diff
changeset | 279 |                 ==> (\<lambda>x \<in> C. snd(g`x)): (\<Pi> x \<in> C. R``{x})"
 | 
| 12776 | 280 | apply (rule lam_type) | 
| 281 | apply (force dest: apply_type) | |
| 282 | done | |
| 283 | ||
| 284 | lemma AC5_AC4: "AC5 ==> AC4" | |
| 285 | apply (unfold AC4_def AC5_def, clarify) | |
| 286 | apply (elim allE ballE) | |
| 13535 | 287 | apply (drule AC5_AC4_aux3 [OF _ AC5_AC4_aux2], assumption) | 
| 288 | apply (fast elim!: AC5_AC4_aux4) | |
| 289 | apply (blast intro: AC5_AC4_aux1) | |
| 12776 | 290 | done | 
| 291 | ||
| 292 | ||
| 293 | (* ********************************************************************** *) | |
| 294 | (* AC1 <-> AC6 *) | |
| 295 | (* ********************************************************************** *) | |
| 296 | ||
| 297 | lemma AC1_iff_AC6: "AC1 <-> AC6" | |
| 298 | by (unfold AC1_def AC6_def, blast) | |
| 299 | ||
| 300 | end |