| author | wenzelm | 
| Sun, 16 Jun 2024 11:50:42 +0200 | |
| changeset 80385 | 605e19319343 | 
| parent 78200 | 264f2b69d09c | 
| child 82802 | 547335b41005 | 
| permissions | -rw-r--r-- | 
| 50134 | 1 | (* Title: HOL/Library/Countable_Set.thy | 
| 2 | Author: Johannes Hölzl | |
| 3 | Author: Andrei Popescu | |
| 4 | *) | |
| 5 | ||
| 60500 | 6 | section \<open>Countable sets\<close> | 
| 50134 | 7 | |
| 8 | theory Countable_Set | |
| 51542 | 9 | imports Countable Infinite_Set | 
| 50134 | 10 | begin | 
| 11 | ||
| 60500 | 12 | subsection \<open>Predicate for countable sets\<close> | 
| 50134 | 13 | |
| 14 | definition countable :: "'a set \<Rightarrow> bool" where | |
| 15 | "countable S \<longleftrightarrow> (\<exists>f::'a \<Rightarrow> nat. inj_on f S)" | |
| 16 | ||
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 17 | lemma countable_as_injective_image_subset: "countable S \<longleftrightarrow> (\<exists>f. \<exists>K::nat set. S = f ` K \<and> inj_on f K)" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 18 | by (metis countable_def inj_on_the_inv_into the_inv_into_onto) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77138diff
changeset | 19 | |
| 50134 | 20 | lemma countableE: | 
| 21 | assumes S: "countable S" obtains f :: "'a \<Rightarrow> nat" where "inj_on f S" | |
| 22 | using S by (auto simp: countable_def) | |
| 23 | ||
| 24 | lemma countableI: "inj_on (f::'a \<Rightarrow> nat) S \<Longrightarrow> countable S" | |
| 25 | by (auto simp: countable_def) | |
| 26 | ||
| 27 | lemma countableI': "inj_on (f::'a \<Rightarrow> 'b::countable) S \<Longrightarrow> countable S" | |
| 28 | using comp_inj_on[of f S to_nat] by (auto intro: countableI) | |
| 29 | ||
| 30 | lemma countableE_bij: | |
| 31 | assumes S: "countable S" obtains f :: "nat \<Rightarrow> 'a" and C :: "nat set" where "bij_betw f C S" | |
| 32 | using S by (blast elim: countableE dest: inj_on_imp_bij_betw bij_betw_inv) | |
| 33 | ||
| 34 | lemma countableI_bij: "bij_betw f (C::nat set) S \<Longrightarrow> countable S" | |
| 35 | by (blast intro: countableI bij_betw_inv_into bij_betw_imp_inj_on) | |
| 36 | ||
| 37 | lemma countable_finite: "finite S \<Longrightarrow> countable S" | |
| 38 | by (blast dest: finite_imp_inj_to_nat_seg countableI) | |
| 39 | ||
| 40 | lemma countableI_bij1: "bij_betw f A B \<Longrightarrow> countable A \<Longrightarrow> countable B" | |
| 41 | by (blast elim: countableE_bij intro: bij_betw_trans countableI_bij) | |
| 42 | ||
| 43 | lemma countableI_bij2: "bij_betw f B A \<Longrightarrow> countable A \<Longrightarrow> countable B" | |
| 44 | by (blast elim: countableE_bij intro: bij_betw_trans bij_betw_inv_into countableI_bij) | |
| 45 | ||
| 46 | lemma countable_iff_bij[simp]: "bij_betw f A B \<Longrightarrow> countable A \<longleftrightarrow> countable B" | |
| 47 | by (blast intro: countableI_bij1 countableI_bij2) | |
| 48 | ||
| 49 | lemma countable_subset: "A \<subseteq> B \<Longrightarrow> countable B \<Longrightarrow> countable A" | |
| 50 | by (auto simp: countable_def intro: subset_inj_on) | |
| 51 | ||
| 52 | lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)" | |
| 53 | using countableI[of to_nat A] by auto | |
| 54 | ||
| 60500 | 55 | subsection \<open>Enumerate a countable set\<close> | 
| 50134 | 56 | |
| 57 | lemma countableE_infinite: | |
| 58 | assumes "countable S" "infinite S" | |
| 59 | obtains e :: "'a \<Rightarrow> nat" where "bij_betw e S UNIV" | |
| 60 | proof - | |
| 53381 | 61 | obtain f :: "'a \<Rightarrow> nat" where "inj_on f S" | 
| 60500 | 62 | using \<open>countable S\<close> by (rule countableE) | 
| 50134 | 63 | then have "bij_betw f S (f`S)" | 
| 64 | unfolding bij_betw_def by simp | |
| 65 | moreover | |
| 60500 | 66 | from \<open>inj_on f S\<close> \<open>infinite S\<close> have inf_fS: "infinite (f`S)" | 
| 50134 | 67 | by (auto dest: finite_imageD) | 
| 68 | then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV" | |
| 69 | by (intro bij_betw_the_inv_into bij_enumerate) | |
| 70 | ultimately have "bij_betw (the_inv_into UNIV (enumerate (f`S)) \<circ> f) S UNIV" | |
| 71 | by (rule bij_betw_trans) | |
| 72 | then show thesis .. | |
| 73 | qed | |
| 74 | ||
| 68860 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
67613diff
changeset | 75 | lemma countable_infiniteE': | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
67613diff
changeset | 76 | assumes "countable A" "infinite A" | 
| 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
67613diff
changeset | 77 | obtains g where "bij_betw g (UNIV :: nat set) A" | 
| 77138 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
74438diff
changeset | 78 | by (meson assms bij_betw_inv countableE_infinite) | 
| 68860 
f443ec10447d
Some basic materials on filters and topology
 Manuel Eberl <eberlm@in.tum.de> parents: 
67613diff
changeset | 79 | |
| 50134 | 80 | lemma countable_enum_cases: | 
| 81 | assumes "countable S" | |
| 82 |   obtains (finite) f :: "'a \<Rightarrow> nat" where "finite S" "bij_betw f S {..<card S}"
 | |
| 83 | | (infinite) f :: "'a \<Rightarrow> nat" where "infinite S" "bij_betw f S UNIV" | |
| 60500 | 84 | using ex_bij_betw_finite_nat[of S] countableE_infinite \<open>countable S\<close> | 
| 50134 | 85 | by (cases "finite S") (auto simp add: atLeast0LessThan) | 
| 86 | ||
| 87 | definition to_nat_on :: "'a set \<Rightarrow> 'a \<Rightarrow> nat" where | |
| 88 |   "to_nat_on S = (SOME f. if finite S then bij_betw f S {..< card S} else bij_betw f S UNIV)"
 | |
| 89 | ||
| 90 | definition from_nat_into :: "'a set \<Rightarrow> nat \<Rightarrow> 'a" where | |
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 91 | "from_nat_into S n = (if n \<in> to_nat_on S ` S then inv_into S (to_nat_on S) n else SOME s. s\<in>S)" | 
| 50134 | 92 | |
| 93 | lemma to_nat_on_finite: "finite S \<Longrightarrow> bij_betw (to_nat_on S) S {..< card S}"
 | |
| 94 | using ex_bij_betw_finite_nat unfolding to_nat_on_def | |
| 95 |   by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S {..<card S}"]) (auto simp add: atLeast0LessThan)
 | |
| 96 | ||
| 97 | lemma to_nat_on_infinite: "countable S \<Longrightarrow> infinite S \<Longrightarrow> bij_betw (to_nat_on S) S UNIV" | |
| 98 | using countableE_infinite unfolding to_nat_on_def | |
| 99 | by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S UNIV"]) auto | |
| 100 | ||
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 101 | lemma bij_betw_from_nat_into_finite: "finite S \<Longrightarrow> bij_betw (from_nat_into S) {..< card S} S"
 | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 102 | unfolding from_nat_into_def[abs_def] | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 103 | using to_nat_on_finite[of S] | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 104 | apply (subst bij_betw_cong) | 
| 62390 | 105 | apply (split if_split) | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 106 | apply (simp add: bij_betw_def) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 107 | apply (auto cong: bij_betw_cong | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 108 | intro: bij_betw_inv_into to_nat_on_finite) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 109 | done | 
| 50134 | 110 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 111 | lemma bij_betw_from_nat_into: "countable S \<Longrightarrow> infinite S \<Longrightarrow> bij_betw (from_nat_into S) UNIV S" | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 112 | unfolding from_nat_into_def[abs_def] | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 113 | using to_nat_on_infinite[of S, unfolded bij_betw_def] | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 114 | by (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_infinite) | 
| 50134 | 115 | |
| 74438 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 116 | text \<open> | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 117 | The sum/product over the enumeration of a finite set equals simply the sum/product over the set | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 118 | \<close> | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 119 | context comm_monoid_set | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 120 | begin | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 121 | |
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 122 | lemma card_from_nat_into: | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 123 |   "F (\<lambda>i. h (from_nat_into A i)) {..<card A} = F h A"
 | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 124 | proof (cases "finite A") | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 125 | case True | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 126 |   have "F (\<lambda>i. h (from_nat_into A i)) {..<card A} = F h (from_nat_into A ` {..<card A})"
 | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 127 | by (metis True bij_betw_def bij_betw_from_nat_into_finite reindex_cong) | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 128 | also have "... = F h A" | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 129 | by (metis True bij_betw_def bij_betw_from_nat_into_finite) | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 130 | finally show ?thesis . | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 131 | qed auto | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 132 | |
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 133 | end | 
| 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 paulson <lp15@cam.ac.uk> parents: 
74325diff
changeset | 134 | |
| 63127 | 135 | lemma countable_as_injective_image: | 
| 136 | assumes "countable A" "infinite A" | |
| 137 | obtains f :: "nat \<Rightarrow> 'a" where "A = range f" "inj f" | |
| 138 | by (metis bij_betw_def bij_betw_from_nat_into [OF assms]) | |
| 139 | ||
| 50134 | 140 | lemma inj_on_to_nat_on[intro]: "countable A \<Longrightarrow> inj_on (to_nat_on A) A" | 
| 141 | using to_nat_on_infinite[of A] to_nat_on_finite[of A] | |
| 142 | by (cases "finite A") (auto simp: bij_betw_def) | |
| 143 | ||
| 144 | lemma to_nat_on_inj[simp]: | |
| 145 | "countable A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> to_nat_on A a = to_nat_on A b \<longleftrightarrow> a = b" | |
| 146 | using inj_on_to_nat_on[of A] by (auto dest: inj_onD) | |
| 147 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 148 | lemma from_nat_into_to_nat_on[simp]: "countable A \<Longrightarrow> a \<in> A \<Longrightarrow> from_nat_into A (to_nat_on A a) = a" | 
| 50134 | 149 | by (auto simp: from_nat_into_def intro!: inv_into_f_f) | 
| 150 | ||
| 151 | lemma subset_range_from_nat_into: "countable A \<Longrightarrow> A \<subseteq> range (from_nat_into A)" | |
| 152 | by (auto intro: from_nat_into_to_nat_on[symmetric]) | |
| 153 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 154 | lemma from_nat_into: "A \<noteq> {} \<Longrightarrow> from_nat_into A n \<in> A"
 | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 155 | unfolding from_nat_into_def by (metis equals0I inv_into_into someI_ex) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 156 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 157 | lemma range_from_nat_into_subset: "A \<noteq> {} \<Longrightarrow> range (from_nat_into A) \<subseteq> A"
 | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 158 | using from_nat_into[of A] by auto | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 159 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 160 | lemma range_from_nat_into[simp]: "A \<noteq> {} \<Longrightarrow> countable A \<Longrightarrow> range (from_nat_into A) = A"
 | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 161 | by (metis equalityI range_from_nat_into_subset subset_range_from_nat_into) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 162 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 163 | lemma image_to_nat_on: "countable A \<Longrightarrow> infinite A \<Longrightarrow> to_nat_on A ` A = UNIV" | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 164 | using to_nat_on_infinite[of A] by (simp add: bij_betw_def) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 165 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 166 | lemma to_nat_on_surj: "countable A \<Longrightarrow> infinite A \<Longrightarrow> \<exists>a\<in>A. to_nat_on A a = n" | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 167 | by (metis (no_types) image_iff iso_tuple_UNIV_I image_to_nat_on) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 168 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 169 | lemma to_nat_on_from_nat_into[simp]: "n \<in> to_nat_on A ` A \<Longrightarrow> to_nat_on A (from_nat_into A n) = n" | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 170 | by (simp add: f_inv_into_f from_nat_into_def) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 171 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 172 | lemma to_nat_on_from_nat_into_infinite[simp]: | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 173 | "countable A \<Longrightarrow> infinite A \<Longrightarrow> to_nat_on A (from_nat_into A n) = n" | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 174 | by (metis image_iff to_nat_on_surj to_nat_on_from_nat_into) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 175 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 176 | lemma from_nat_into_inj: | 
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 177 | "countable A \<Longrightarrow> m \<in> to_nat_on A ` A \<Longrightarrow> n \<in> to_nat_on A ` A \<Longrightarrow> | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 178 | from_nat_into A m = from_nat_into A n \<longleftrightarrow> m = n" | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 179 | by (subst to_nat_on_inj[symmetric, of A]) auto | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 180 | |
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 181 | lemma from_nat_into_inj_infinite[simp]: | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 182 | "countable A \<Longrightarrow> infinite A \<Longrightarrow> from_nat_into A m = from_nat_into A n \<longleftrightarrow> m = n" | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 183 | using image_to_nat_on[of A] from_nat_into_inj[of A m n] by simp | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 184 | |
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 185 | lemma eq_from_nat_into_iff: | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 186 | "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> i \<in> to_nat_on A ` A \<Longrightarrow> x = from_nat_into A i \<longleftrightarrow> i = to_nat_on A x" | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 187 | by auto | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 188 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 189 | lemma from_nat_into_surj: "countable A \<Longrightarrow> a \<in> A \<Longrightarrow> \<exists>n. from_nat_into A n = a" | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 190 | by (rule exI[of _ "to_nat_on A a"]) simp | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 191 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 192 | lemma from_nat_into_inject[simp]: | 
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 193 |   "A \<noteq> {} \<Longrightarrow> countable A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> countable B \<Longrightarrow> from_nat_into A = from_nat_into B \<longleftrightarrow> A = B"
 | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 194 | by (metis range_from_nat_into) | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 195 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 196 | lemma inj_on_from_nat_into: "inj_on from_nat_into ({A. A \<noteq> {} \<and> countable A})"
 | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 197 | unfolding inj_on_def by auto | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 198 | |
| 60500 | 199 | subsection \<open>Closure properties of countability\<close> | 
| 50134 | 200 | |
| 201 | lemma countable_SIGMA[intro, simp]: | |
| 202 | "countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (A i)) \<Longrightarrow> countable (SIGMA i : I. A i)" | |
| 203 | by (intro countableI'[of "\<lambda>(i, a). (to_nat_on I i, to_nat_on (A i) a)"]) (auto simp: inj_on_def) | |
| 204 | ||
| 53381 | 205 | lemma countable_image[intro, simp]: | 
| 206 | assumes "countable A" | |
| 207 | shows "countable (f`A)" | |
| 50134 | 208 | proof - | 
| 53381 | 209 | obtain g :: "'a \<Rightarrow> nat" where "inj_on g A" | 
| 210 | using assms by (rule countableE) | |
| 50134 | 211 | moreover have "inj_on (inv_into A f) (f`A)" "inv_into A f ` f ` A \<subseteq> A" | 
| 212 | by (auto intro: inj_on_inv_into inv_into_into) | |
| 213 | ultimately show ?thesis | |
| 214 | by (blast dest: comp_inj_on subset_inj_on intro: countableI) | |
| 215 | qed | |
| 216 | ||
| 60303 | 217 | lemma countable_image_inj_on: "countable (f ` A) \<Longrightarrow> inj_on f A \<Longrightarrow> countable A" | 
| 218 | by (metis countable_image the_inv_into_onto) | |
| 219 | ||
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 220 | lemma countable_image_inj_Int_vimage: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 221 | "\<lbrakk>inj_on f S; countable A\<rbrakk> \<Longrightarrow> countable (S \<inter> f -` A)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 222 | by (meson countable_image_inj_on countable_subset image_subset_iff_subset_vimage inf_le2 inj_on_Int) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 223 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 224 | lemma countable_image_inj_gen: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 225 |    "\<lbrakk>inj_on f S; countable A\<rbrakk> \<Longrightarrow> countable {x \<in> S. f x \<in> A}"
 | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 226 | using countable_image_inj_Int_vimage | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 227 | by (auto simp: vimage_def Collect_conj_eq) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 228 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 229 | lemma countable_image_inj_eq: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 230 | "inj_on f S \<Longrightarrow> countable(f ` S) \<longleftrightarrow> countable S" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 231 | using countable_image_inj_on by blast | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 232 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 233 | lemma countable_image_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 234 |    "\<lbrakk>countable A; inj f\<rbrakk> \<Longrightarrow> countable {x. f x \<in> A}"
 | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 235 | by (metis (mono_tags, lifting) countable_image_inj_eq countable_subset image_Collect_subsetI inj_on_inverseI the_inv_f_f) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 236 | |
| 50134 | 237 | lemma countable_UN[intro, simp]: | 
| 238 | fixes I :: "'i set" and A :: "'i => 'a set" | |
| 239 | assumes I: "countable I" | |
| 240 | assumes A: "\<And>i. i \<in> I \<Longrightarrow> countable (A i)" | |
| 241 | shows "countable (\<Union>i\<in>I. A i)" | |
| 242 | proof - | |
| 243 | have "(\<Union>i\<in>I. A i) = snd ` (SIGMA i : I. A i)" by (auto simp: image_iff) | |
| 244 | then show ?thesis by (simp add: assms) | |
| 245 | qed | |
| 246 | ||
| 247 | lemma countable_Un[intro]: "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A \<union> B)" | |
| 248 |   by (rule countable_UN[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B", simplified])
 | |
| 249 | (simp split: bool.split) | |
| 250 | ||
| 251 | lemma countable_Un_iff[simp]: "countable (A \<union> B) \<longleftrightarrow> countable A \<and> countable B" | |
| 252 | by (metis countable_Un countable_subset inf_sup_ord(3,4)) | |
| 253 | ||
| 254 | lemma countable_Plus[intro, simp]: | |
| 255 | "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A <+> B)" | |
| 256 | by (simp add: Plus_def) | |
| 257 | ||
| 258 | lemma countable_empty[intro, simp]: "countable {}"
 | |
| 259 | by (blast intro: countable_finite) | |
| 260 | ||
| 261 | lemma countable_insert[intro, simp]: "countable A \<Longrightarrow> countable (insert a A)" | |
| 262 |   using countable_Un[of "{a}" A] by (auto simp: countable_finite)
 | |
| 263 | ||
| 264 | lemma countable_Int1[intro, simp]: "countable A \<Longrightarrow> countable (A \<inter> B)" | |
| 265 | by (force intro: countable_subset) | |
| 266 | ||
| 267 | lemma countable_Int2[intro, simp]: "countable B \<Longrightarrow> countable (A \<inter> B)" | |
| 268 | by (blast intro: countable_subset) | |
| 269 | ||
| 270 | lemma countable_INT[intro, simp]: "i \<in> I \<Longrightarrow> countable (A i) \<Longrightarrow> countable (\<Inter>i\<in>I. A i)" | |
| 271 | by (blast intro: countable_subset) | |
| 272 | ||
| 273 | lemma countable_Diff[intro, simp]: "countable A \<Longrightarrow> countable (A - B)" | |
| 274 | by (blast intro: countable_subset) | |
| 275 | ||
| 60303 | 276 | lemma countable_insert_eq [simp]: "countable (insert x A) = countable A" | 
| 277 | by auto (metis Diff_insert_absorb countable_Diff insert_absorb) | |
| 278 | ||
| 50134 | 279 | lemma countable_vimage: "B \<subseteq> range f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B" | 
| 63092 | 280 | by (metis Int_absorb2 countable_image image_vimage_eq) | 
| 50134 | 281 | |
| 282 | lemma surj_countable_vimage: "surj f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B" | |
| 283 | by (metis countable_vimage top_greatest) | |
| 284 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 285 | lemma countable_Collect[simp]: "countable A \<Longrightarrow> countable {a \<in> A. \<phi> a}"
 | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 286 | by (metis Collect_conj_eq Int_absorb Int_commute Int_def countable_Int1) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 287 | |
| 54410 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 288 | lemma countable_Image: | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 289 |   assumes "\<And>y. y \<in> Y \<Longrightarrow> countable (X `` {y})"
 | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 290 | assumes "countable Y" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 291 | shows "countable (X `` Y)" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 292 | proof - | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 293 |   have "countable (X `` (\<Union>y\<in>Y. {y}))"
 | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 294 | unfolding Image_UN by (intro countable_UN assms) | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 295 | then show ?thesis by simp | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 296 | qed | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 297 | |
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 298 | lemma countable_relpow: | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 299 | fixes X :: "'a rel" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 300 | assumes Image_X: "\<And>Y. countable Y \<Longrightarrow> countable (X `` Y)" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 301 | assumes Y: "countable Y" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 302 | shows "countable ((X ^^ i) `` Y)" | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 303 | using Y by (induct i arbitrary: Y) (auto simp: relcomp_Image Image_X) | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 304 | |
| 60058 | 305 | lemma countable_funpow: | 
| 306 | fixes f :: "'a set \<Rightarrow> 'a set" | |
| 307 | assumes "\<And>A. countable A \<Longrightarrow> countable (f A)" | |
| 308 | and "countable A" | |
| 309 | shows "countable ((f ^^ n) A)" | |
| 310 | by(induction n)(simp_all add: assms) | |
| 311 | ||
| 54410 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 312 | lemma countable_rtrancl: | 
| 67613 | 313 | "(\<And>Y. countable Y \<Longrightarrow> countable (X `` Y)) \<Longrightarrow> countable Y \<Longrightarrow> countable (X\<^sup>* `` Y)" | 
| 54410 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 314 | unfolding rtrancl_is_UN_relpow UN_Image by (intro countable_UN countableI_type countable_relpow) | 
| 
0a578fb7fb73
countability of the image of a reflexive transitive closure
 hoelzl parents: 
53381diff
changeset | 315 | |
| 50134 | 316 | lemma countable_lists[intro, simp]: | 
| 317 | assumes A: "countable A" shows "countable (lists A)" | |
| 318 | proof - | |
| 319 | have "countable (lists (range (from_nat_into A)))" | |
| 320 | by (auto simp: lists_image) | |
| 321 | with A show ?thesis | |
| 322 | by (auto dest: subset_range_from_nat_into countable_subset lists_mono) | |
| 323 | qed | |
| 324 | ||
| 50245 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 325 | lemma Collect_finite_eq_lists: "Collect finite = set ` lists UNIV" | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 326 | using finite_list by auto | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 327 | |
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 328 | lemma countable_Collect_finite: "countable (Collect (finite::'a::countable set\<Rightarrow>bool))" | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 329 | by (simp add: Collect_finite_eq_lists) | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 330 | |
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63649diff
changeset | 331 | lemma countable_int: "countable \<int>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63649diff
changeset | 332 | unfolding Ints_def by auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63649diff
changeset | 333 | |
| 50936 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 334 | lemma countable_rat: "countable \<rat>" | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 335 | unfolding Rats_def by auto | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 336 | |
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 337 | lemma Collect_finite_subset_eq_lists: "{A. finite A \<and> A \<subseteq> T} = set ` lists T"
 | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 338 | using finite_list by (auto simp: lists_eq_set) | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 339 | |
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 340 | lemma countable_Collect_finite_subset: | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 341 |   "countable T \<Longrightarrow> countable {A. finite A \<and> A \<subseteq> T}"
 | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 342 | unfolding Collect_finite_subset_eq_lists by auto | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 343 | |
| 71848 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 344 | lemma countable_Fpow: "countable S \<Longrightarrow> countable (Fpow S)" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 345 | using countable_Collect_finite_subset | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 346 | by (force simp add: Fpow_def conj_commute) | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 347 | |
| 60058 | 348 | lemma countable_set_option [simp]: "countable (set_option x)" | 
| 71848 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 349 | by (cases x) auto | 
| 60058 | 350 | |
| 60500 | 351 | subsection \<open>Misc lemmas\<close> | 
| 50134 | 352 | |
| 63301 | 353 | lemma countable_subset_image: | 
| 354 | "countable B \<and> B \<subseteq> (f ` A) \<longleftrightarrow> (\<exists>A'. countable A' \<and> A' \<subseteq> A \<and> (B = f ` A'))" | |
| 355 | (is "?lhs = ?rhs") | |
| 356 | proof | |
| 357 | assume ?lhs | |
| 63649 | 358 | show ?rhs | 
| 359 | by (rule exI [where x="inv_into A f ` B"]) | |
| 360 | (use \<open>?lhs\<close> in \<open>auto simp: f_inv_into_f subset_iff image_inv_into_cancel inv_into_into\<close>) | |
| 63301 | 361 | next | 
| 362 | assume ?rhs | |
| 363 | then show ?lhs by force | |
| 364 | qed | |
| 365 | ||
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 366 | lemma ex_subset_image_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 367 | "(\<exists>T. T \<subseteq> f ` S \<and> P T) \<longleftrightarrow> (\<exists>T. T \<subseteq> S \<and> inj_on f T \<and> P (f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 368 | by (auto simp: subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 369 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 370 | lemma all_subset_image_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 371 | "(\<forall>T. T \<subseteq> f ` S \<longrightarrow> P T) \<longleftrightarrow> (\<forall>T. T \<subseteq> S \<and> inj_on f T \<longrightarrow> P(f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 372 | by (metis subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 373 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 374 | lemma ex_countable_subset_image_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 375 | "(\<exists>T. countable T \<and> T \<subseteq> f ` S \<and> P T) \<longleftrightarrow> | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 376 | (\<exists>T. countable T \<and> T \<subseteq> S \<and> inj_on f T \<and> P (f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 377 | by (metis countable_image_inj_eq subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 378 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 379 | lemma all_countable_subset_image_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 380 | "(\<forall>T. countable T \<and> T \<subseteq> f ` S \<longrightarrow> P T) \<longleftrightarrow> (\<forall>T. countable T \<and> T \<subseteq> S \<and> inj_on f T \<longrightarrow>P(f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 381 | by (metis countable_image_inj_eq subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 382 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 383 | lemma ex_countable_subset_image: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 384 | "(\<exists>T. countable T \<and> T \<subseteq> f ` S \<and> P T) \<longleftrightarrow> (\<exists>T. countable T \<and> T \<subseteq> S \<and> P (f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 385 | by (metis countable_subset_image) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 386 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 387 | lemma all_countable_subset_image: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 388 | "(\<forall>T. countable T \<and> T \<subseteq> f ` S \<longrightarrow> P T) \<longleftrightarrow> (\<forall>T. countable T \<and> T \<subseteq> S \<longrightarrow> P(f ` T))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 389 | by (metis countable_subset_image) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 390 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 391 | lemma countable_image_eq: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 392 | "countable(f ` S) \<longleftrightarrow> (\<exists>T. countable T \<and> T \<subseteq> S \<and> f ` S = f ` T)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 393 | by (metis countable_image countable_image_inj_eq order_refl subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 394 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 395 | lemma countable_image_eq_inj: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 396 | "countable(f ` S) \<longleftrightarrow> (\<exists>T. countable T \<and> T \<subseteq> S \<and> f ` S = f ` T \<and> inj_on f T)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 397 | by (metis countable_image_inj_eq order_refl subset_image_inj) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69661diff
changeset | 398 | |
| 62648 | 399 | lemma infinite_countable_subset': | 
| 400 | assumes X: "infinite X" shows "\<exists>C\<subseteq>X. countable C \<and> infinite C" | |
| 401 | proof - | |
| 74325 | 402 | obtain f :: "nat \<Rightarrow> 'a" where "inj f" "range f \<subseteq> X" | 
| 403 | using infinite_countable_subset [OF X] by blast | |
| 62648 | 404 | then show ?thesis | 
| 405 | by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite) | |
| 406 | qed | |
| 407 | ||
| 50134 | 408 | lemma countable_all: | 
| 409 | assumes S: "countable S" | |
| 410 | shows "(\<forall>s\<in>S. P s) \<longleftrightarrow> (\<forall>n::nat. from_nat_into S n \<in> S \<longrightarrow> P (from_nat_into S n))" | |
| 411 | using S[THEN subset_range_from_nat_into] by auto | |
| 412 | ||
| 57025 | 413 | lemma finite_sequence_to_countable_set: | 
| 69661 | 414 | assumes "countable X" | 
| 415 | obtains F where "\<And>i. F i \<subseteq> X" "\<And>i. F i \<subseteq> F (Suc i)" "\<And>i. finite (F i)" "(\<Union>i. F i) = X" | |
| 416 | proof - | |
| 417 | show thesis | |
| 57025 | 418 |     apply (rule that[of "\<lambda>i. if X = {} then {} else from_nat_into X ` {..i}"])
 | 
| 69661 | 419 | apply (auto simp add: image_iff intro: from_nat_into split: if_splits) | 
| 420 | using assms from_nat_into_surj by (fastforce cong: image_cong) | |
| 57025 | 421 | qed | 
| 422 | ||
| 62370 | 423 | lemma transfer_countable[transfer_rule]: | 
| 67399 | 424 | "bi_unique R \<Longrightarrow> rel_fun (rel_set R) (=) countable countable" | 
| 62370 | 425 | by (rule rel_funI, erule (1) bi_unique_rel_set_lemma) | 
| 426 | (auto dest: countable_image_inj_on) | |
| 427 | ||
| 60500 | 428 | subsection \<open>Uncountable\<close> | 
| 57234 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 429 | |
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 430 | abbreviation uncountable where | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 431 | "uncountable A \<equiv> \<not> countable A" | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 432 | |
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 433 | lemma uncountable_def: "uncountable A \<longleftrightarrow> A \<noteq> {} \<and> \<not> (\<exists>f::(nat \<Rightarrow> 'a). range f = A)"
 | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 434 | by (auto intro: inj_on_inv_into simp: countable_def) | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 435 | (metis all_not_in_conv inj_on_iff_surj subset_UNIV) | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 436 | |
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 437 | lemma uncountable_bij_betw: "bij_betw f A B \<Longrightarrow> uncountable B \<Longrightarrow> uncountable A" | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 438 | unfolding bij_betw_def by (metis countable_image) | 
| 62370 | 439 | |
| 57234 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 440 | lemma uncountable_infinite: "uncountable A \<Longrightarrow> infinite A" | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 441 | by (metis countable_finite) | 
| 62370 | 442 | |
| 57234 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 443 | lemma uncountable_minus_countable: | 
| 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 444 | "uncountable A \<Longrightarrow> countable B \<Longrightarrow> uncountable (A - B)" | 
| 63092 | 445 | using countable_Un[of B "A - B"] by auto | 
| 57234 
596a499318ab
clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57025diff
changeset | 446 | |
| 60303 | 447 | lemma countable_Diff_eq [simp]: "countable (A - {x}) = countable A"
 | 
| 448 | by (meson countable_Diff countable_empty countable_insert uncountable_minus_countable) | |
| 449 | ||
| 71848 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 450 | text \<open>Every infinite set can be covered by a pairwise disjoint family of infinite sets. | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 451 | This version doesn't achieve equality, as it only covers a countable subset\<close> | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 452 | lemma infinite_infinite_partition: | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 453 | assumes "infinite A" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 454 | obtains C :: "nat \<Rightarrow> 'a set" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 455 | where "pairwise (\<lambda>i j. disjnt (C i) (C j)) UNIV" "(\<Union>i. C i) \<subseteq> A" "\<And>i. infinite (C i)" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 456 | proof - | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 457 | obtain f :: "nat\<Rightarrow>'a" where "range f \<subseteq> A" "inj f" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 458 | using assms infinite_countable_subset by blast | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 459 | let ?C = "\<lambda>i. range (\<lambda>j. f (prod_encode (i,j)))" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 460 | show thesis | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 461 | proof | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 462 | show "pairwise (\<lambda>i j. disjnt (?C i) (?C j)) UNIV" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 463 | by (auto simp: pairwise_def disjnt_def inj_on_eq_iff [OF \<open>inj f\<close>] inj_on_eq_iff [OF inj_prod_encode, of _ UNIV]) | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 464 | show "(\<Union>i. ?C i) \<subseteq> A" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 465 | using \<open>range f \<subseteq> A\<close> by blast | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 466 | have "infinite (range (\<lambda>j. f (prod_encode (i, j))))" for i | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 467 | by (rule range_inj_infinite) (meson Pair_inject \<open>inj f\<close> inj_def prod_encode_eq) | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 468 | then show "\<And>i. infinite (?C i)" | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 469 | using that by auto | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 470 | qed | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 471 | qed | 
| 
3c7852327787
A few new theorems, plus some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 472 | |
| 50134 | 473 | end |