| author | sultana | 
| Wed, 19 Feb 2014 15:57:02 +0000 | |
| changeset 55586 | c94f1a72d9c5 | 
| parent 54612 | 7e291ae244ea | 
| child 58881 | b9556a055632 | 
| 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 | ||
| 5 | header {* Infinite Sets and Related Concepts *}
 | |
| 6 | ||
| 7 | theory Infinite_Set | |
| 54612 
7e291ae244ea
Backed out changeset: a8ad7f6dd217---bypassing Main breaks theories that use \<inf> or \<sup>
 traytel parents: 
54607diff
changeset | 8 | imports Main | 
| 20809 | 9 | begin | 
| 10 | ||
| 11 | subsection "Infinite Sets" | |
| 12 | ||
| 13 | text {*
 | |
| 54557 
d71c2737ee21
correctly account for dead variables when naming set functions
 blanchet parents: 
53239diff
changeset | 14 | Some elementary facts about infinite sets, mostly by Stephan Merz. | 
| 20809 | 15 | Beware! Because "infinite" merely abbreviates a negation, these | 
| 16 |   lemmas may not work well with @{text "blast"}.
 | |
| 17 | *} | |
| 18 | ||
| 53239 | 19 | abbreviation infinite :: "'a set \<Rightarrow> bool" | 
| 20 | where "infinite S \<equiv> \<not> finite S" | |
| 20809 | 21 | |
| 22 | text {*
 | |
| 23 | Infinite sets are non-empty, and if we remove some elements from an | |
| 24 | infinite set, the result is still infinite. | |
| 25 | *} | |
| 26 | ||
| 53239 | 27 | lemma infinite_imp_nonempty: "infinite S \<Longrightarrow> S \<noteq> {}"
 | 
| 20809 | 28 | by auto | 
| 29 | ||
| 53239 | 30 | lemma infinite_remove: "infinite S \<Longrightarrow> infinite (S - {a})"
 | 
| 20809 | 31 | by simp | 
| 32 | ||
| 33 | lemma Diff_infinite_finite: | |
| 34 | assumes T: "finite T" and S: "infinite S" | |
| 35 | shows "infinite (S - T)" | |
| 36 | using T | |
| 37 | proof induct | |
| 38 | from S | |
| 39 |   show "infinite (S - {})" by auto
 | |
| 40 | next | |
| 41 | fix T x | |
| 42 | assume ih: "infinite (S - T)" | |
| 43 |   have "S - (insert x T) = (S - T) - {x}"
 | |
| 44 | by (rule Diff_insert) | |
| 45 | with ih | |
| 46 | show "infinite (S - (insert x T))" | |
| 47 | by (simp add: infinite_remove) | |
| 48 | qed | |
| 49 | ||
| 50 | lemma Un_infinite: "infinite S \<Longrightarrow> infinite (S \<union> T)" | |
| 51 | by simp | |
| 52 | ||
| 35844 
65258d2c3214
added lemma infinite_Un
 Christian Urban <urbanc@in.tum.de> parents: 
35056diff
changeset | 53 | lemma infinite_Un: "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T" | 
| 
65258d2c3214
added lemma infinite_Un
 Christian Urban <urbanc@in.tum.de> parents: 
35056diff
changeset | 54 | by simp | 
| 
65258d2c3214
added lemma infinite_Un
 Christian Urban <urbanc@in.tum.de> parents: 
35056diff
changeset | 55 | |
| 20809 | 56 | lemma infinite_super: | 
| 57 | assumes T: "S \<subseteq> T" and S: "infinite S" | |
| 58 | shows "infinite T" | |
| 59 | proof | |
| 60 | assume "finite T" | |
| 61 | with T have "finite S" by (simp add: finite_subset) | |
| 62 | with S show False by simp | |
| 63 | qed | |
| 64 | ||
| 65 | text {*
 | |
| 66 | As a concrete example, we prove that the set of natural numbers is | |
| 67 | infinite. | |
| 68 | *} | |
| 69 | ||
| 70 | lemma finite_nat_bounded: | |
| 71 | assumes S: "finite (S::nat set)" | |
| 72 |   shows "\<exists>k. S \<subseteq> {..<k}"  (is "\<exists>k. ?bounded S k")
 | |
| 53239 | 73 | using S | 
| 20809 | 74 | proof induct | 
| 75 |   have "?bounded {} 0" by simp
 | |
| 76 |   then show "\<exists>k. ?bounded {} k" ..
 | |
| 77 | next | |
| 78 | fix S x | |
| 79 | assume "\<exists>k. ?bounded S k" | |
| 80 | then obtain k where k: "?bounded S k" .. | |
| 81 | show "\<exists>k. ?bounded (insert x S) k" | |
| 82 | proof (cases "x < k") | |
| 83 | case True | |
| 84 | with k show ?thesis by auto | |
| 85 | next | |
| 86 | case False | |
| 87 | with k have "?bounded S (Suc x)" by auto | |
| 88 | then show ?thesis by auto | |
| 89 | qed | |
| 90 | qed | |
| 91 | ||
| 92 | lemma finite_nat_iff_bounded: | |
| 53239 | 93 |   "finite (S::nat set) \<longleftrightarrow> (\<exists>k. S \<subseteq> {..<k})"  (is "?lhs \<longleftrightarrow> ?rhs")
 | 
| 20809 | 94 | proof | 
| 95 | assume ?lhs | |
| 96 | then show ?rhs by (rule finite_nat_bounded) | |
| 97 | next | |
| 98 | assume ?rhs | |
| 99 |   then obtain k where "S \<subseteq> {..<k}" ..
 | |
| 100 | then show "finite S" | |
| 101 | by (rule finite_subset) simp | |
| 102 | qed | |
| 103 | ||
| 104 | lemma finite_nat_iff_bounded_le: | |
| 53239 | 105 |   "finite (S::nat set) \<longleftrightarrow> (\<exists>k. S \<subseteq> {..k})"  (is "?lhs \<longleftrightarrow> ?rhs")
 | 
| 20809 | 106 | proof | 
| 107 | assume ?lhs | |
| 108 |   then obtain k where "S \<subseteq> {..<k}"
 | |
| 109 | by (blast dest: finite_nat_bounded) | |
| 110 |   then have "S \<subseteq> {..k}" by auto
 | |
| 111 | then show ?rhs .. | |
| 112 | next | |
| 113 | assume ?rhs | |
| 114 |   then obtain k where "S \<subseteq> {..k}" ..
 | |
| 115 | then show "finite S" | |
| 116 | by (rule finite_subset) simp | |
| 117 | qed | |
| 118 | ||
| 119 | lemma infinite_nat_iff_unbounded: | |
| 53239 | 120 | "infinite (S::nat set) \<longleftrightarrow> (\<forall>m. \<exists>n. m < n \<and> n \<in> S)" | 
| 121 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 20809 | 122 | proof | 
| 123 | assume ?lhs | |
| 124 | show ?rhs | |
| 125 | proof (rule ccontr) | |
| 126 | assume "\<not> ?rhs" | |
| 53239 | 127 | then obtain m where m: "\<forall>n. m < n \<longrightarrow> n \<notin> S" by blast | 
| 20809 | 128 |     then have "S \<subseteq> {..m}"
 | 
| 129 | by (auto simp add: sym [OF linorder_not_less]) | |
| 130 | with `?lhs` show False | |
| 131 | by (simp add: finite_nat_iff_bounded_le) | |
| 132 | qed | |
| 133 | next | |
| 134 | assume ?rhs | |
| 135 | show ?lhs | |
| 136 | proof | |
| 137 | assume "finite S" | |
| 138 |     then obtain m where "S \<subseteq> {..m}"
 | |
| 139 | by (auto simp add: finite_nat_iff_bounded_le) | |
| 53239 | 140 | then have "\<forall>n. m < n \<longrightarrow> n \<notin> S" by auto | 
| 20809 | 141 | with `?rhs` show False by blast | 
| 142 | qed | |
| 143 | qed | |
| 144 | ||
| 145 | lemma infinite_nat_iff_unbounded_le: | |
| 53239 | 146 | "infinite (S::nat set) \<longleftrightarrow> (\<forall>m. \<exists>n. m \<le> n \<and> n \<in> S)" | 
| 147 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 20809 | 148 | proof | 
| 149 | assume ?lhs | |
| 150 | show ?rhs | |
| 151 | proof | |
| 152 | fix m | |
| 53239 | 153 | from `?lhs` obtain n where "m < n \<and> n \<in> S" | 
| 20809 | 154 | by (auto simp add: infinite_nat_iff_unbounded) | 
| 53239 | 155 | then have "m \<le> n \<and> n \<in> S" by simp | 
| 20809 | 156 | then show "\<exists>n. m \<le> n \<and> n \<in> S" .. | 
| 157 | qed | |
| 158 | next | |
| 159 | assume ?rhs | |
| 160 | show ?lhs | |
| 161 | proof (auto simp add: infinite_nat_iff_unbounded) | |
| 162 | fix m | |
| 53239 | 163 | from `?rhs` obtain n where "Suc m \<le> n \<and> n \<in> S" | 
| 20809 | 164 | by blast | 
| 53239 | 165 | then have "m < n \<and> n \<in> S" by simp | 
| 20809 | 166 | then show "\<exists>n. m < n \<and> n \<in> S" .. | 
| 167 | qed | |
| 168 | qed | |
| 169 | ||
| 170 | text {*
 | |
| 171 | For a set of natural numbers to be infinite, it is enough to know | |
| 172 |   that for any number larger than some @{text k}, there is some larger
 | |
| 173 | number that is an element of the set. | |
| 174 | *} | |
| 175 | ||
| 176 | lemma unbounded_k_infinite: | |
| 53239 | 177 | assumes k: "\<forall>m. k < m \<longrightarrow> (\<exists>n. m < n \<and> n \<in> S)" | 
| 20809 | 178 | shows "infinite (S::nat set)" | 
| 179 | proof - | |
| 180 |   {
 | |
| 53239 | 181 | fix m have "\<exists>n. m < n \<and> n \<in> S" | 
| 182 | proof (cases "k < m") | |
| 20809 | 183 | case True | 
| 184 | with k show ?thesis by blast | |
| 185 | next | |
| 186 | case False | |
| 53239 | 187 | from k obtain n where "Suc k < n \<and> n \<in> S" by auto | 
| 188 | with False have "m < n \<and> n \<in> S" by auto | |
| 20809 | 189 | then show ?thesis .. | 
| 190 | qed | |
| 191 | } | |
| 192 | then show ?thesis | |
| 193 | by (auto simp add: infinite_nat_iff_unbounded) | |
| 194 | qed | |
| 195 | ||
| 35056 | 196 | lemma nat_not_finite: "finite (UNIV::nat set) \<Longrightarrow> R" | 
| 20809 | 197 | by simp | 
| 198 | ||
| 199 | lemma range_inj_infinite: | |
| 200 | "inj (f::nat \<Rightarrow> 'a) \<Longrightarrow> infinite (range f)" | |
| 201 | proof | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 202 | assume "finite (range f)" and "inj f" | 
| 20809 | 203 | then 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 | 204 | by (rule finite_imageD) | 
| 20809 | 205 | then show False by simp | 
| 206 | qed | |
| 207 | ||
| 208 | text {*
 | |
| 209 | For any function with infinite domain and finite range there is some | |
| 210 | element that is the image of infinitely many domain elements. In | |
| 211 | particular, any infinite sequence of elements from a finite set | |
| 212 | contains some element that occurs infinitely often. | |
| 213 | *} | |
| 214 | ||
| 215 | lemma inf_img_fin_dom: | |
| 216 | assumes img: "finite (f`A)" and dom: "infinite A" | |
| 217 |   shows "\<exists>y \<in> f`A. infinite (f -` {y})"
 | |
| 218 | proof (rule ccontr) | |
| 219 | assume "\<not> ?thesis" | |
| 40786 
0a54cfc9add3
gave more standard finite set rules simp and intro attribute
 nipkow parents: 
35844diff
changeset | 220 |   with img have "finite (UN y:f`A. f -` {y})" by blast
 | 
| 20809 | 221 |   moreover have "A \<subseteq> (UN y:f`A. f -` {y})" by auto
 | 
| 222 | moreover note dom | |
| 223 | ultimately show False by (simp add: infinite_super) | |
| 224 | qed | |
| 225 | ||
| 226 | lemma inf_img_fin_domE: | |
| 227 | assumes "finite (f`A)" and "infinite A" | |
| 228 |   obtains y where "y \<in> f`A" and "infinite (f -` {y})"
 | |
| 23394 | 229 | using assms by (blast dest: inf_img_fin_dom) | 
| 20809 | 230 | |
| 231 | ||
| 232 | subsection "Infinitely Many and Almost All" | |
| 233 | ||
| 234 | text {*
 | |
| 235 | We often need to reason about the existence of infinitely many | |
| 236 | (resp., all but finitely many) objects satisfying some predicate, so | |
| 237 | we introduce corresponding binders and their proof rules. | |
| 238 | *} | |
| 239 | ||
| 53239 | 240 | definition Inf_many :: "('a \<Rightarrow> bool) \<Rightarrow> bool"  (binder "INFM " 10)
 | 
| 241 |   where "Inf_many P \<longleftrightarrow> infinite {x. P x}"
 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 242 | |
| 53239 | 243 | definition Alm_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool"  (binder "MOST " 10)
 | 
| 244 | where "Alm_all P \<longleftrightarrow> \<not> (INFM x. \<not> P x)" | |
| 20809 | 245 | |
| 21210 | 246 | notation (xsymbols) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 247 | Inf_many (binder "\<exists>\<^sub>\<infinity>" 10) and | 
| 20809 | 248 | Alm_all (binder "\<forall>\<^sub>\<infinity>" 10) | 
| 249 | ||
| 21210 | 250 | notation (HTML output) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21256diff
changeset | 251 | Inf_many (binder "\<exists>\<^sub>\<infinity>" 10) and | 
| 20809 | 252 | Alm_all (binder "\<forall>\<^sub>\<infinity>" 10) | 
| 253 | ||
| 34112 | 254 | lemma INFM_iff_infinite: "(INFM x. P x) \<longleftrightarrow> infinite {x. P x}"
 | 
| 255 | unfolding Inf_many_def .. | |
| 256 | ||
| 257 | lemma MOST_iff_cofinite: "(MOST x. P x) \<longleftrightarrow> finite {x. \<not> P x}"
 | |
| 258 | unfolding Alm_all_def Inf_many_def by simp | |
| 259 | ||
| 260 | (* legacy name *) | |
| 261 | lemmas MOST_iff_finiteNeg = MOST_iff_cofinite | |
| 262 | ||
| 263 | lemma not_INFM [simp]: "\<not> (INFM x. P x) \<longleftrightarrow> (MOST x. \<not> P x)" | |
| 264 | unfolding Alm_all_def not_not .. | |
| 20809 | 265 | |
| 34112 | 266 | lemma not_MOST [simp]: "\<not> (MOST x. P x) \<longleftrightarrow> (INFM x. \<not> P x)" | 
| 267 | unfolding Alm_all_def not_not .. | |
| 268 | ||
| 269 | lemma INFM_const [simp]: "(INFM x::'a. P) \<longleftrightarrow> P \<and> infinite (UNIV::'a set)" | |
| 270 | unfolding Inf_many_def by simp | |
| 271 | ||
| 272 | lemma MOST_const [simp]: "(MOST x::'a. P) \<longleftrightarrow> P \<or> finite (UNIV::'a set)" | |
| 273 | unfolding Alm_all_def by simp | |
| 274 | ||
| 275 | lemma INFM_EX: "(\<exists>\<^sub>\<infinity>x. P x) \<Longrightarrow> (\<exists>x. P x)" | |
| 53239 | 276 | apply (erule contrapos_pp) | 
| 277 | apply simp | |
| 278 | done | |
| 20809 | 279 | |
| 280 | lemma ALL_MOST: "\<forall>x. P x \<Longrightarrow> \<forall>\<^sub>\<infinity>x. P x" | |
| 34112 | 281 | by simp | 
| 282 | ||
| 53239 | 283 | lemma INFM_E: | 
| 284 | assumes "INFM x. P x" | |
| 285 | obtains x where "P x" | |
| 34112 | 286 | using INFM_EX [OF assms] by (rule exE) | 
| 287 | ||
| 53239 | 288 | lemma MOST_I: | 
| 289 | assumes "\<And>x. P x" | |
| 290 | shows "MOST x. P x" | |
| 34112 | 291 | using assms by simp | 
| 20809 | 292 | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 293 | lemma INFM_mono: | 
| 20809 | 294 | assumes inf: "\<exists>\<^sub>\<infinity>x. P x" and q: "\<And>x. P x \<Longrightarrow> Q x" | 
| 295 | shows "\<exists>\<^sub>\<infinity>x. Q x" | |
| 296 | proof - | |
| 297 |   from inf have "infinite {x. P x}" unfolding Inf_many_def .
 | |
| 298 |   moreover from q have "{x. P x} \<subseteq> {x. Q x}" by auto
 | |
| 299 | ultimately show ?thesis | |
| 300 | by (simp add: Inf_many_def infinite_super) | |
| 301 | qed | |
| 302 | ||
| 303 | lemma MOST_mono: "\<forall>\<^sub>\<infinity>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> \<forall>\<^sub>\<infinity>x. Q x" | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 304 | unfolding Alm_all_def by (blast intro: INFM_mono) | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 305 | |
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 306 | lemma INFM_disj_distrib: | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 307 | "(\<exists>\<^sub>\<infinity>x. P x \<or> Q x) \<longleftrightarrow> (\<exists>\<^sub>\<infinity>x. P x) \<or> (\<exists>\<^sub>\<infinity>x. Q x)" | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 308 | unfolding Inf_many_def by (simp add: Collect_disj_eq) | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 309 | |
| 34112 | 310 | lemma INFM_imp_distrib: | 
| 311 | "(INFM x. P x \<longrightarrow> Q x) \<longleftrightarrow> ((MOST x. P x) \<longrightarrow> (INFM x. Q x))" | |
| 312 | by (simp only: imp_conv_disj INFM_disj_distrib not_MOST) | |
| 313 | ||
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 314 | lemma MOST_conj_distrib: | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 315 | "(\<forall>\<^sub>\<infinity>x. P x \<and> Q x) \<longleftrightarrow> (\<forall>\<^sub>\<infinity>x. P x) \<and> (\<forall>\<^sub>\<infinity>x. Q x)" | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 316 | unfolding Alm_all_def by (simp add: INFM_disj_distrib del: disj_not1) | 
| 20809 | 317 | |
| 34112 | 318 | lemma MOST_conjI: | 
| 319 | "MOST x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> MOST x. P x \<and> Q x" | |
| 320 | by (simp add: MOST_conj_distrib) | |
| 321 | ||
| 34113 | 322 | lemma INFM_conjI: | 
| 323 | "INFM x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> INFM x. P x \<and> Q x" | |
| 324 | unfolding MOST_iff_cofinite INFM_iff_infinite | |
| 325 | apply (drule (1) Diff_infinite_finite) | |
| 326 | apply (simp add: Collect_conj_eq Collect_neg_eq) | |
| 327 | done | |
| 328 | ||
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 329 | lemma MOST_rev_mp: | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 330 | assumes "\<forall>\<^sub>\<infinity>x. P x" and "\<forall>\<^sub>\<infinity>x. P x \<longrightarrow> Q x" | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 331 | shows "\<forall>\<^sub>\<infinity>x. Q x" | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 332 | proof - | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 333 | have "\<forall>\<^sub>\<infinity>x. P x \<and> (P x \<longrightarrow> Q x)" | 
| 34112 | 334 | using assms by (rule MOST_conjI) | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 335 | thus ?thesis by (rule MOST_mono) simp | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 336 | qed | 
| 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 337 | |
| 34112 | 338 | lemma MOST_imp_iff: | 
| 339 | assumes "MOST x. P x" | |
| 340 | shows "(MOST x. P x \<longrightarrow> Q x) \<longleftrightarrow> (MOST x. Q x)" | |
| 341 | proof | |
| 342 | assume "MOST x. P x \<longrightarrow> Q x" | |
| 343 | with assms show "MOST x. Q x" by (rule MOST_rev_mp) | |
| 344 | next | |
| 345 | assume "MOST x. Q x" | |
| 346 | then show "MOST x. P x \<longrightarrow> Q x" by (rule MOST_mono) simp | |
| 347 | qed | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 348 | |
| 34112 | 349 | lemma INFM_MOST_simps [simp]: | 
| 350 | "\<And>P Q. (INFM x. P x \<and> Q) \<longleftrightarrow> (INFM x. P x) \<and> Q" | |
| 351 | "\<And>P Q. (INFM x. P \<and> Q x) \<longleftrightarrow> P \<and> (INFM x. Q x)" | |
| 352 | "\<And>P Q. (MOST x. P x \<or> Q) \<longleftrightarrow> (MOST x. P x) \<or> Q" | |
| 353 | "\<And>P Q. (MOST x. P \<or> Q x) \<longleftrightarrow> P \<or> (MOST x. Q x)" | |
| 354 | "\<And>P Q. (MOST x. P x \<longrightarrow> Q) \<longleftrightarrow> ((INFM x. P x) \<longrightarrow> Q)" | |
| 355 | "\<And>P Q. (MOST x. P \<longrightarrow> Q x) \<longleftrightarrow> (P \<longrightarrow> (MOST x. Q x))" | |
| 356 | unfolding Alm_all_def Inf_many_def | |
| 357 | by (simp_all add: Collect_conj_eq) | |
| 358 | ||
| 359 | text {* Properties of quantifiers with injective functions. *}
 | |
| 360 | ||
| 53239 | 361 | lemma INFM_inj: "INFM x. P (f x) \<Longrightarrow> inj f \<Longrightarrow> INFM x. P x" | 
| 34112 | 362 | unfolding INFM_iff_infinite | 
| 53239 | 363 | apply clarify | 
| 364 | apply (drule (1) finite_vimageI) | |
| 365 | apply simp | |
| 366 | done | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 367 | |
| 53239 | 368 | lemma MOST_inj: "MOST x. P x \<Longrightarrow> inj f \<Longrightarrow> MOST x. P (f x)" | 
| 34112 | 369 | unfolding MOST_iff_cofinite | 
| 53239 | 370 | apply (drule (1) finite_vimageI) | 
| 371 | apply simp | |
| 372 | done | |
| 34112 | 373 | |
| 374 | text {* Properties of quantifiers with singletons. *}
 | |
| 375 | ||
| 376 | lemma not_INFM_eq [simp]: | |
| 377 | "\<not> (INFM x. x = a)" | |
| 378 | "\<not> (INFM x. a = x)" | |
| 379 | unfolding INFM_iff_infinite by simp_all | |
| 380 | ||
| 381 | lemma MOST_neq [simp]: | |
| 382 | "MOST x. x \<noteq> a" | |
| 383 | "MOST x. a \<noteq> x" | |
| 384 | unfolding MOST_iff_cofinite by simp_all | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 385 | |
| 34112 | 386 | lemma INFM_neq [simp]: | 
| 387 | "(INFM x::'a. x \<noteq> a) \<longleftrightarrow> infinite (UNIV::'a set)" | |
| 388 | "(INFM x::'a. a \<noteq> x) \<longleftrightarrow> infinite (UNIV::'a set)" | |
| 389 | unfolding INFM_iff_infinite by simp_all | |
| 390 | ||
| 391 | lemma MOST_eq [simp]: | |
| 392 | "(MOST x::'a. x = a) \<longleftrightarrow> finite (UNIV::'a set)" | |
| 393 | "(MOST x::'a. a = x) \<longleftrightarrow> finite (UNIV::'a set)" | |
| 394 | unfolding MOST_iff_cofinite by simp_all | |
| 395 | ||
| 396 | lemma MOST_eq_imp: | |
| 397 | "MOST x. x = a \<longrightarrow> P x" | |
| 398 | "MOST x. a = x \<longrightarrow> P x" | |
| 399 | unfolding MOST_iff_cofinite by simp_all | |
| 400 | ||
| 401 | text {* Properties of quantifiers over the naturals. *}
 | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 402 | |
| 53239 | 403 | lemma INFM_nat: "(\<exists>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<forall>m. \<exists>n. m < n \<and> P n)" | 
| 20809 | 404 | by (simp add: Inf_many_def infinite_nat_iff_unbounded) | 
| 405 | ||
| 53239 | 406 | lemma INFM_nat_le: "(\<exists>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<forall>m. \<exists>n. m \<le> n \<and> P n)" | 
| 20809 | 407 | by (simp add: Inf_many_def infinite_nat_iff_unbounded_le) | 
| 408 | ||
| 53239 | 409 | lemma MOST_nat: "(\<forall>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<exists>m. \<forall>n. m < n \<longrightarrow> P n)" | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 410 | by (simp add: Alm_all_def INFM_nat) | 
| 20809 | 411 | |
| 53239 | 412 | lemma MOST_nat_le: "(\<forall>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<exists>m. \<forall>n. m \<le> n \<longrightarrow> P n)" | 
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 413 | by (simp add: Alm_all_def INFM_nat_le) | 
| 20809 | 414 | |
| 415 | ||
| 416 | subsection "Enumeration of an Infinite Set" | |
| 417 | ||
| 418 | text {*
 | |
| 419 | The set's element type must be wellordered (e.g. the natural numbers). | |
| 420 | *} | |
| 421 | ||
| 53239 | 422 | primrec (in wellorder) enumerate :: "'a set \<Rightarrow> nat \<Rightarrow> 'a" | 
| 423 | where | |
| 424 | enumerate_0: "enumerate S 0 = (LEAST n. n \<in> S)" | |
| 425 | | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \<in> S}) n"
 | |
| 20809 | 426 | |
| 53239 | 427 | lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n"
 | 
| 20809 | 428 | by simp | 
| 429 | ||
| 430 | lemma enumerate_in_set: "infinite S \<Longrightarrow> enumerate S n : S" | |
| 53239 | 431 | apply (induct n arbitrary: S) | 
| 432 | apply (fastforce intro: LeastI dest!: infinite_imp_nonempty) | |
| 433 | apply simp | |
| 434 | apply (metis DiffE infinite_remove) | |
| 435 | done | |
| 20809 | 436 | |
| 437 | declare enumerate_0 [simp del] enumerate_Suc [simp del] | |
| 438 | ||
| 439 | lemma enumerate_step: "infinite S \<Longrightarrow> enumerate S n < enumerate S (Suc n)" | |
| 440 | apply (induct n arbitrary: S) | |
| 441 | apply (rule order_le_neq_trans) | |
| 442 | apply (simp add: enumerate_0 Least_le enumerate_in_set) | |
| 443 | apply (simp only: enumerate_Suc') | |
| 444 |    apply (subgoal_tac "enumerate (S - {enumerate S 0}) 0 : S - {enumerate S 0}")
 | |
| 445 | apply (blast intro: sym) | |
| 446 | apply (simp add: enumerate_in_set del: Diff_iff) | |
| 447 | apply (simp add: enumerate_Suc') | |
| 448 | done | |
| 449 | ||
| 450 | lemma enumerate_mono: "m<n \<Longrightarrow> infinite S \<Longrightarrow> enumerate S m < enumerate S n" | |
| 451 | apply (erule less_Suc_induct) | |
| 452 | apply (auto intro: enumerate_step) | |
| 453 | done | |
| 454 | ||
| 455 | ||
| 50134 | 456 | lemma le_enumerate: | 
| 457 | assumes S: "infinite S" | |
| 458 | shows "n \<le> enumerate S n" | |
| 459 | using S | |
| 460 | proof (induct n) | |
| 53239 | 461 | case 0 | 
| 462 | then show ?case by simp | |
| 463 | next | |
| 50134 | 464 | case (Suc n) | 
| 465 | then have "n \<le> enumerate S n" by simp | |
| 466 | also note enumerate_mono[of n "Suc n", OF _ `infinite S`] | |
| 467 | finally show ?case by simp | |
| 53239 | 468 | qed | 
| 50134 | 469 | |
| 470 | lemma enumerate_Suc'': | |
| 471 | fixes S :: "'a::wellorder set" | |
| 53239 | 472 | assumes "infinite S" | 
| 473 | shows "enumerate S (Suc n) = (LEAST s. s \<in> S \<and> enumerate S n < s)" | |
| 474 | using assms | |
| 50134 | 475 | proof (induct n arbitrary: S) | 
| 476 | case 0 | |
| 53239 | 477 | then have "\<forall>s \<in> S. enumerate S 0 \<le> s" | 
| 50134 | 478 | by (auto simp: enumerate.simps intro: Least_le) | 
| 479 | then show ?case | |
| 480 |     unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"]
 | |
| 53239 | 481 | by (intro arg_cong[where f = Least] ext) auto | 
| 50134 | 482 | next | 
| 483 | case (Suc n S) | |
| 484 | show ?case | |
| 485 | using enumerate_mono[OF zero_less_Suc `infinite S`, of n] `infinite S` | |
| 486 | apply (subst (1 2) enumerate_Suc') | |
| 487 | apply (subst Suc) | |
| 53239 | 488 | using `infinite S` | 
| 489 | apply simp | |
| 490 | apply (intro arg_cong[where f = Least] ext) | |
| 491 | apply (auto simp: enumerate_Suc'[symmetric]) | |
| 492 | done | |
| 50134 | 493 | qed | 
| 494 | ||
| 495 | lemma enumerate_Ex: | |
| 496 | assumes S: "infinite (S::nat set)" | |
| 497 | shows "s \<in> S \<Longrightarrow> \<exists>n. enumerate S n = s" | |
| 498 | proof (induct s rule: less_induct) | |
| 499 | case (less s) | |
| 500 | show ?case | |
| 501 | proof cases | |
| 502 |     let ?y = "Max {s'\<in>S. s' < s}"
 | |
| 503 | assume "\<exists>y\<in>S. y < s" | |
| 53239 | 504 | then have y: "\<And>x. ?y < x \<longleftrightarrow> (\<forall>s'\<in>S. s' < s \<longrightarrow> s' < x)" | 
| 505 | by (subst Max_less_iff) auto | |
| 506 |     then have y_in: "?y \<in> {s'\<in>S. s' < s}"
 | |
| 507 | by (intro Max_in) auto | |
| 508 | with less.hyps[of ?y] obtain n where "enumerate S n = ?y" | |
| 509 | by auto | |
| 50134 | 510 | with S have "enumerate S (Suc n) = s" | 
| 511 | by (auto simp: y less enumerate_Suc'' intro!: Least_equality) | |
| 512 | then show ?case by auto | |
| 513 | next | |
| 514 | assume *: "\<not> (\<exists>y\<in>S. y < s)" | |
| 515 | then have "\<forall>t\<in>S. s \<le> t" by auto | |
| 516 | with `s \<in> S` show ?thesis | |
| 517 | by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) | |
| 518 | qed | |
| 519 | qed | |
| 520 | ||
| 521 | lemma bij_enumerate: | |
| 522 | fixes S :: "nat set" | |
| 523 | assumes S: "infinite S" | |
| 524 | shows "bij_betw (enumerate S) UNIV S" | |
| 525 | proof - | |
| 526 | have "\<And>n m. n \<noteq> m \<Longrightarrow> enumerate S n \<noteq> enumerate S m" | |
| 527 | using enumerate_mono[OF _ `infinite S`] by (auto simp: neq_iff) | |
| 528 | then have "inj (enumerate S)" | |
| 529 | by (auto simp: inj_on_def) | |
| 53239 | 530 | moreover have "\<forall>s \<in> S. \<exists>i. enumerate S i = s" | 
| 50134 | 531 | using enumerate_Ex[OF S] by auto | 
| 532 | moreover note `infinite S` | |
| 533 | ultimately show ?thesis | |
| 534 | unfolding bij_betw_def by (auto intro: enumerate_in_set) | |
| 535 | qed | |
| 536 | ||
| 20809 | 537 | subsection "Miscellaneous" | 
| 538 | ||
| 539 | text {*
 | |
| 540 | A few trivial lemmas about sets that contain at most one element. | |
| 541 | These simplify the reasoning about deterministic automata. | |
| 542 | *} | |
| 543 | ||
| 53239 | 544 | definition atmost_one :: "'a set \<Rightarrow> bool" | 
| 545 | where "atmost_one S \<longleftrightarrow> (\<forall>x y. x\<in>S \<and> y\<in>S \<longrightarrow> x = y)" | |
| 20809 | 546 | |
| 547 | lemma atmost_one_empty: "S = {} \<Longrightarrow> atmost_one S"
 | |
| 548 | by (simp add: atmost_one_def) | |
| 549 | ||
| 550 | lemma atmost_one_singleton: "S = {x} \<Longrightarrow> atmost_one S"
 | |
| 551 | by (simp add: atmost_one_def) | |
| 552 | ||
| 553 | lemma atmost_one_unique [elim]: "atmost_one S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> y = x" | |
| 554 | by (simp add: atmost_one_def) | |
| 555 | ||
| 556 | end | |
| 54612 
7e291ae244ea
Backed out changeset: a8ad7f6dd217---bypassing Main breaks theories that use \<inf> or \<sup>
 traytel parents: 
54607diff
changeset | 557 |