| author | hoelzl | 
| Thu, 08 Oct 2015 11:19:43 +0200 | |
| changeset 61362 | 48d1b147f094 | 
| parent 60828 | e9e272fa8daf | 
| child 61585 | a9599d3d7610 | 
| 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 | |
| 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 | ||
| 60500 | 13 | text \<open> | 
| 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"}.
 | |
| 60500 | 17 | \<close> | 
| 20809 | 18 | |
| 53239 | 19 | abbreviation infinite :: "'a set \<Rightarrow> bool" | 
| 20 | where "infinite S \<equiv> \<not> finite S" | |
| 20809 | 21 | |
| 60500 | 22 | text \<open> | 
| 20809 | 23 | Infinite sets are non-empty, and if we remove some elements from an | 
| 24 | infinite set, the result is still infinite. | |
| 60500 | 25 | \<close> | 
| 20809 | 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 | ||
| 60828 | 65 | lemma infinite_coinduct [consumes 1, case_names infinite]: | 
| 66 | assumes "X A" | |
| 67 |   and step: "\<And>A. X A \<Longrightarrow> \<exists>x\<in>A. X (A - {x}) \<or> infinite (A - {x})"
 | |
| 68 | shows "infinite A" | |
| 69 | proof | |
| 70 | assume "finite A" | |
| 71 | then show False using \<open>X A\<close> | |
| 72 | by(induction rule: finite_psubset_induct)(meson Diff_subset card_Diff1_less card_psubset finite_Diff step) | |
| 73 | qed | |
| 74 | ||
| 60500 | 75 | text \<open> | 
| 20809 | 76 | As a concrete example, we prove that the set of natural numbers is | 
| 77 | infinite. | |
| 60500 | 78 | \<close> | 
| 20809 | 79 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 80 | lemma infinite_nat_iff_unbounded_le: "infinite (S::nat set) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. n \<in> S)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 81 | 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 | 82 | 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 | 83 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 84 | lemma infinite_nat_iff_unbounded: "infinite (S::nat set) \<longleftrightarrow> (\<forall>m. \<exists>n>m. n \<in> S)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 85 | 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 | 86 | by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) | 
| 20809 | 87 | |
| 59000 | 88 | lemma finite_nat_iff_bounded: "finite (S::nat set) \<longleftrightarrow> (\<exists>k. S \<subseteq> {..<k})"
 | 
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 89 | using infinite_nat_iff_unbounded_le[of S] by (simp add: subset_eq) (metis not_le) | 
| 20809 | 90 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 91 | lemma finite_nat_iff_bounded_le: "finite (S::nat set) \<longleftrightarrow> (\<exists>k. S \<subseteq> {.. k})"
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 92 | using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) | 
| 20809 | 93 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 94 | lemma finite_nat_bounded: "finite (S::nat set) \<Longrightarrow> \<exists>k. S \<subseteq> {..<k}"
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 95 | by (simp add: finite_nat_iff_bounded) | 
| 20809 | 96 | |
| 60500 | 97 | text \<open> | 
| 20809 | 98 | For a set of natural numbers to be infinite, it is enough to know | 
| 99 |   that for any number larger than some @{text k}, there is some larger
 | |
| 100 | number that is an element of the set. | |
| 60500 | 101 | \<close> | 
| 20809 | 102 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 103 | lemma unbounded_k_infinite: "\<forall>m>k. \<exists>n>m. n \<in> S \<Longrightarrow> infinite (S::nat set)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 104 | by (metis (full_types) infinite_nat_iff_unbounded_le le_imp_less_Suc not_less | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 105 | not_less_iff_gr_or_eq sup.bounded_iff) | 
| 20809 | 106 | |
| 35056 | 107 | lemma nat_not_finite: "finite (UNIV::nat set) \<Longrightarrow> R" | 
| 20809 | 108 | by simp | 
| 109 | ||
| 110 | lemma range_inj_infinite: | |
| 111 | "inj (f::nat \<Rightarrow> 'a) \<Longrightarrow> infinite (range f)" | |
| 112 | proof | |
| 27407 
68e111812b83
rename lemmas INF_foo to INFM_foo; add new lemmas about MOST and INFM
 huffman parents: 
27368diff
changeset | 113 | assume "finite (range f)" and "inj f" | 
| 20809 | 114 | 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 | 115 | by (rule finite_imageD) | 
| 20809 | 116 | then show False by simp | 
| 117 | qed | |
| 118 | ||
| 60500 | 119 | text \<open> | 
| 20809 | 120 | For any function with infinite domain and finite range there is some | 
| 121 | element that is the image of infinitely many domain elements. In | |
| 122 | particular, any infinite sequence of elements from a finite set | |
| 123 | contains some element that occurs infinitely often. | |
| 60500 | 124 | \<close> | 
| 20809 | 125 | |
| 59488 | 126 | lemma inf_img_fin_dom': | 
| 127 | assumes img: "finite (f ` A)" and dom: "infinite A" | |
| 128 |   shows "\<exists>y \<in> f ` A. infinite (f -` {y} \<inter> A)"
 | |
| 129 | proof (rule ccontr) | |
| 130 |   have "A \<subseteq> (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by auto
 | |
| 131 | moreover | |
| 132 | assume "\<not> ?thesis" | |
| 133 |   with img have "finite (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by blast
 | |
| 134 | ultimately have "finite A" by(rule finite_subset) | |
| 135 | with dom show False by contradiction | |
| 136 | qed | |
| 137 | ||
| 138 | lemma inf_img_fin_domE': | |
| 139 | assumes "finite (f ` A)" and "infinite A" | |
| 140 |   obtains y where "y \<in> f`A" and "infinite (f -` {y} \<inter> A)"
 | |
| 141 | using assms by (blast dest: inf_img_fin_dom') | |
| 142 | ||
| 20809 | 143 | lemma inf_img_fin_dom: | 
| 144 | assumes img: "finite (f`A)" and dom: "infinite A" | |
| 145 |   shows "\<exists>y \<in> f`A. infinite (f -` {y})"
 | |
| 59488 | 146 | using inf_img_fin_dom'[OF assms] by auto | 
| 20809 | 147 | |
| 148 | lemma inf_img_fin_domE: | |
| 149 | assumes "finite (f`A)" and "infinite A" | |
| 150 |   obtains y where "y \<in> f`A" and "infinite (f -` {y})"
 | |
| 23394 | 151 | using assms by (blast dest: inf_img_fin_dom) | 
| 20809 | 152 | |
| 153 | subsection "Infinitely Many and Almost All" | |
| 154 | ||
| 60500 | 155 | text \<open> | 
| 20809 | 156 | We often need to reason about the existence of infinitely many | 
| 157 | (resp., all but finitely many) objects satisfying some predicate, so | |
| 158 | we introduce corresponding binders and their proof rules. | |
| 60500 | 159 | \<close> | 
| 20809 | 160 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 161 | (* The following two lemmas are available as filter-rules, but not in the simp-set *) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 162 | lemma not_INFM [simp]: "\<not> (INFM x. P x) \<longleftrightarrow> (MOST x. \<not> P x)" by (fact not_frequently) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 163 | lemma not_MOST [simp]: "\<not> (MOST x. P x) \<longleftrightarrow> (INFM x. \<not> P x)" by (fact not_eventually) | 
| 34112 | 164 | |
| 165 | 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 | 166 | by (simp add: frequently_const_iff) | 
| 34112 | 167 | |
| 168 | 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 | 169 | by (simp add: eventually_const_iff) | 
| 20809 | 170 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 171 | lemma INFM_imp_distrib: "(INFM x. P x \<longrightarrow> Q x) \<longleftrightarrow> ((MOST x. P x) \<longrightarrow> (INFM x. Q x))" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 172 | by (simp only: imp_conv_disj frequently_disj_iff not_eventually) | 
| 34112 | 173 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 174 | lemma MOST_imp_iff: "MOST x. P x \<Longrightarrow> (MOST x. P x \<longrightarrow> Q x) \<longleftrightarrow> (MOST x. Q x)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 175 | by (auto intro: eventually_rev_mp eventually_elim1) | 
| 34113 | 176 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 177 | lemma INFM_conjI: "INFM x. P x \<Longrightarrow> MOST x. Q x \<Longrightarrow> INFM x. P x \<and> Q x" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 178 | by (rule frequently_rev_mp[of P]) (auto elim: eventually_elim1) | 
| 34112 | 179 | |
| 60500 | 180 | text \<open>Properties of quantifiers with injective functions.\<close> | 
| 34112 | 181 | |
| 53239 | 182 | 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 | 183 |   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 | 184 | |
| 53239 | 185 | 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 | 186 |   using finite_vimageI[of "{x. \<not> P x}" f] by (auto simp: eventually_cofinite)
 | 
| 34112 | 187 | |
| 60500 | 188 | text \<open>Properties of quantifiers with singletons.\<close> | 
| 34112 | 189 | |
| 190 | lemma not_INFM_eq [simp]: | |
| 191 | "\<not> (INFM x. x = a)" | |
| 192 | "\<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 | 193 | unfolding frequently_cofinite by simp_all | 
| 34112 | 194 | |
| 195 | lemma MOST_neq [simp]: | |
| 196 | "MOST x. x \<noteq> a" | |
| 197 | "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 | 198 | 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 | 199 | |
| 34112 | 200 | lemma INFM_neq [simp]: | 
| 201 | "(INFM x::'a. x \<noteq> a) \<longleftrightarrow> infinite (UNIV::'a set)" | |
| 202 | "(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 | 203 | unfolding frequently_cofinite by simp_all | 
| 34112 | 204 | |
| 205 | lemma MOST_eq [simp]: | |
| 206 | "(MOST x::'a. x = a) \<longleftrightarrow> finite (UNIV::'a set)" | |
| 207 | "(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 | 208 | unfolding eventually_cofinite by simp_all | 
| 34112 | 209 | |
| 210 | lemma MOST_eq_imp: | |
| 211 | "MOST x. x = a \<longrightarrow> P x" | |
| 212 | "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 | 213 | unfolding eventually_cofinite by simp_all | 
| 34112 | 214 | |
| 60500 | 215 | 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 | 216 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 217 | lemma MOST_nat: "(\<forall>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<exists>m. \<forall>n>m. P n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 218 | by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq not_le[symmetric]) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 219 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 220 | lemma MOST_nat_le: "(\<forall>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<exists>m. \<forall>n\<ge>m. P n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 221 | by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq not_le[symmetric]) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 222 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 223 | lemma INFM_nat: "(\<exists>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<forall>m. \<exists>n>m. P n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 224 | by (simp add: frequently_cofinite infinite_nat_iff_unbounded) | 
| 20809 | 225 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 226 | lemma INFM_nat_le: "(\<exists>\<^sub>\<infinity>n. P (n::nat)) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. P n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 227 | 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 | 228 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 229 | 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 | 230 | 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 | 231 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 232 | lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 233 | by (simp add: cofinite_eq_sequentially eventually_sequentially_Suc) | 
| 20809 | 234 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 235 | lemma | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 236 | shows MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 237 | and MOST_SucD: "MOST n. P (Suc n) \<Longrightarrow> MOST n. P n" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 238 | 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 | 239 | |
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 240 | lemma MOST_ge_nat: "MOST n::nat. m \<le> n" | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 241 | by (simp add: cofinite_eq_sequentially eventually_ge_at_top) | 
| 20809 | 242 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 243 | (* legacy names *) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 244 | 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 | 245 | 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 | 246 | 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 | 247 | 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 | 248 | 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 | 249 | 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 | 250 | 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) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 251 | 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_elim1) | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 252 | 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 | 253 | 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 | 254 | 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 | 255 | 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 | 256 | 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 | 257 | 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 | 258 | 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 | 259 | 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 | 260 | lemmas MOST_iff_finiteNeg = MOST_iff_cofinite | 
| 20809 | 261 | |
| 262 | subsection "Enumeration of an Infinite Set" | |
| 263 | ||
| 60500 | 264 | text \<open> | 
| 20809 | 265 | The set's element type must be wellordered (e.g. the natural numbers). | 
| 60500 | 266 | \<close> | 
| 20809 | 267 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 268 | text \<open> | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 269 | Could be generalized to | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 270 |     @{term "enumerate' S n = (SOME t. t \<in> s \<and> finite {s\<in>S. s < t} \<and> card {s\<in>S. s < t} = n)"}.
 | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 271 | \<close> | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 272 | |
| 53239 | 273 | primrec (in wellorder) enumerate :: "'a set \<Rightarrow> nat \<Rightarrow> 'a" | 
| 274 | where | |
| 275 | enumerate_0: "enumerate S 0 = (LEAST n. n \<in> S)" | |
| 276 | | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \<in> S}) n"
 | |
| 20809 | 277 | |
| 53239 | 278 | lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n"
 | 
| 20809 | 279 | by simp | 
| 280 | ||
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 281 | lemma enumerate_in_set: "infinite S \<Longrightarrow> enumerate S n \<in> S" | 
| 53239 | 282 | apply (induct n arbitrary: S) | 
| 283 | apply (fastforce intro: LeastI dest!: infinite_imp_nonempty) | |
| 284 | apply simp | |
| 285 | apply (metis DiffE infinite_remove) | |
| 286 | done | |
| 20809 | 287 | |
| 288 | declare enumerate_0 [simp del] enumerate_Suc [simp del] | |
| 289 | ||
| 290 | lemma enumerate_step: "infinite S \<Longrightarrow> enumerate S n < enumerate S (Suc n)" | |
| 291 | apply (induct n arbitrary: S) | |
| 292 | apply (rule order_le_neq_trans) | |
| 293 | apply (simp add: enumerate_0 Least_le enumerate_in_set) | |
| 294 | apply (simp only: enumerate_Suc') | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 295 |    apply (subgoal_tac "enumerate (S - {enumerate S 0}) 0 \<in> S - {enumerate S 0}")
 | 
| 20809 | 296 | apply (blast intro: sym) | 
| 297 | apply (simp add: enumerate_in_set del: Diff_iff) | |
| 298 | apply (simp add: enumerate_Suc') | |
| 299 | done | |
| 300 | ||
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
59506diff
changeset | 301 | lemma enumerate_mono: "m < n \<Longrightarrow> infinite S \<Longrightarrow> enumerate S m < enumerate S n" | 
| 20809 | 302 | apply (erule less_Suc_induct) | 
| 303 | apply (auto intro: enumerate_step) | |
| 304 | done | |
| 305 | ||
| 306 | ||
| 50134 | 307 | lemma le_enumerate: | 
| 308 | assumes S: "infinite S" | |
| 309 | shows "n \<le> enumerate S n" | |
| 310 | using S | |
| 311 | proof (induct n) | |
| 53239 | 312 | case 0 | 
| 313 | then show ?case by simp | |
| 314 | next | |
| 50134 | 315 | case (Suc n) | 
| 316 | then have "n \<le> enumerate S n" by simp | |
| 60500 | 317 | also note enumerate_mono[of n "Suc n", OF _ \<open>infinite S\<close>] | 
| 50134 | 318 | finally show ?case by simp | 
| 53239 | 319 | qed | 
| 50134 | 320 | |
| 321 | lemma enumerate_Suc'': | |
| 322 | fixes S :: "'a::wellorder set" | |
| 53239 | 323 | assumes "infinite S" | 
| 324 | shows "enumerate S (Suc n) = (LEAST s. s \<in> S \<and> enumerate S n < s)" | |
| 325 | using assms | |
| 50134 | 326 | proof (induct n arbitrary: S) | 
| 327 | case 0 | |
| 53239 | 328 | then have "\<forall>s \<in> S. enumerate S 0 \<le> s" | 
| 50134 | 329 | by (auto simp: enumerate.simps intro: Least_le) | 
| 330 | then show ?case | |
| 331 |     unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"]
 | |
| 53239 | 332 | by (intro arg_cong[where f = Least] ext) auto | 
| 50134 | 333 | next | 
| 334 | case (Suc n S) | |
| 335 | show ?case | |
| 60500 | 336 | using enumerate_mono[OF zero_less_Suc \<open>infinite S\<close>, of n] \<open>infinite S\<close> | 
| 50134 | 337 | apply (subst (1 2) enumerate_Suc') | 
| 338 | apply (subst Suc) | |
| 60500 | 339 | using \<open>infinite S\<close> | 
| 53239 | 340 | apply simp | 
| 341 | apply (intro arg_cong[where f = Least] ext) | |
| 342 | apply (auto simp: enumerate_Suc'[symmetric]) | |
| 343 | done | |
| 50134 | 344 | qed | 
| 345 | ||
| 346 | lemma enumerate_Ex: | |
| 347 | assumes S: "infinite (S::nat set)" | |
| 348 | shows "s \<in> S \<Longrightarrow> \<exists>n. enumerate S n = s" | |
| 349 | proof (induct s rule: less_induct) | |
| 350 | case (less s) | |
| 351 | show ?case | |
| 352 | proof cases | |
| 353 |     let ?y = "Max {s'\<in>S. s' < s}"
 | |
| 354 | assume "\<exists>y\<in>S. y < s" | |
| 53239 | 355 | then have y: "\<And>x. ?y < x \<longleftrightarrow> (\<forall>s'\<in>S. s' < s \<longrightarrow> s' < x)" | 
| 356 | by (subst Max_less_iff) auto | |
| 357 |     then have y_in: "?y \<in> {s'\<in>S. s' < s}"
 | |
| 358 | by (intro Max_in) auto | |
| 359 | with less.hyps[of ?y] obtain n where "enumerate S n = ?y" | |
| 360 | by auto | |
| 50134 | 361 | with S have "enumerate S (Suc n) = s" | 
| 362 | by (auto simp: y less enumerate_Suc'' intro!: Least_equality) | |
| 363 | then show ?case by auto | |
| 364 | next | |
| 365 | assume *: "\<not> (\<exists>y\<in>S. y < s)" | |
| 366 | then have "\<forall>t\<in>S. s \<le> t" by auto | |
| 60500 | 367 | with \<open>s \<in> S\<close> show ?thesis | 
| 50134 | 368 | by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) | 
| 369 | qed | |
| 370 | qed | |
| 371 | ||
| 372 | lemma bij_enumerate: | |
| 373 | fixes S :: "nat set" | |
| 374 | assumes S: "infinite S" | |
| 375 | shows "bij_betw (enumerate S) UNIV S" | |
| 376 | proof - | |
| 377 | have "\<And>n m. n \<noteq> m \<Longrightarrow> enumerate S n \<noteq> enumerate S m" | |
| 60500 | 378 | using enumerate_mono[OF _ \<open>infinite S\<close>] by (auto simp: neq_iff) | 
| 50134 | 379 | then have "inj (enumerate S)" | 
| 380 | by (auto simp: inj_on_def) | |
| 53239 | 381 | moreover have "\<forall>s \<in> S. \<exists>i. enumerate S i = s" | 
| 50134 | 382 | using enumerate_Ex[OF S] by auto | 
| 60500 | 383 | moreover note \<open>infinite S\<close> | 
| 50134 | 384 | ultimately show ?thesis | 
| 385 | unfolding bij_betw_def by (auto intro: enumerate_in_set) | |
| 386 | qed | |
| 387 | ||
| 20809 | 388 | subsection "Miscellaneous" | 
| 389 | ||
| 60500 | 390 | text \<open> | 
| 20809 | 391 | A few trivial lemmas about sets that contain at most one element. | 
| 392 | These simplify the reasoning about deterministic automata. | |
| 60500 | 393 | \<close> | 
| 20809 | 394 | |
| 53239 | 395 | definition atmost_one :: "'a set \<Rightarrow> bool" | 
| 396 | where "atmost_one S \<longleftrightarrow> (\<forall>x y. x\<in>S \<and> y\<in>S \<longrightarrow> x = y)" | |
| 20809 | 397 | |
| 398 | lemma atmost_one_empty: "S = {} \<Longrightarrow> atmost_one S"
 | |
| 399 | by (simp add: atmost_one_def) | |
| 400 | ||
| 401 | lemma atmost_one_singleton: "S = {x} \<Longrightarrow> atmost_one S"
 | |
| 402 | by (simp add: atmost_one_def) | |
| 403 | ||
| 404 | lemma atmost_one_unique [elim]: "atmost_one S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> y = x" | |
| 405 | by (simp add: atmost_one_def) | |
| 406 | ||
| 407 | end | |
| 54612 
7e291ae244ea
Backed out changeset: a8ad7f6dd217---bypassing Main breaks theories that use \<inf> or \<sup>
 traytel parents: 
54607diff
changeset | 408 |