| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 19 Dec 2023 17:26:15 +0100 | |
| changeset 79292 | fb86fa1c6af9 | 
| parent 75711 | 32d45952c12d | 
| child 81974 | f30022be9213 | 
| permissions | -rw-r--r-- | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 1 | (* Title: HOL/Library/Infinite_Set.thy | 
| 20809 | 2 | Author: Stephan Merz | 
| 3 | *) | |
| 4 | ||
| 60500 | 5 | section \<open>Infinite Sets and Related Concepts\<close> | 
| 20809 | 6 | |
| 7 | theory Infinite_Set | |
| 64967 | 8 | imports Main | 
| 20809 | 9 | begin | 
| 10 | ||
| 64967 | 11 | subsection \<open>The set of natural numbers is infinite\<close> | 
| 20809 | 12 | |
| 64967 | 13 | lemma infinite_nat_iff_unbounded_le: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. n \<in> S)" | 
| 14 | for S :: "nat set" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 15 | using frequently_cofinite[of "\<lambda>x. x \<in> S"] | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 16 | by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 17 | |
| 64967 | 18 | lemma infinite_nat_iff_unbounded: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n>m. n \<in> S)" | 
| 19 | for S :: "nat set" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 20 | using frequently_cofinite[of "\<lambda>x. x \<in> S"] | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 21 | by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) | 
| 20809 | 22 | |
| 64967 | 23 | lemma finite_nat_iff_bounded: "finite S \<longleftrightarrow> (\<exists>k. S \<subseteq> {..<k})"
 | 
| 24 | for S :: "nat set" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 25 | using infinite_nat_iff_unbounded_le[of S] by (simp add: subset_eq) (metis not_le) | 
| 20809 | 26 | |
| 64967 | 27 | lemma finite_nat_iff_bounded_le: "finite S \<longleftrightarrow> (\<exists>k. S \<subseteq> {.. k})"
 | 
| 28 | for S :: "nat set" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 29 | using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) | 
| 20809 | 30 | |
| 64967 | 31 | lemma finite_nat_bounded: "finite S \<Longrightarrow> \<exists>k. S \<subseteq> {..<k}"
 | 
| 32 | for S :: "nat set" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 33 | by (simp add: finite_nat_iff_bounded) | 
| 20809 | 34 | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 35 | |
| 60500 | 36 | text \<open> | 
| 20809 | 37 | For a set of natural numbers to be infinite, it is enough to know | 
| 61585 | 38 | that for any number larger than some \<open>k\<close>, there is some larger | 
| 20809 | 39 | number that is an element of the set. | 
| 60500 | 40 | \<close> | 
| 20809 | 41 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 42 | lemma unbounded_k_infinite: "\<forall>m>k. \<exists>n>m. n \<in> S \<Longrightarrow> infinite (S::nat set)" | 
| 64967 | 43 | apply (clarsimp simp add: finite_nat_set_iff_bounded) | 
| 44 | apply (drule_tac x="Suc (max m k)" in spec) | |
| 45 | using less_Suc_eq apply fastforce | |
| 46 | done | |
| 20809 | 47 | |
| 35056 | 48 | lemma nat_not_finite: "finite (UNIV::nat set) \<Longrightarrow> R" | 
| 20809 | 49 | by simp | 
| 50 | ||
| 51 | lemma range_inj_infinite: | |
| 64967 | 52 | fixes f :: "nat \<Rightarrow> 'a" | 
| 53 | assumes "inj f" | |
| 54 | shows "infinite (range f)" | |
| 20809 | 55 | proof | 
| 64967 | 56 | assume "finite (range f)" | 
| 57 | from this assms have "finite (UNIV::nat set)" | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 58 | by (rule finite_imageD) | 
| 20809 | 59 | then show False by simp | 
| 60 | qed | |
| 61 | ||
| 64967 | 62 | |
| 63 | subsection \<open>The set of integers is also infinite\<close> | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 64 | |
| 64967 | 65 | lemma infinite_int_iff_infinite_nat_abs: "infinite S \<longleftrightarrow> infinite ((nat \<circ> abs) ` S)" | 
| 66 | for S :: "int set" | |
| 69661 | 67 | proof (unfold Not_eq_iff, rule iffI) | 
| 68 | assume "finite ((nat \<circ> abs) ` S)" | |
| 69 | then have "finite (nat ` (abs ` S))" | |
| 70 | by (simp add: image_image cong: image_cong) | |
| 71 | moreover have "inj_on nat (abs ` S)" | |
| 66837 | 72 | by (rule inj_onI) auto | 
| 69661 | 73 | ultimately have "finite (abs ` S)" | 
| 74 | by (rule finite_imageD) | |
| 75 | then show "finite S" | |
| 76 | by (rule finite_image_absD) | |
| 77 | qed simp | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 78 | |
| 64967 | 79 | proposition infinite_int_iff_unbounded_le: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n. \<bar>n\<bar> \<ge> m \<and> n \<in> S)" | 
| 80 | for S :: "int set" | |
| 81 | by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) | |
| 82 | (metis abs_ge_zero nat_le_eq_zle le_nat_iff) | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 83 | |
| 64967 | 84 | proposition infinite_int_iff_unbounded: "infinite S \<longleftrightarrow> (\<forall>m. \<exists>n. \<bar>n\<bar> > m \<and> n \<in> S)" | 
| 85 | for S :: "int set" | |
| 86 | by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) | |
| 87 | (metis (full_types) nat_le_iff nat_mono not_le) | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 88 | |
| 64967 | 89 | proposition finite_int_iff_bounded: "finite S \<longleftrightarrow> (\<exists>k. abs ` S \<subseteq> {..<k})"
 | 
| 90 | for S :: "int set" | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 91 | using infinite_int_iff_unbounded_le[of S] by (simp add: subset_eq) (metis not_le) | 
| 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 92 | |
| 64967 | 93 | proposition finite_int_iff_bounded_le: "finite S \<longleftrightarrow> (\<exists>k. abs ` S \<subseteq> {.. k})"
 | 
| 94 | for S :: "int set" | |
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 95 | using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) | 
| 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61585diff
changeset | 96 | |
| 64967 | 97 | |
| 98 | subsection \<open>Infinitely Many and Almost All\<close> | |
| 20809 | 99 | |
| 60500 | 100 | text \<open> | 
| 20809 | 101 | We often need to reason about the existence of infinitely many | 
| 102 | (resp., all but finitely many) objects satisfying some predicate, so | |
| 103 | we introduce corresponding binders and their proof rules. | |
| 60500 | 104 | \<close> | 
| 20809 | 105 | |
| 64967 | 106 | lemma not_INFM [simp]: "\<not> (INFM x. P x) \<longleftrightarrow> (MOST x. \<not> P x)" | 
| 107 | by (rule not_frequently) | |
| 108 | ||
| 109 | lemma not_MOST [simp]: "\<not> (MOST x. P x) \<longleftrightarrow> (INFM x. \<not> P x)" | |
| 110 | by (rule not_eventually) | |
| 34112 | 111 | |
| 112 | lemma INFM_const [simp]: "(INFM x::'a. P) \<longleftrightarrow> P \<and> infinite (UNIV::'a set)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 113 | by (simp add: frequently_const_iff) | 
| 34112 | 114 | |
| 115 | lemma MOST_const [simp]: "(MOST x::'a. P) \<longleftrightarrow> P \<or> finite (UNIV::'a set)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 116 | by (simp add: eventually_const_iff) | 
| 20809 | 117 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 118 | lemma INFM_imp_distrib: "(INFM x. P x \<longrightarrow> Q x) \<longleftrightarrow> ((MOST x. P x) \<longrightarrow> (INFM x. Q x))" | 
| 64967 | 119 | by (rule frequently_imp_iff) | 
| 34112 | 120 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 121 | lemma MOST_imp_iff: "MOST x. P x \<Longrightarrow> (MOST x. P x \<longrightarrow> Q x) \<longleftrightarrow> (MOST x. Q x)" | 
| 61810 | 122 | by (auto intro: eventually_rev_mp eventually_mono) | 
| 34113 | 123 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 124 | lemma INFM_conjI: "INFM x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> INFM x. P x \<and> Q x" | 
| 61810 | 125 | by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) | 
| 34112 | 126 | |
| 64967 | 127 | |
| 60500 | 128 | text \<open>Properties of quantifiers with injective functions.\<close> | 
| 34112 | 129 | |
| 53239 | 130 | lemma INFM_inj: "INFM x. P (f x) \<Longrightarrow> inj f \<Longrightarrow> INFM x. P x" | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 131 |   using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite)
 | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 132 | |
| 53239 | 133 | lemma MOST_inj: "MOST x. P x \<Longrightarrow> inj f \<Longrightarrow> MOST x. P (f x)" | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 134 |   using finite_vimageI[of "{x. \<not> P x}" f] by (auto simp: eventually_cofinite)
 | 
| 34112 | 135 | |
| 64967 | 136 | |
| 60500 | 137 | text \<open>Properties of quantifiers with singletons.\<close> | 
| 34112 | 138 | |
| 139 | lemma not_INFM_eq [simp]: | |
| 140 | "\<not> (INFM x. x = a)" | |
| 141 | "\<not> (INFM x. a = x)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 142 | unfolding frequently_cofinite by simp_all | 
| 34112 | 143 | |
| 144 | lemma MOST_neq [simp]: | |
| 145 | "MOST x. x \<noteq> a" | |
| 146 | "MOST x. a \<noteq> x" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 147 | unfolding eventually_cofinite by simp_all | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 148 | |
| 34112 | 149 | lemma INFM_neq [simp]: | 
| 150 | "(INFM x::'a. x \<noteq> a) \<longleftrightarrow> infinite (UNIV::'a set)" | |
| 151 | "(INFM x::'a. a \<noteq> x) \<longleftrightarrow> infinite (UNIV::'a set)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 152 | unfolding frequently_cofinite by simp_all | 
| 34112 | 153 | |
| 154 | lemma MOST_eq [simp]: | |
| 155 | "(MOST x::'a. x = a) \<longleftrightarrow> finite (UNIV::'a set)" | |
| 156 | "(MOST x::'a. a = x) \<longleftrightarrow> finite (UNIV::'a set)" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 157 | unfolding eventually_cofinite by simp_all | 
| 34112 | 158 | |
| 159 | lemma MOST_eq_imp: | |
| 160 | "MOST x. x = a \<longrightarrow> P x" | |
| 161 | "MOST x. a = x \<longrightarrow> P x" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 162 | unfolding eventually_cofinite by simp_all | 
| 34112 | 163 | |
| 64967 | 164 | |
| 60500 | 165 | text \<open>Properties of quantifiers over the naturals.\<close> | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 166 | |
| 64967 | 167 | lemma MOST_nat: "(\<forall>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<exists>m. \<forall>n>m. P n)" | 
| 168 | for P :: "nat \<Rightarrow> bool" | |
| 68406 | 169 | by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq simp flip: not_le) | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 170 | |
| 64967 | 171 | lemma MOST_nat_le: "(\<forall>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<exists>m. \<forall>n\<ge>m. P n)" | 
| 172 | for P :: "nat \<Rightarrow> bool" | |
| 68406 | 173 | by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq simp flip: not_le) | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 174 | |
| 64967 | 175 | lemma INFM_nat: "(\<exists>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<forall>m. \<exists>n>m. P n)" | 
| 176 | for P :: "nat \<Rightarrow> bool" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 177 | by (simp add: frequently_cofinite infinite_nat_iff_unbounded) | 
| 20809 | 178 | |
| 64967 | 179 | lemma INFM_nat_le: "(\<exists>\<^sub>\<infinity>n. P n) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. P n)" | 
| 180 | for P :: "nat \<Rightarrow> bool" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 181 | by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 182 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 183 | lemma MOST_INFM: "infinite (UNIV::'a set) \<Longrightarrow> MOST x::'a. P x \<Longrightarrow> INFM x::'a. P x" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 184 | by (simp add: eventually_frequently) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 185 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 186 | lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)" | 
| 64697 | 187 | by (simp add: cofinite_eq_sequentially) | 
| 20809 | 188 | |
| 64967 | 189 | lemma MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)" | 
| 190 | and MOST_SucD: "MOST n. P (Suc n) \<Longrightarrow> MOST n. P n" | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 191 | by (simp_all add: MOST_Suc_iff) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 192 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 193 | lemma MOST_ge_nat: "MOST n::nat. m \<le> n" | 
| 66837 | 194 | by (simp add: cofinite_eq_sequentially) | 
| 20809 | 195 | |
| 67408 | 196 | \<comment> \<open>legacy names\<close> | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 197 | lemma Inf_many_def: "Inf_many P \<longleftrightarrow> infinite {x. P x}" by (fact frequently_cofinite)
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 198 | lemma Alm_all_def: "Alm_all P \<longleftrightarrow> \<not> (INFM x. \<not> P x)" by simp | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 199 | lemma INFM_iff_infinite: "(INFM x. P x) \<longleftrightarrow> infinite {x. P x}" by (fact frequently_cofinite)
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 200 | lemma MOST_iff_cofinite: "(MOST x. P x) \<longleftrightarrow> finite {x. \<not> P x}" by (fact eventually_cofinite)
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 201 | lemma INFM_EX: "(\<exists>\<^sub>\<infinity>x. P x) \<Longrightarrow> (\<exists>x. P x)" by (fact frequently_ex) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 202 | lemma ALL_MOST: "\<forall>x. P x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. P x" by (fact always_eventually) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 203 | lemma INFM_mono: "\<exists>\<^sub>\<infinity>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> \<exists>\<^sub>\<infinity>x. Q x" by (fact frequently_elim1) | 
| 61810 | 204 | lemma MOST_mono: "\<forall>\<^sub>\<infinity>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> \<forall>\<^sub>\<infinity>x. Q x" by (fact eventually_mono) | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 205 | lemma INFM_disj_distrib: "(\<exists>\<^sub>\<infinity>x. P x \<or> Q x) \<longleftrightarrow> (\<exists>\<^sub>\<infinity>x. P x) \<or> (\<exists>\<^sub>\<infinity>x. Q x)" by (fact frequently_disj_iff) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 206 | lemma MOST_rev_mp: "\<forall>\<^sub>\<infinity>x. P x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. P x \<longrightarrow> Q x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. Q x" by (fact eventually_rev_mp) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 207 | lemma MOST_conj_distrib: "(\<forall>\<^sub>\<infinity>x. P x \<and> Q x) \<longleftrightarrow> (\<forall>\<^sub>\<infinity>x. P x) \<and> (\<forall>\<^sub>\<infinity>x. Q x)" by (fact eventually_conj_iff) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 208 | lemma MOST_conjI: "MOST x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> MOST x. P x \<and> Q x" by (fact eventually_conj) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 209 | lemma INFM_finite_Bex_distrib: "finite A \<Longrightarrow> (INFM y. \<exists>x\<in>A. P x y) \<longleftrightarrow> (\<exists>x\<in>A. INFM y. P x y)" by (fact frequently_bex_finite_distrib) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 210 | lemma MOST_finite_Ball_distrib: "finite A \<Longrightarrow> (MOST y. \<forall>x\<in>A. P x y) \<longleftrightarrow> (\<forall>x\<in>A. MOST y. P x y)" by (fact eventually_ball_finite_distrib) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 211 | lemma INFM_E: "INFM x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> thesis) \<Longrightarrow> thesis" by (fact frequentlyE) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 212 | lemma MOST_I: "(\<And>x. P x) \<Longrightarrow> MOST x. P x" by (rule eventuallyI) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 213 | lemmas MOST_iff_finiteNeg = MOST_iff_cofinite | 
| 20809 | 214 | |
| 215 | ||
| 64967 | 216 | subsection \<open>Enumeration of an Infinite Set\<close> | 
| 217 | ||
| 218 | text \<open>The set's element type must be wellordered (e.g. the natural numbers).\<close> | |
| 20809 | 219 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 220 | text \<open> | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 221 | Could be generalized to | 
| 69593 | 222 |     \<^prop>\<open>enumerate' S n = (SOME t. t \<in> s \<and> finite {s\<in>S. s < t} \<and> card {s\<in>S. s < t} = n)\<close>.
 | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 223 | \<close> | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 224 | |
| 53239 | 225 | primrec (in wellorder) enumerate :: "'a set \<Rightarrow> nat \<Rightarrow> 'a" | 
| 64967 | 226 | where | 
| 227 | enumerate_0: "enumerate S 0 = (LEAST n. n \<in> S)" | |
| 228 |   | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \<in> S}) n"
 | |
| 20809 | 229 | |
| 53239 | 230 | lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n"
 | 
| 20809 | 231 | by simp | 
| 232 | ||
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 233 | lemma enumerate_in_set: "infinite S \<Longrightarrow> enumerate S n \<in> S" | 
| 64967 | 234 | proof (induct n arbitrary: S) | 
| 235 | case 0 | |
| 236 | then show ?case | |
| 237 | by (fastforce intro: LeastI dest!: infinite_imp_nonempty) | |
| 238 | next | |
| 239 | case (Suc n) | |
| 240 | then show ?case | |
| 241 | by simp (metis DiffE infinite_remove) | |
| 242 | qed | |
| 20809 | 243 | |
| 244 | declare enumerate_0 [simp del] enumerate_Suc [simp del] | |
| 245 | ||
| 246 | lemma enumerate_step: "infinite S \<Longrightarrow> enumerate S n < enumerate S (Suc n)" | |
| 71813 | 247 | proof (induction n arbitrary: S) | 
| 248 | case 0 | |
| 249 | then have "enumerate S 0 \<le> enumerate S (Suc 0)" | |
| 250 | by (simp add: enumerate_0 Least_le enumerate_in_set) | |
| 251 |   moreover have "enumerate (S - {enumerate S 0}) 0 \<in> S - {enumerate S 0}"
 | |
| 252 | by (meson "0.prems" enumerate_in_set infinite_remove) | |
| 253 |   then have "enumerate S 0 \<noteq> enumerate (S - {enumerate S 0}) 0"
 | |
| 254 | by auto | |
| 255 | ultimately show ?case | |
| 256 | by (simp add: enumerate_Suc') | |
| 257 | next | |
| 258 | case (Suc n) | |
| 259 | then show ?case | |
| 260 | by (simp add: enumerate_Suc') | |
| 261 | qed | |
| 20809 | 262 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 263 | lemma enumerate_mono: "m < n \<Longrightarrow> infinite S \<Longrightarrow> enumerate S m < enumerate S n" | 
| 64967 | 264 | by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) | 
| 20809 | 265 | |
| 72095 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 266 | lemma enumerate_mono_iff [simp]: | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 267 | "infinite S \<Longrightarrow> enumerate S m < enumerate S n \<longleftrightarrow> m < n" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 268 | by (metis enumerate_mono less_asym less_linear) | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 269 | |
| 75711 | 270 | lemma enumerate_mono_le_iff [simp]: | 
| 271 | "infinite S \<Longrightarrow> enumerate S m \<le> enumerate S n \<longleftrightarrow> m \<le> n" | |
| 272 | by (meson enumerate_mono_iff not_le) | |
| 273 | ||
| 50134 | 274 | lemma le_enumerate: | 
| 275 | assumes S: "infinite S" | |
| 276 | shows "n \<le> enumerate S n" | |
| 61810 | 277 | using S | 
| 50134 | 278 | proof (induct n) | 
| 53239 | 279 | case 0 | 
| 280 | then show ?case by simp | |
| 281 | next | |
| 50134 | 282 | case (Suc n) | 
| 283 | then have "n \<le> enumerate S n" by simp | |
| 60500 | 284 | also note enumerate_mono[of n "Suc n", OF _ \<open>infinite S\<close>] | 
| 50134 | 285 | finally show ?case by simp | 
| 53239 | 286 | qed | 
| 50134 | 287 | |
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
68406diff
changeset | 288 | lemma infinite_enumerate: | 
| 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
68406diff
changeset | 289 | assumes fS: "infinite S" | 
| 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
68406diff
changeset | 290 | shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (\<forall>n. r n \<in> S)" | 
| 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
68406diff
changeset | 291 | unfolding strict_mono_def | 
| 72095 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 292 | using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by blast | 
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
68406diff
changeset | 293 | |
| 50134 | 294 | lemma enumerate_Suc'': | 
| 295 | fixes S :: "'a::wellorder set" | |
| 53239 | 296 | assumes "infinite S" | 
| 297 | shows "enumerate S (Suc n) = (LEAST s. s \<in> S \<and> enumerate S n < s)" | |
| 298 | using assms | |
| 50134 | 299 | proof (induct n arbitrary: S) | 
| 300 | case 0 | |
| 53239 | 301 | then have "\<forall>s \<in> S. enumerate S 0 \<le> s" | 
| 50134 | 302 | by (auto simp: enumerate.simps intro: Least_le) | 
| 303 | then show ?case | |
| 304 |     unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"]
 | |
| 53239 | 305 | by (intro arg_cong[where f = Least] ext) auto | 
| 50134 | 306 | next | 
| 307 | case (Suc n S) | |
| 308 | show ?case | |
| 60500 | 309 | using enumerate_mono[OF zero_less_Suc \<open>infinite S\<close>, of n] \<open>infinite S\<close> | 
| 50134 | 310 | apply (subst (1 2) enumerate_Suc') | 
| 311 | apply (subst Suc) | |
| 64967 | 312 | apply (use \<open>infinite S\<close> in simp) | 
| 53239 | 313 | apply (intro arg_cong[where f = Least] ext) | 
| 68406 | 314 | apply (auto simp flip: enumerate_Suc') | 
| 53239 | 315 | done | 
| 50134 | 316 | qed | 
| 317 | ||
| 318 | lemma enumerate_Ex: | |
| 64967 | 319 | fixes S :: "nat set" | 
| 320 | assumes S: "infinite S" | |
| 321 | and s: "s \<in> S" | |
| 322 | shows "\<exists>n. enumerate S n = s" | |
| 323 | using s | |
| 50134 | 324 | proof (induct s rule: less_induct) | 
| 325 | case (less s) | |
| 326 | show ?case | |
| 64967 | 327 | proof (cases "\<exists>y\<in>S. y < s") | 
| 328 | case True | |
| 50134 | 329 |     let ?y = "Max {s'\<in>S. s' < s}"
 | 
| 64967 | 330 | from True have y: "\<And>x. ?y < x \<longleftrightarrow> (\<forall>s'\<in>S. s' < s \<longrightarrow> s' < x)" | 
| 53239 | 331 | by (subst Max_less_iff) auto | 
| 332 |     then have y_in: "?y \<in> {s'\<in>S. s' < s}"
 | |
| 333 | by (intro Max_in) auto | |
| 334 | with less.hyps[of ?y] obtain n where "enumerate S n = ?y" | |
| 335 | by auto | |
| 50134 | 336 | with S have "enumerate S (Suc n) = s" | 
| 337 | by (auto simp: y less enumerate_Suc'' intro!: Least_equality) | |
| 64967 | 338 | then show ?thesis by auto | 
| 50134 | 339 | next | 
| 64967 | 340 | case False | 
| 50134 | 341 | then have "\<forall>t\<in>S. s \<le> t" by auto | 
| 60500 | 342 | with \<open>s \<in> S\<close> show ?thesis | 
| 50134 | 343 | by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) | 
| 344 | qed | |
| 345 | qed | |
| 346 | ||
| 71813 | 347 | lemma inj_enumerate: | 
| 348 | fixes S :: "'a::wellorder set" | |
| 349 | assumes S: "infinite S" | |
| 350 | shows "inj (enumerate S)" | |
| 351 | unfolding inj_on_def | |
| 352 | proof clarsimp | |
| 353 | show "\<And>x y. enumerate S x = enumerate S y \<Longrightarrow> x = y" | |
| 354 | by (metis neq_iff enumerate_mono[OF _ \<open>infinite S\<close>]) | |
| 355 | qed | |
| 356 | ||
| 357 | text \<open>To generalise this, we'd need a condition that all initial segments were finite\<close> | |
| 50134 | 358 | lemma bij_enumerate: | 
| 359 | fixes S :: "nat set" | |
| 360 | assumes S: "infinite S" | |
| 361 | shows "bij_betw (enumerate S) UNIV S" | |
| 362 | proof - | |
| 71813 | 363 | have "\<forall>s \<in> S. \<exists>i. enumerate S i = s" | 
| 50134 | 364 | using enumerate_Ex[OF S] by auto | 
| 71813 | 365 | moreover note \<open>infinite S\<close> inj_enumerate | 
| 50134 | 366 | ultimately show ?thesis | 
| 367 | unfolding bij_betw_def by (auto intro: enumerate_in_set) | |
| 368 | qed | |
| 369 | ||
| 72097 | 370 | lemma | 
| 371 | fixes S :: "nat set" | |
| 372 | assumes S: "infinite S" | |
| 373 | shows range_enumerate: "range (enumerate S) = S" | |
| 374 | and strict_mono_enumerate: "strict_mono (enumerate S)" | |
| 375 | by (auto simp add: bij_betw_imp_surj_on bij_enumerate assms strict_mono_def) | |
| 376 | ||
| 64967 | 377 | text \<open>A pair of weird and wonderful lemmas from HOL Light.\<close> | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 378 | lemma finite_transitivity_chain: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 379 | assumes "finite A" | 
| 64967 | 380 | and R: "\<And>x. \<not> R x x" "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" | 
| 381 | and A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y. y \<in> A \<and> R x y" | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 382 |   shows "A = {}"
 | 
| 64967 | 383 | using \<open>finite A\<close> A | 
| 384 | proof (induct A) | |
| 385 | case empty | |
| 386 | then show ?case by simp | |
| 387 | next | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 388 | case (insert a A) | 
| 71840 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 389 | have False | 
| 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 390 | using R(1)[of a] R(2)[of _ a] insert(3,4) by blast | 
| 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 391 | thus ?case .. | 
| 64967 | 392 | qed | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 393 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 394 | corollary Union_maximal_sets: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 395 | assumes "finite \<F>" | 
| 64967 | 396 |   shows "\<Union>{T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} = \<Union>\<F>"
 | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 397 | (is "?lhs = ?rhs") | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 398 | proof | 
| 64967 | 399 | show "?lhs \<subseteq> ?rhs" by force | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 400 | show "?rhs \<subseteq> ?lhs" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 401 | proof (rule Union_subsetI) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 402 | fix S | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 403 | assume "S \<in> \<F>" | 
| 64967 | 404 |     have "{T \<in> \<F>. S \<subseteq> T} = {}"
 | 
| 405 |       if "\<not> (\<exists>y. y \<in> {T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} \<and> S \<subseteq> y)"
 | |
| 71813 | 406 | proof - | 
| 407 | have \<section>: "\<And>x. x \<in> \<F> \<and> S \<subseteq> x \<Longrightarrow> \<exists>y. y \<in> \<F> \<and> S \<subseteq> y \<and> x \<subset> y" | |
| 408 | using that by (blast intro: dual_order.trans psubset_imp_subset) | |
| 409 | show ?thesis | |
| 410 | proof (rule finite_transitivity_chain [of _ "\<lambda>T U. S \<subseteq> T \<and> T \<subset> U"]) | |
| 411 | qed (use assms in \<open>auto intro: \<section>\<close>) | |
| 412 | qed | |
| 64967 | 413 |     with \<open>S \<in> \<F>\<close> show "\<exists>y. y \<in> {T \<in> \<F>. \<forall>U\<in>\<F>. \<not> T \<subset> U} \<and> S \<subseteq> y"
 | 
| 414 | by blast | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 415 | qed | 
| 64967 | 416 | qed | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
61945diff
changeset | 417 | |
| 71827 | 418 | subsection \<open>Properties of @{term enumerate} on finite sets\<close>
 | 
| 419 | ||
| 420 | lemma finite_enumerate_in_set: "\<lbrakk>finite S; n < card S\<rbrakk> \<Longrightarrow> enumerate S n \<in> S" | |
| 421 | proof (induction n arbitrary: S) | |
| 422 | case 0 | |
| 423 | then show ?case | |
| 72302 
d7d90ed4c74e
fixed some remarkably ugly proofs
 paulson <lp15@cam.ac.uk> parents: 
72097diff
changeset | 424 | by (metis all_not_in_conv card.empty enumerate.simps(1) not_less0 wellorder_Least_lemma(1)) | 
| 71827 | 425 | next | 
| 426 | case (Suc n) | |
| 427 | show ?case | |
| 428 |     using Suc.prems Suc.IH [of "S - {LEAST n. n \<in> S}"]
 | |
| 429 | apply (simp add: enumerate.simps) | |
| 430 | by (metis Diff_empty Diff_insert0 Suc_lessD card.remove less_Suc_eq) | |
| 431 | qed | |
| 432 | ||
| 433 | lemma finite_enumerate_step: "\<lbrakk>finite S; Suc n < card S\<rbrakk> \<Longrightarrow> enumerate S n < enumerate S (Suc n)" | |
| 434 | proof (induction n arbitrary: S) | |
| 435 | case 0 | |
| 436 | then have "enumerate S 0 \<le> enumerate S (Suc 0)" | |
| 437 | by (simp add: Least_le enumerate.simps(1) finite_enumerate_in_set) | |
| 438 |   moreover have "enumerate (S - {enumerate S 0}) 0 \<in> S - {enumerate S 0}"
 | |
| 439 | by (metis 0 Suc_lessD Suc_less_eq card_Suc_Diff1 enumerate_in_set finite_enumerate_in_set) | |
| 440 |   then have "enumerate S 0 \<noteq> enumerate (S - {enumerate S 0}) 0"
 | |
| 441 | by auto | |
| 442 | ultimately show ?case | |
| 443 | by (simp add: enumerate_Suc') | |
| 444 | next | |
| 445 | case (Suc n) | |
| 446 | then show ?case | |
| 447 | by (simp add: enumerate_Suc' finite_enumerate_in_set) | |
| 448 | qed | |
| 449 | ||
| 450 | lemma finite_enumerate_mono: "\<lbrakk>m < n; finite S; n < card S\<rbrakk> \<Longrightarrow> enumerate S m < enumerate S n" | |
| 451 | by (induct m n rule: less_Suc_induct) (auto intro: finite_enumerate_step) | |
| 452 | ||
| 72095 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 453 | lemma finite_enumerate_mono_iff [simp]: | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 454 | "\<lbrakk>finite S; m < card S; n < card S\<rbrakk> \<Longrightarrow> enumerate S m < enumerate S n \<longleftrightarrow> m < n" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 455 | by (metis finite_enumerate_mono less_asym less_linear) | 
| 71827 | 456 | |
| 457 | lemma finite_le_enumerate: | |
| 458 | assumes "finite S" "n < card S" | |
| 459 | shows "n \<le> enumerate S n" | |
| 460 | using assms | |
| 461 | proof (induction n) | |
| 462 | case 0 | |
| 463 | then show ?case by simp | |
| 464 | next | |
| 465 | case (Suc n) | |
| 466 | then have "n \<le> enumerate S n" by simp | |
| 467 | also note finite_enumerate_mono[of n "Suc n", OF _ \<open>finite S\<close>] | |
| 468 | finally show ?case | |
| 469 | using Suc.prems(2) Suc_leI by blast | |
| 470 | qed | |
| 471 | ||
| 472 | lemma finite_enumerate: | |
| 473 | assumes fS: "finite S" | |
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
72302diff
changeset | 474 |   shows "\<exists>r::nat\<Rightarrow>nat. strict_mono_on {..<card S} r \<and> (\<forall>n<card S. r n \<in> S)"
 | 
| 71827 | 475 | unfolding strict_mono_def | 
| 476 | using finite_enumerate_in_set[OF fS] finite_enumerate_mono[of _ _ S] fS | |
| 477 | by (metis lessThan_iff strict_mono_on_def) | |
| 478 | ||
| 479 | lemma finite_enumerate_Suc'': | |
| 480 | fixes S :: "'a::wellorder set" | |
| 481 | assumes "finite S" "Suc n < card S" | |
| 482 | shows "enumerate S (Suc n) = (LEAST s. s \<in> S \<and> enumerate S n < s)" | |
| 483 | using assms | |
| 484 | proof (induction n arbitrary: S) | |
| 485 | case 0 | |
| 486 | then have "\<forall>s \<in> S. enumerate S 0 \<le> s" | |
| 487 | by (auto simp: enumerate.simps intro: Least_le) | |
| 488 | then show ?case | |
| 489 |     unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"]
 | |
| 490 | by (metis Diff_iff dual_order.strict_iff_order singletonD singletonI) | |
| 491 | next | |
| 492 | case (Suc n S) | |
| 493 |   then have "Suc n < card (S - {enumerate S 0})"
 | |
| 494 | using Suc.prems(2) finite_enumerate_in_set by force | |
| 495 | then show ?case | |
| 496 | apply (subst (1 2) enumerate_Suc') | |
| 497 | apply (simp add: Suc) | |
| 498 | apply (intro arg_cong[where f = Least] HOL.ext) | |
| 499 | using finite_enumerate_mono[OF zero_less_Suc \<open>finite S\<close>, of n] Suc.prems | |
| 500 | by (auto simp flip: enumerate_Suc') | |
| 501 | qed | |
| 502 | ||
| 503 | lemma finite_enumerate_initial_segment: | |
| 504 | fixes S :: "'a::wellorder set" | |
| 72090 | 505 |   assumes "finite S" and n: "n < card (S \<inter> {..<s})"
 | 
| 71827 | 506 |   shows "enumerate (S \<inter> {..<s}) n = enumerate S n"
 | 
| 507 | using n | |
| 508 | proof (induction n) | |
| 509 | case 0 | |
| 510 | have "(LEAST n. n \<in> S \<and> n < s) = (LEAST n. n \<in> S)" | |
| 511 | proof (rule Least_equality) | |
| 512 | have "\<exists>t. t \<in> S \<and> t < s" | |
| 513 | by (metis "0" card_gt_0_iff disjoint_iff_not_equal lessThan_iff) | |
| 514 | then show "(LEAST n. n \<in> S) \<in> S \<and> (LEAST n. n \<in> S) < s" | |
| 515 | by (meson LeastI Least_le le_less_trans) | |
| 516 | qed (simp add: Least_le) | |
| 517 | then show ?case | |
| 518 | by (auto simp: enumerate_0) | |
| 519 | next | |
| 520 | case (Suc n) | |
| 521 | then have less_card: "Suc n < card S" | |
| 522 | by (meson assms(1) card_mono inf_sup_ord(1) leD le_less_linear order.trans) | |
| 72090 | 523 |   obtain T where T: "T \<in> {s \<in> S. enumerate S n < s}"
 | 
| 71827 | 524 | by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) | 
| 72090 | 525 | have "(LEAST x. x \<in> S \<and> x < s \<and> enumerate S n < x) = (LEAST x. x \<in> S \<and> enumerate S n < x)" | 
| 71827 | 526 | (is "_ = ?r") | 
| 527 | proof (intro Least_equality conjI) | |
| 528 | show "?r \<in> S" | |
| 72090 | 529 | by (metis (mono_tags, lifting) LeastI mem_Collect_eq T) | 
| 530 | have "\<not> s \<le> ?r" | |
| 531 | using not_less_Least [of _ "\<lambda>x. x \<in> S \<and> enumerate S n < x"] Suc assms | |
| 532 | by (metis (mono_tags, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def less_le_trans) | |
| 533 | then show "?r < s" | |
| 534 | by auto | |
| 71827 | 535 | show "enumerate S n < ?r" | 
| 72090 | 536 | by (metis (no_types, lifting) LeastI mem_Collect_eq T) | 
| 71827 | 537 | qed (auto simp: Least_le) | 
| 538 | then show ?case | |
| 539 | using Suc assms by (simp add: finite_enumerate_Suc'' less_card) | |
| 540 | qed | |
| 541 | ||
| 542 | lemma finite_enumerate_Ex: | |
| 543 | fixes S :: "'a::wellorder set" | |
| 544 | assumes S: "finite S" | |
| 545 | and s: "s \<in> S" | |
| 546 | shows "\<exists>n<card S. enumerate S n = s" | |
| 547 | using s S | |
| 548 | proof (induction s arbitrary: S rule: less_induct) | |
| 549 | case (less s) | |
| 550 | show ?case | |
| 551 | proof (cases "\<exists>y\<in>S. y < s") | |
| 552 | case True | |
| 553 |     let ?T = "S \<inter> {..<s}"
 | |
| 554 | have "finite ?T" | |
| 555 | using less.prems(2) by blast | |
| 556 | have TS: "card ?T < card S" | |
| 557 | using less.prems by (blast intro: psubset_card_mono [OF \<open>finite S\<close>]) | |
| 558 | from True have y: "\<And>x. Max ?T < x \<longleftrightarrow> (\<forall>s'\<in>S. s' < s \<longrightarrow> s' < x)" | |
| 559 | by (subst Max_less_iff) (auto simp: \<open>finite ?T\<close>) | |
| 560 |     then have y_in: "Max ?T \<in> {s'\<in>S. s' < s}"
 | |
| 561 | using Max_in \<open>finite ?T\<close> by fastforce | |
| 562 | with less.IH[of "Max ?T" ?T] obtain n where n: "enumerate ?T n = Max ?T" "n < card ?T" | |
| 563 | using \<open>finite ?T\<close> by blast | |
| 564 | then have "Suc n < card S" | |
| 565 | using TS less_trans_Suc by blast | |
| 566 | with S n have "enumerate S (Suc n) = s" | |
| 567 | by (subst finite_enumerate_Suc'') (auto simp: y finite_enumerate_initial_segment less finite_enumerate_Suc'' intro!: Least_equality) | |
| 568 | then show ?thesis | |
| 569 | using \<open>Suc n < card S\<close> by blast | |
| 570 | next | |
| 571 | case False | |
| 572 | then have "\<forall>t\<in>S. s \<le> t" by auto | |
| 573 | moreover have "0 < card S" | |
| 574 | using card_0_eq less.prems by blast | |
| 575 | ultimately show ?thesis | |
| 576 | using \<open>s \<in> S\<close> | |
| 577 | by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) | |
| 578 | qed | |
| 579 | qed | |
| 580 | ||
| 72095 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 581 | lemma finite_enum_subset: | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 582 | assumes "\<And>i. i < card X \<Longrightarrow> enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X \<le> card Y" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 583 | shows "X \<subseteq> Y" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 584 | by (metis assms finite_enumerate_Ex finite_enumerate_in_set less_le_trans subsetI) | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 585 | |
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 586 | lemma finite_enum_ext: | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 587 | assumes "\<And>i. i < card X \<Longrightarrow> enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X = card Y" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 588 | shows "X = Y" | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 589 | by (intro antisym finite_enum_subset) (auto simp: assms) | 
| 
cfb6c22a5636
lemmas about sets and the enumerate operator
 paulson <lp15@cam.ac.uk> parents: 
72090diff
changeset | 590 | |
| 71827 | 591 | lemma finite_bij_enumerate: | 
| 592 | fixes S :: "'a::wellorder set" | |
| 593 | assumes S: "finite S" | |
| 594 |   shows "bij_betw (enumerate S) {..<card S} S"
 | |
| 595 | proof - | |
| 596 | have "\<And>n m. \<lbrakk>n \<noteq> m; n < card S; m < card S\<rbrakk> \<Longrightarrow> enumerate S n \<noteq> enumerate S m" | |
| 597 | using finite_enumerate_mono[OF _ \<open>finite S\<close>] by (auto simp: neq_iff) | |
| 598 |   then have "inj_on (enumerate S) {..<card S}"
 | |
| 599 | by (auto simp: inj_on_def) | |
| 600 | moreover have "\<forall>s \<in> S. \<exists>i<card S. enumerate S i = s" | |
| 601 | using finite_enumerate_Ex[OF S] by auto | |
| 602 | moreover note \<open>finite S\<close> | |
| 603 | ultimately show ?thesis | |
| 604 | unfolding bij_betw_def by (auto intro: finite_enumerate_in_set) | |
| 605 | qed | |
| 606 | ||
| 607 | lemma ex_bij_betw_strict_mono_card: | |
| 608 | fixes M :: "'a::wellorder set" | |
| 609 | assumes "finite M" | |
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
72302diff
changeset | 610 |   obtains h where "bij_betw h {..<card M} M" and "strict_mono_on {..<card M} h"
 | 
| 71827 | 611 | proof | 
| 612 |   show "bij_betw (enumerate M) {..<card M} M"
 | |
| 613 | by (simp add: assms finite_bij_enumerate) | |
| 75607 
3c544d64c218
changed argument order of mono_on and strict_mono_on to uniformize with monotone_on and other predicates
 desharna parents: 
72302diff
changeset | 614 |   show "strict_mono_on {..<card M} (enumerate M)"
 | 
| 71827 | 615 | by (simp add: assms finite_enumerate_mono strict_mono_on_def) | 
| 616 | qed | |
| 617 | ||
| 20809 | 618 | end |