| author | wenzelm | 
| Sun, 17 Dec 2023 21:30:21 +0100 | |
| changeset 79271 | b14b289caaf6 | 
| parent 76219 | cf7db6353322 | 
| permissions | -rw-r--r-- | 
| 12776 | 1 | (* Title: ZF/AC/WO2_AC16.thy | 
| 2 | Author: Krzysztof Grabczewski | |
| 3 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 4 | The proof of WO2 \<Longrightarrow> AC16(k #+ m, k) | 
| 12776 | 5 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 6 | The main part of the proof is the inductive reasoning concerning | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 7 | properties of constructed family T_gamma. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 8 | The proof deals with three cases for ordinals: 0, succ and limit ordinal. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 9 | The first instance is trivial, the third not difficult, but the second | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 10 | is very complicated requiring many lemmas. | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 11 | We also need to prove that at any stage gamma the set | 
| 76219 | 12 | (s - \<Union>(...) - k_gamma) (Rubin & Rubin page 15) | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 13 | contains m distinct elements (in fact is equipollent to s) | 
| 12776 | 14 | *) | 
| 15 | ||
| 16417 | 16 | theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin | 
| 12776 | 17 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 18 | (**** A recursive definition used in the proof of WO2 \<Longrightarrow> AC16 ****) | 
| 12776 | 19 | |
| 24893 | 20 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 21 | recfunAC16 :: "[i,i,i,i] \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 22 | "recfunAC16(f,h,i,a) \<equiv> | 
| 12776 | 23 | transrec2(i, 0, | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 24 | \<lambda>g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r | 
| 76214 | 25 |                     else r \<union> {f`(\<mu> i. h`g \<subseteq> f`i \<and> 
 | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 26 | (\<forall>b<a. (h`b \<subseteq> f`i \<longrightarrow> (\<forall>t \<in> r. \<not> h`b \<subseteq> t))))})" | 
| 12776 | 27 | |
| 28 | (* ********************************************************************** *) | |
| 29 | (* Basic properties of recfunAC16 *) | |
| 30 | (* ********************************************************************** *) | |
| 31 | ||
| 32 | lemma recfunAC16_0: "recfunAC16(f,h,0,a) = 0" | |
| 12820 | 33 | by (simp add: recfunAC16_def) | 
| 12776 | 34 | |
| 35 | lemma recfunAC16_succ: | |
| 36 | "recfunAC16(f,h,succ(i),a) = | |
| 37 | (if (\<exists>y \<in> recfunAC16(f,h,i,a). h ` i \<subseteq> y) then recfunAC16(f,h,i,a) | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 38 | else recfunAC16(f,h,i,a) \<union> | 
| 76214 | 39 |             {f ` (\<mu> j. h ` i \<subseteq> f ` j \<and>   
 | 
| 12776 | 40 | (\<forall>b<a. (h`b \<subseteq> f`j | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 41 | \<longrightarrow> (\<forall>t \<in> recfunAC16(f,h,i,a). \<not> h`b \<subseteq> t))))})" | 
| 12820 | 42 | apply (simp add: recfunAC16_def) | 
| 12776 | 43 | done | 
| 44 | ||
| 45 | lemma recfunAC16_Limit: "Limit(i) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 46 | \<Longrightarrow> recfunAC16(f,h,i,a) = (\<Union>j<i. recfunAC16(f,h,j,a))" | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 47 | by (simp add: recfunAC16_def transrec2_Limit) | 
| 12776 | 48 | |
| 49 | (* ********************************************************************** *) | |
| 50 | (* Monotonicity of transrec2 *) | |
| 51 | (* ********************************************************************** *) | |
| 52 | ||
| 53 | lemma transrec2_mono_lemma [rule_format]: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 54 | "\<lbrakk>\<And>g r. r \<subseteq> B(g,r); Ord(i)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 55 | \<Longrightarrow> j<i \<longrightarrow> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)" | 
| 12820 | 56 | apply (erule trans_induct) | 
| 57 | apply (rule Ord_cases, assumption+, fast) | |
| 12776 | 58 | apply (simp (no_asm_simp)) | 
| 12820 | 59 | apply (blast elim!: leE) | 
| 60 | apply (simp add: transrec2_Limit) | |
| 12776 | 61 | apply (blast intro: OUN_I ltI Ord_in_Ord [THEN le_refl] | 
| 62 | elim!: Limit_has_succ [THEN ltE]) | |
| 63 | done | |
| 64 | ||
| 65 | lemma transrec2_mono: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 66 | "\<lbrakk>\<And>g r. r \<subseteq> B(g,r); j\<le>i\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 67 | \<Longrightarrow> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)" | 
| 12820 | 68 | apply (erule leE) | 
| 12776 | 69 | apply (rule transrec2_mono_lemma) | 
| 12820 | 70 | apply (auto intro: lt_Ord2 ) | 
| 12776 | 71 | done | 
| 72 | ||
| 73 | (* ********************************************************************** *) | |
| 74 | (* Monotonicity of recfunAC16 *) | |
| 75 | (* ********************************************************************** *) | |
| 76 | ||
| 77 | lemma recfunAC16_mono: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 78 | "i\<le>j \<Longrightarrow> recfunAC16(f, g, i, a) \<subseteq> recfunAC16(f, g, j, a)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 79 | unfolding recfunAC16_def | 
| 12820 | 80 | apply (rule transrec2_mono, auto) | 
| 12776 | 81 | done | 
| 82 | ||
| 83 | (* ********************************************************************** *) | |
| 84 | (* case of limit ordinal *) | |
| 85 | (* ********************************************************************** *) | |
| 86 | ||
| 87 | ||
| 88 | lemma lemma3_1: | |
| 76214 | 89 | "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) \<and> f(z)<=Y); | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 90 | \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a; | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 91 | V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 92 | \<Longrightarrow> V = W" | 
| 12776 | 93 | apply (erule asm_rl allE impE)+ | 
| 12820 | 94 | apply (drule subsetD, assumption, blast) | 
| 12776 | 95 | done | 
| 96 | ||
| 97 | ||
| 98 | lemma lemma3: | |
| 76214 | 99 | "\<lbrakk>\<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) \<longrightarrow> (\<exists>! Y. Y \<in> F(y) \<and> f(z)<=Y); | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 100 | \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j); i<x; j<x; z<a; | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 101 | V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 102 | \<Longrightarrow> V = W" | 
| 12776 | 103 | apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+) | 
| 104 | apply (erule lemma3_1 [symmetric], assumption+) | |
| 105 | apply (erule lemma3_1, assumption+) | |
| 106 | done | |
| 107 | ||
| 108 | lemma lemma4: | |
| 76214 | 109 | "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X \<and> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 110 | (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow> | 
| 76214 | 111 | (\<exists>! Y. Y \<in> F(y) \<and> h(x) \<subseteq> Y)); | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 112 | x < a\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 113 | \<Longrightarrow> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) \<longrightarrow> | 
| 76214 | 114 | (\<exists>! Y. Y \<in> F(y) \<and> h(z) \<subseteq> Y)" | 
| 12776 | 115 | apply (intro oallI impI) | 
| 12820 | 116 | apply (drule ospec, assumption, clarify) | 
| 12776 | 117 | apply (blast elim!: oallE ) | 
| 118 | done | |
| 119 | ||
| 120 | lemma lemma5: | |
| 76214 | 121 | "\<lbrakk>\<forall>y<x. F(y) \<subseteq> X \<and> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 122 | (\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) \<longrightarrow> | 
| 76214 | 123 | (\<exists>! Y. Y \<in> F(y) \<and> h(x) \<subseteq> Y)); | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 124 | x < a; Limit(x); \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk> | 
| 76214 | 125 | \<Longrightarrow> (\<Union>x<x. F(x)) \<subseteq> X \<and> | 
| 12776 | 126 | (\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x) | 
| 76214 | 127 | \<longrightarrow> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) \<and> h(xa) \<subseteq> Y))" | 
| 12776 | 128 | apply (rule conjI) | 
| 129 | apply (rule subsetI) | |
| 130 | apply (erule OUN_E) | |
| 12820 | 131 | apply (drule ospec, assumption, fast) | 
| 12776 | 132 | apply (drule lemma4, assumption) | 
| 133 | apply (rule oallI) | |
| 134 | apply (rule impI) | |
| 135 | apply (erule disjE) | |
| 136 | apply (frule ospec, erule Limit_has_succ, assumption) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 137 | apply (drule_tac A = a and x = xa in ospec, assumption) | 
| 12776 | 138 | apply (erule impE, rule le_refl [THEN disjI1], erule lt_Ord) | 
| 139 | apply (blast intro: lemma3 Limit_has_succ) | |
| 140 | apply (blast intro: lemma3) | |
| 141 | done | |
| 142 | ||
| 143 | (* ********************************************************************** *) | |
| 144 | (* case of successor ordinal *) | |
| 145 | (* ********************************************************************** *) | |
| 146 | ||
| 147 | (* | |
| 148 | First quite complicated proof of the fact used in the recursive construction | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 149 | of the family T_gamma (WO2 \<Longrightarrow> AC16(k #+ m, k)) - the fact that at any stage | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 150 | gamma the set (s - \<Union>(...) - k_gamma) is equipollent to s | 
| 76219 | 151 | (Rubin & Rubin page 15). | 
| 12776 | 152 | *) | 
| 153 | ||
| 154 | (* ********************************************************************** *) | |
| 155 | (* dbl_Diff_eqpoll_Card *) | |
| 156 | (* ********************************************************************** *) | |
| 157 | ||
| 158 | ||
| 159 | lemma dbl_Diff_eqpoll_Card: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 160 | "\<lbrakk>A\<approx>a; Card(a); \<not>Finite(a); B\<prec>a; C\<prec>a\<rbrakk> \<Longrightarrow> A - B - C\<approx>a" | 
| 12776 | 161 | by (blast intro: Diff_lesspoll_eqpoll_Card) | 
| 162 | ||
| 163 | (* ********************************************************************** *) | |
| 164 | (* Case of finite ordinals *) | |
| 165 | (* ********************************************************************** *) | |
| 166 | ||
| 167 | ||
| 168 | lemma Finite_lesspoll_infinite_Ord: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 169 | "\<lbrakk>Finite(X); \<not>Finite(a); Ord(a)\<rbrakk> \<Longrightarrow> X\<prec>a" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 170 | unfolding lesspoll_def | 
| 12776 | 171 | apply (rule conjI) | 
| 172 | apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption) | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 173 | unfolding Finite_def | 
| 12776 | 174 | apply (blast intro: leI [THEN le_imp_subset, THEN subset_imp_lepoll] | 
| 175 | ltI eqpoll_imp_lepoll lepoll_trans) | |
| 176 | apply (blast intro: eqpoll_sym [THEN eqpoll_trans]) | |
| 177 | done | |
| 178 | ||
| 179 | lemma Union_lesspoll: | |
| 76214 | 180 | "\<lbrakk>\<forall>x \<in> X. x \<lesssim> n \<and> x \<subseteq> T; well_ord(T, R); X \<lesssim> b; | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 181 | b<a; \<not>Finite(a); Card(a); n \<in> nat\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 182 | \<Longrightarrow> \<Union>(X)\<prec>a" | 
| 12776 | 183 | apply (case_tac "Finite (X)") | 
| 184 | apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord | |
| 185 | lepoll_nat_imp_Finite Finite_Union) | |
| 186 | apply (drule lepoll_imp_ex_le_eqpoll) | |
| 187 | apply (erule lt_Ord) | |
| 188 | apply (elim exE conjE) | |
| 189 | apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) | |
| 190 | apply (erule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], | |
| 191 | THEN exE]) | |
| 192 | apply (frule bij_is_surj [THEN surj_image_eq]) | |
| 193 | apply (drule image_fun [OF bij_is_fun subset_refl]) | |
| 194 | apply (drule sym [THEN trans], assumption) | |
| 195 | apply (blast intro: lt_Ord UN_lepoll lt_Card_imp_lesspoll | |
| 196 | lt_trans1 lesspoll_trans1) | |
| 197 | done | |
| 198 | ||
| 199 | (* ********************************************************************** *) | |
| 200 | (* recfunAC16_lepoll_index *) | |
| 201 | (* ********************************************************************** *) | |
| 202 | ||
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 203 | lemma Un_sing_eq_cons: "A \<union> {a} = cons(a, A)"
 | 
| 12776 | 204 | by fast | 
| 205 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 206 | lemma Un_lepoll_succ: "A \<lesssim> B \<Longrightarrow> A \<union> {a} \<lesssim> succ(B)"
 | 
| 12776 | 207 | apply (simp add: Un_sing_eq_cons succ_def) | 
| 208 | apply (blast elim!: mem_irrefl intro: cons_lepoll_cong) | |
| 209 | done | |
| 210 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 211 | lemma Diff_UN_succ_empty: "Ord(a) \<Longrightarrow> F(a) - (\<Union>b<succ(a). F(b)) = 0" | 
| 12776 | 212 | by (fast intro!: le_refl) | 
| 213 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 214 | lemma Diff_UN_succ_subset: "Ord(a) \<Longrightarrow> F(a) \<union> X - (\<Union>b<succ(a). F(b)) \<subseteq> X" | 
| 12776 | 215 | by blast | 
| 216 | ||
| 217 | lemma recfunAC16_Diff_lepoll_1: | |
| 218 | "Ord(x) | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 219 | \<Longrightarrow> recfunAC16(f, g, x, a) - (\<Union>i<x. recfunAC16(f, g, i, a)) \<lesssim> 1" | 
| 12776 | 220 | apply (erule Ord_cases) | 
| 221 | apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll]) | |
| 222 | (*Limit case*) | |
| 223 | prefer 2 apply (simp add: recfunAC16_Limit Diff_cancel | |
| 224 | empty_subsetI [THEN subset_imp_lepoll]) | |
| 225 | (*succ case*) | |
| 226 | apply (simp add: recfunAC16_succ | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 227 | Diff_UN_succ_empty [of _ "\<lambda>j. recfunAC16(f,g,j,a)"] | 
| 12776 | 228 | empty_subsetI [THEN subset_imp_lepoll]) | 
| 229 | apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll] | |
| 230 | singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans) | |
| 231 | done | |
| 232 | ||
| 233 | lemma in_Least_Diff: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 234 | "\<lbrakk>z \<in> F(x); Ord(x)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 235 | \<Longrightarrow> z \<in> F(\<mu> i. z \<in> F(i)) - (\<Union>j<(\<mu> i. z \<in> F(i)). F(j))" | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 236 | by (fast elim: less_LeastE elim!: LeastI) | 
| 12776 | 237 | |
| 238 | lemma Least_eq_imp_ex: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 239 | "\<lbrakk>(\<mu> i. w \<in> F(i)) = (\<mu> i. z \<in> F(i)); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 240 | w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i))\<rbrakk> | 
| 76214 | 241 | \<Longrightarrow> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) \<and> z \<in> (F(b) - (\<Union>c<b. F(c)))" | 
| 12776 | 242 | apply (elim OUN_E) | 
| 243 | apply (drule in_Least_Diff, erule lt_Ord) | |
| 244 | apply (frule in_Least_Diff, erule lt_Ord) | |
| 245 | apply (rule oexI, force) | |
| 246 | apply (blast intro: lt_Ord Least_le [THEN lt_trans1]) | |
| 247 | done | |
| 248 | ||
| 249 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 250 | lemma two_in_lepoll_1: "\<lbrakk>A \<lesssim> 1; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a=b" | 
| 12776 | 251 | by (fast dest!: lepoll_1_is_sing) | 
| 252 | ||
| 253 | ||
| 254 | lemma UN_lepoll_index: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 255 | "\<lbrakk>\<forall>i<a. F(i)-(\<Union>j<i. F(j)) \<lesssim> 1; Limit(a)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 256 | \<Longrightarrow> (\<Union>x<a. F(x)) \<lesssim> a" | 
| 12776 | 257 | apply (rule lepoll_def [THEN def_imp_iff [THEN iffD2]]) | 
| 61394 | 258 | apply (rule_tac x = "\<lambda>z \<in> (\<Union>x<a. F (x)). \<mu> i. z \<in> F (i) " in exI) | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 259 | unfolding inj_def | 
| 12776 | 260 | apply (rule CollectI) | 
| 261 | apply (rule lam_type) | |
| 262 | apply (erule OUN_E) | |
| 263 | apply (erule Least_in_Ord) | |
| 264 | apply (erule ltD) | |
| 265 | apply (erule lt_Ord2) | |
| 266 | apply (intro ballI) | |
| 267 | apply (simp (no_asm_simp)) | |
| 268 | apply (rule impI) | |
| 269 | apply (drule Least_eq_imp_ex, assumption+) | |
| 270 | apply (fast elim!: two_in_lepoll_1) | |
| 271 | done | |
| 272 | ||
| 273 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 274 | lemma recfunAC16_lepoll_index: "Ord(y) \<Longrightarrow> recfunAC16(f, h, y, a) \<lesssim> y" | 
| 12776 | 275 | apply (erule trans_induct3) | 
| 276 | (*0 case*) | |
| 277 | apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl) | |
| 278 | (*succ case*) | |
| 279 | apply (simp (no_asm_simp) add: recfunAC16_succ) | |
| 280 | apply (blast dest!: succI1 [THEN rev_bspec] | |
| 281 | intro: subset_succI [THEN subset_imp_lepoll] Un_lepoll_succ | |
| 282 | lepoll_trans) | |
| 283 | apply (simp (no_asm_simp) add: recfunAC16_Limit) | |
| 284 | apply (blast intro: lt_Ord [THEN recfunAC16_Diff_lepoll_1] UN_lepoll_index) | |
| 285 | done | |
| 286 | ||
| 1196 | 287 | |
| 12776 | 288 | lemma Union_recfunAC16_lesspoll: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 289 |      "\<lbrakk>recfunAC16(f,g,y,a) \<subseteq> {X \<in> Pow(A). X\<approx>n};   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 290 | A\<approx>a; y<a; \<not>Finite(a); Card(a); n \<in> nat\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 291 | \<Longrightarrow> \<Union>(recfunAC16(f,g,y,a))\<prec>a" | 
| 12776 | 292 | apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE]) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 293 | apply (rule_tac T=A in Union_lesspoll, simp_all) | 
| 12776 | 294 | apply (blast intro!: eqpoll_imp_lepoll) | 
| 295 | apply (blast intro: bij_is_inj Card_is_Ord [THEN well_ord_Memrel] | |
| 296 | well_ord_rvimage) | |
| 297 | apply (erule lt_Ord [THEN recfunAC16_lepoll_index]) | |
| 298 | done | |
| 299 | ||
| 300 | lemma dbl_Diff_eqpoll: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 301 |      "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 302 | Card(a); \<not> Finite(a); A\<approx>a; | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 303 | k \<in> nat; y<a; | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 304 |          h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)})\<rbrakk>   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 305 | \<Longrightarrow> A - \<Union>(recfunAC16(f, h, y, a)) - h`y\<approx>a" | 
| 12776 | 306 | apply (rule dbl_Diff_eqpoll_Card, simp_all) | 
| 307 | apply (simp add: Union_recfunAC16_lesspoll) | |
| 308 | apply (rule Finite_lesspoll_infinite_Ord) | |
| 309 | apply (rule Finite_def [THEN def_imp_iff, THEN iffD2]) | |
| 12820 | 310 | apply (blast dest: ltD bij_is_fun [THEN apply_type], assumption) | 
| 12776 | 311 | apply (blast intro: Card_is_Ord) | 
| 58860 | 312 | done | 
| 12776 | 313 | |
| 314 | (* back to the proof *) | |
| 315 | ||
| 316 | lemmas disj_Un_eqpoll_nat_sum = | |
| 317 | eqpoll_trans [THEN eqpoll_trans, | |
| 58860 | 318 | OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum] | 
| 12776 | 319 | |
| 320 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 321 | lemma Un_in_Collect: "\<lbrakk>x \<in> Pow(A - B - h`i); x\<approx>m; | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 322 |         h \<in> bij(a, {x \<in> Pow(A) . x\<approx>k}); i<a; k \<in> nat; m \<in> nat\<rbrakk>   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 323 |         \<Longrightarrow> h ` i \<union> x \<in> {x \<in> Pow(A) . x\<approx>k #+ m}"
 | 
| 12776 | 324 | by (blast intro: disj_Un_eqpoll_nat_sum | 
| 325 | dest: ltD bij_is_fun [THEN apply_type]) | |
| 326 | ||
| 327 | ||
| 328 | (* ********************************************************************** *) | |
| 329 | (* Lemmas simplifying assumptions *) | |
| 330 | (* ********************************************************************** *) | |
| 331 | ||
| 332 | lemma lemma6: | |
| 76214 | 333 | "\<lbrakk>\<forall>y<succ(j). F(y)<=X \<and> (\<forall>x<a. x<y | P(x,y) \<longrightarrow> Q(x,y)); succ(j)<a\<rbrakk> | 
| 334 | \<Longrightarrow> F(j)<=X \<and> (\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j))" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 335 | by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]) | 
| 12776 | 336 | |
| 337 | ||
| 338 | lemma lemma7: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 339 | "\<lbrakk>\<forall>x<a. x<j | P(x,j) \<longrightarrow> Q(x,j); succ(j)<a\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 340 | \<Longrightarrow> P(j,j) \<longrightarrow> (\<forall>x<a. x\<le>j | P(x,j) \<longrightarrow> Q(x,j))" | 
| 12776 | 341 | by (fast elim!: leE) | 
| 342 | ||
| 343 | (* ********************************************************************** *) | |
| 344 | (* Lemmas needed to prove ex_next_set, which means that for any successor *) | |
| 345 | (* ordinal there is a set satisfying certain properties *) | |
| 346 | (* ********************************************************************** *) | |
| 347 | ||
| 348 | lemma ex_subset_eqpoll: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 349 | "\<lbrakk>A\<approx>a; \<not> Finite(a); Ord(a); m \<in> nat\<rbrakk> \<Longrightarrow> \<exists>X \<in> Pow(A). X\<approx>m" | 
| 12776 | 350 | apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE]) | 
| 351 | apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll]) | |
| 352 | apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat) | |
| 353 | apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) | |
| 354 | apply (fast elim!: eqpoll_sym) | |
| 355 | done | |
| 356 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 357 | lemma subset_Un_disjoint: "\<lbrakk>A \<subseteq> B \<union> C; A \<inter> C = 0\<rbrakk> \<Longrightarrow> A \<subseteq> B" | 
| 12776 | 358 | by blast | 
| 359 | ||
| 360 | ||
| 361 | lemma Int_empty: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 362 | "\<lbrakk>X \<in> Pow(A - \<Union>(B) -C); T \<in> B; F \<subseteq> T\<rbrakk> \<Longrightarrow> F \<inter> X = 0" | 
| 12776 | 363 | by blast | 
| 364 | ||
| 365 | ||
| 366 | (* ********************************************************************** *) | |
| 367 | (* equipollent subset (and finite) is the whole set *) | |
| 368 | (* ********************************************************************** *) | |
| 369 | ||
| 370 | ||
| 371 | lemma subset_imp_eq_lemma: | |
| 76214 | 372 | "m \<in> nat \<Longrightarrow> \<forall>A B. A \<subseteq> B \<and> m \<lesssim> A \<and> B \<lesssim> m \<longrightarrow> A=B" | 
| 12776 | 373 | apply (induct_tac "m") | 
| 374 | apply (fast dest!: lepoll_0_is_0) | |
| 375 | apply (intro allI impI) | |
| 376 | apply (elim conjE) | |
| 377 | apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], assumption) | |
| 378 | apply (frule subsetD [THEN Diff_sing_lepoll], assumption+) | |
| 379 | apply (frule lepoll_Diff_sing) | |
| 380 | apply (erule allE impE)+ | |
| 381 | apply (rule conjI) | |
| 382 | prefer 2 apply fast | |
| 383 | apply fast | |
| 384 | apply (blast elim: equalityE) | |
| 385 | done | |
| 386 | ||
| 387 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 388 | lemma subset_imp_eq: "\<lbrakk>A \<subseteq> B; m \<lesssim> A; B \<lesssim> m; m \<in> nat\<rbrakk> \<Longrightarrow> A=B" | 
| 12776 | 389 | by (blast dest!: subset_imp_eq_lemma) | 
| 390 | ||
| 391 | ||
| 392 | lemma bij_imp_arg_eq: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 393 |      "\<lbrakk>f \<in> bij(a, {Y \<in> X. Y\<approx>succ(k)}); k \<in> nat; f`b \<subseteq> f`y; b<a; y<a\<rbrakk> 
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 394 | \<Longrightarrow> b=y" | 
| 12776 | 395 | apply (drule subset_imp_eq) | 
| 396 | apply (erule_tac [3] nat_succI) | |
| 76217 | 397 | unfolding bij_def inj_def | 
| 12776 | 398 | apply (blast intro: eqpoll_sym eqpoll_imp_lepoll | 
| 399 | dest: ltD apply_type)+ | |
| 400 | done | |
| 401 | ||
| 402 | ||
| 403 | lemma ex_next_set: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 404 |      "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 405 | Card(a); \<not> Finite(a); A\<approx>a; | 
| 12776 | 406 | k \<in> nat; m \<in> nat; y<a; | 
| 407 |          h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 408 | \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk> | 
| 76214 | 409 |       \<Longrightarrow> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X \<and>   
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 410 | (\<forall>b<a. h`b \<subseteq> X \<longrightarrow> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 411 | (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))" | 
| 12776 | 412 | apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], | 
| 413 | assumption+) | |
| 414 | apply (erule Card_is_Ord, assumption) | |
| 415 | apply (frule Un_in_Collect, (erule asm_rl nat_succI)+) | |
| 416 | apply (erule CollectE) | |
| 417 | apply (rule rev_bexI, simp) | |
| 418 | apply (rule conjI, blast) | |
| 419 | apply (intro ballI impI oallI notI) | |
| 420 | apply (drule subset_Un_disjoint, rule Int_empty, assumption+) | |
| 421 | apply (blast dest: bij_imp_arg_eq) | |
| 422 | done | |
| 423 | ||
| 424 | (* ********************************************************************** *) | |
| 425 | (* Lemma ex_next_Ord states that for any successor *) | |
| 426 | (* ordinal there is a number of the set satisfying certain properties *) | |
| 427 | (* ********************************************************************** *) | |
| 428 | ||
| 429 | lemma ex_next_Ord: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 430 |      "\<lbrakk>recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)};   
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 431 | Card(a); \<not> Finite(a); A\<approx>a; | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 432 | k \<in> nat; m \<in> nat; y<a; | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 433 |          h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)});   
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 434 |          f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)});   
 | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 435 | \<not> (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y)\<rbrakk> | 
| 76214 | 436 | \<Longrightarrow> \<exists>c<a. h`y \<subseteq> f`c \<and> | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 437 | (\<forall>b<a. h`b \<subseteq> f`c \<longrightarrow> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 438 | (\<forall>T \<in> recfunAC16(f, h, y, a). \<not> h`b \<subseteq> T))" | 
| 12776 | 439 | apply (drule ex_next_set, assumption+) | 
| 440 | apply (erule bexE) | |
| 15481 | 441 | apply (rule_tac x="converse(f)`X" in oexI) | 
| 442 | apply (simp add: right_inverse_bij) | |
| 12776 | 443 | apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI | 
| 444 | Card_is_Ord) | |
| 445 | done | |
| 446 | ||
| 447 | ||
| 448 | (* ********************************************************************** *) | |
| 449 | (* Lemma simplifying assumptions *) | |
| 450 | (* ********************************************************************** *) | |
| 451 | ||
| 452 | lemma lemma8: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 453 | "\<lbrakk>\<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa)) | 
| 76214 | 454 | \<longrightarrow> (\<exists>! Y. Y \<in> F(j) \<and> P(x, Y)); F(j) \<subseteq> X; | 
| 455 | L \<in> X; P(j, L) \<and> (\<forall>x<a. P(x, L) \<longrightarrow> (\<forall>xa \<in> F(j). \<not>P(x, xa)))\<rbrakk> | |
| 456 |       \<Longrightarrow> F(j) \<union> {L} \<subseteq> X \<and>   
 | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 457 |           (\<forall>x<a. x\<le>j | (\<exists>xa \<in> (F(j) \<union> {L}). P(x, xa)) \<longrightarrow>   
 | 
| 76214 | 458 |                  (\<exists>! Y. Y \<in> (F(j) \<union> {L}) \<and> P(x, Y)))"
 | 
| 12776 | 459 | apply (rule conjI) | 
| 460 | apply (fast intro!: singleton_subsetI) | |
| 461 | apply (rule oallI) | |
| 462 | apply (blast elim!: leE oallE) | |
| 463 | done | |
| 464 | ||
| 465 | (* ********************************************************************** *) | |
| 466 | (* The main part of the proof: inductive proof of the property of T_gamma *) | |
| 467 | (* lemma main_induct *) | |
| 468 | (* ********************************************************************** *) | |
| 469 | ||
| 470 | lemma main_induct: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 471 |      "\<lbrakk>b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)});   
 | 
| 12776 | 472 |          h \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k)});   
 | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 473 | \<not>Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat\<rbrakk> | 
| 76214 | 474 |       \<Longrightarrow> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} \<and>   
 | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
45602diff
changeset | 475 | (\<forall>x<a. x < b | (\<exists>Y \<in> recfunAC16(f, h, b, a). h ` x \<subseteq> Y) \<longrightarrow> | 
| 76214 | 476 | (\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) \<and> h ` x \<subseteq> Y))" | 
| 12776 | 477 | apply (erule lt_induct) | 
| 478 | apply (frule lt_Ord) | |
| 479 | apply (erule Ord_cases) | |
| 480 | (* case 0 *) | |
| 481 | apply (simp add: recfunAC16_0) | |
| 482 | (* case Limit *) | |
| 483 | prefer 2 apply (simp add: recfunAC16_Limit) | |
| 484 | apply (rule lemma5, assumption+) | |
| 485 | apply (blast dest!: recfunAC16_mono) | |
| 486 | (* case succ *) | |
| 487 | apply clarify | |
| 488 | apply (erule lemma6 [THEN conjE], assumption) | |
| 489 | apply (simp (no_asm_simp) split del: split_if add: recfunAC16_succ) | |
| 490 | apply (rule conjI [THEN split_if [THEN iffD2]]) | |
| 491 | apply (simp, erule lemma7, assumption) | |
| 492 | apply (rule impI) | |
| 493 | apply (rule ex_next_Ord [THEN oexE], | |
| 494 | assumption+, rule le_refl [THEN lt_trans], assumption+) | |
| 495 | apply (erule lemma8, assumption) | |
| 496 | apply (rule bij_is_fun [THEN apply_type], assumption) | |
| 497 | apply (erule Least_le [THEN lt_trans2, THEN ltD]) | |
| 498 | apply (erule lt_Ord) | |
| 499 | apply (erule succ_leI) | |
| 500 | apply (erule LeastI) | |
| 501 | apply (erule lt_Ord) | |
| 502 | done | |
| 503 | ||
| 504 | (* ********************************************************************** *) | |
| 505 | (* Lemma to simplify the inductive proof *) | |
| 506 | (* - the desired property is a consequence of the inductive assumption *) | |
| 507 | (* ********************************************************************** *) | |
| 508 | ||
| 509 | lemma lemma_simp_induct: | |
| 76214 | 510 | "\<lbrakk>\<forall>b. b<a \<longrightarrow> F(b) \<subseteq> S \<and> (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y)) | 
| 511 | \<longrightarrow> (\<exists>! Y. Y \<in> F(b) \<and> f`x \<subseteq> Y)); | |
| 12776 | 512 | f \<in> a->f``(a); Limit(a); | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 513 | \<forall>i j. i\<le>j \<longrightarrow> F(i) \<subseteq> F(j)\<rbrakk> | 
| 76214 | 514 | \<Longrightarrow> (\<Union>j<a. F(j)) \<subseteq> S \<and> | 
| 515 | (\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) \<and> x \<subseteq> Y)" | |
| 12776 | 516 | apply (rule conjI) | 
| 517 | apply (rule subsetI) | |
| 518 | apply (erule OUN_E, blast) | |
| 519 | apply (rule ballI) | |
| 520 | apply (erule imageE) | |
| 521 | apply (drule ltI, erule Limit_is_Ord) | |
| 522 | apply (drule Limit_has_succ, assumption) | |
| 12820 | 523 | apply (frule_tac x1="succ(xa)" in spec [THEN mp], assumption) | 
| 12776 | 524 | apply (erule conjE) | 
| 525 | apply (drule ospec) | |
| 526 | (** LEVEL 10 **) | |
| 527 | apply (erule leI [THEN succ_leE]) | |
| 528 | apply (erule impE) | |
| 529 | apply (fast elim!: leI [THEN succ_leE, THEN lt_Ord, THEN le_refl]) | |
| 530 | apply (drule apply_equality, assumption) | |
| 531 | apply (elim conjE ex1E) | |
| 532 | (** LEVEL 15 **) | |
| 533 | apply (rule ex1I, blast) | |
| 534 | apply (elim conjE OUN_E) | |
| 535 | apply (erule_tac i="succ(xa)" and j=aa | |
| 536 | in Ord_linear_le [OF lt_Ord lt_Ord], assumption) | |
| 537 | prefer 2 | |
| 538 | apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) | |
| 539 | (** LEVEL 20 **) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
12820diff
changeset | 540 | apply (drule_tac x1=aa in spec [THEN mp], assumption) | 
| 12776 | 541 | apply (frule succ_leE) | 
| 542 | apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) | |
| 543 | done | |
| 544 | ||
| 545 | (* ********************************************************************** *) | |
| 546 | (* The target theorem *) | |
| 547 | (* ********************************************************************** *) | |
| 548 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
61394diff
changeset | 549 | theorem WO2_AC16: "\<lbrakk>WO2; 0<m; k \<in> nat; m \<in> nat\<rbrakk> \<Longrightarrow> AC16(k #+ m,k)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 550 | unfolding AC16_def | 
| 12776 | 551 | apply (rule allI) | 
| 552 | apply (rule impI) | |
| 553 | apply (frule WO2_infinite_subsets_eqpoll_X, assumption+) | |
| 12820 | 554 | apply (frule_tac n="k #+ m" in WO2_infinite_subsets_eqpoll_X, simp, simp) | 
| 12776 | 555 | apply (frule WO2_imp_ex_Card) | 
| 556 | apply (elim exE conjE) | |
| 557 | apply (drule eqpoll_trans [THEN eqpoll_sym, | |
| 558 | THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], | |
| 559 | assumption) | |
| 560 | apply (drule eqpoll_trans [THEN eqpoll_sym, | |
| 561 | THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], | |
| 562 | assumption+) | |
| 563 | apply (elim exE) | |
| 564 | apply (rename_tac h) | |
| 565 | apply (rule_tac x = "\<Union>j<a. recfunAC16 (h,f,j,a) " in exI) | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 566 | apply (rule_tac P="\<lambda>z. Y \<and> (\<forall>x \<in> z. Z(x))" for Y Z | 
| 12776 | 567 | in bij_is_surj [THEN surj_image_eq, THEN subst], assumption) | 
| 568 | apply (rule lemma_simp_induct) | |
| 569 | apply (blast del: conjI notI | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 570 | intro!: main_induct eqpoll_imp_lepoll [THEN lepoll_infinite] ) | 
| 12776 | 571 | apply (blast intro: bij_is_fun [THEN surj_image, THEN surj_is_fun]) | 
| 572 | apply (erule eqpoll_imp_lepoll [THEN lepoll_infinite, | |
| 573 | THEN infinite_Card_is_InfCard, | |
| 574 | THEN InfCard_is_Limit], | |
| 575 | assumption+) | |
| 576 | apply (blast dest!: recfunAC16_mono) | |
| 577 | done | |
| 578 | ||
| 579 | end |