| author | wenzelm | 
| Wed, 03 Jul 2013 21:55:15 +0200 | |
| changeset 52514 | 8dd8ab81f1d7 | 
| parent 51542 | 738598beeb26 | 
| child 53381 | 355a4cac5440 | 
| permissions | -rw-r--r-- | 
| 50134 | 1 | (* Title: HOL/Library/Countable_Set.thy | 
| 2 | Author: Johannes Hölzl | |
| 3 | Author: Andrei Popescu | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Countable sets *}
 | |
| 7 | ||
| 8 | theory Countable_Set | |
| 51542 | 9 | imports Countable Infinite_Set | 
| 50134 | 10 | begin | 
| 11 | ||
| 12 | subsection {* Predicate for countable sets *}
 | |
| 13 | ||
| 14 | definition countable :: "'a set \<Rightarrow> bool" where | |
| 15 | "countable S \<longleftrightarrow> (\<exists>f::'a \<Rightarrow> nat. inj_on f S)" | |
| 16 | ||
| 17 | lemma countableE: | |
| 18 | assumes S: "countable S" obtains f :: "'a \<Rightarrow> nat" where "inj_on f S" | |
| 19 | using S by (auto simp: countable_def) | |
| 20 | ||
| 21 | lemma countableI: "inj_on (f::'a \<Rightarrow> nat) S \<Longrightarrow> countable S" | |
| 22 | by (auto simp: countable_def) | |
| 23 | ||
| 24 | lemma countableI': "inj_on (f::'a \<Rightarrow> 'b::countable) S \<Longrightarrow> countable S" | |
| 25 | using comp_inj_on[of f S to_nat] by (auto intro: countableI) | |
| 26 | ||
| 27 | lemma countableE_bij: | |
| 28 | assumes S: "countable S" obtains f :: "nat \<Rightarrow> 'a" and C :: "nat set" where "bij_betw f C S" | |
| 29 | using S by (blast elim: countableE dest: inj_on_imp_bij_betw bij_betw_inv) | |
| 30 | ||
| 31 | lemma countableI_bij: "bij_betw f (C::nat set) S \<Longrightarrow> countable S" | |
| 32 | by (blast intro: countableI bij_betw_inv_into bij_betw_imp_inj_on) | |
| 33 | ||
| 34 | lemma countable_finite: "finite S \<Longrightarrow> countable S" | |
| 35 | by (blast dest: finite_imp_inj_to_nat_seg countableI) | |
| 36 | ||
| 37 | lemma countableI_bij1: "bij_betw f A B \<Longrightarrow> countable A \<Longrightarrow> countable B" | |
| 38 | by (blast elim: countableE_bij intro: bij_betw_trans countableI_bij) | |
| 39 | ||
| 40 | lemma countableI_bij2: "bij_betw f B A \<Longrightarrow> countable A \<Longrightarrow> countable B" | |
| 41 | by (blast elim: countableE_bij intro: bij_betw_trans bij_betw_inv_into countableI_bij) | |
| 42 | ||
| 43 | lemma countable_iff_bij[simp]: "bij_betw f A B \<Longrightarrow> countable A \<longleftrightarrow> countable B" | |
| 44 | by (blast intro: countableI_bij1 countableI_bij2) | |
| 45 | ||
| 46 | lemma countable_subset: "A \<subseteq> B \<Longrightarrow> countable B \<Longrightarrow> countable A" | |
| 47 | by (auto simp: countable_def intro: subset_inj_on) | |
| 48 | ||
| 49 | lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)" | |
| 50 | using countableI[of to_nat A] by auto | |
| 51 | ||
| 52 | subsection {* Enumerate a countable set *}
 | |
| 53 | ||
| 54 | lemma countableE_infinite: | |
| 55 | assumes "countable S" "infinite S" | |
| 56 | obtains e :: "'a \<Rightarrow> nat" where "bij_betw e S UNIV" | |
| 57 | proof - | |
| 58 | from `countable S`[THEN countableE] guess f . | |
| 59 | then have "bij_betw f S (f`S)" | |
| 60 | unfolding bij_betw_def by simp | |
| 61 | moreover | |
| 62 | from `inj_on f S` `infinite S` have inf_fS: "infinite (f`S)" | |
| 63 | by (auto dest: finite_imageD) | |
| 64 | then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV" | |
| 65 | by (intro bij_betw_the_inv_into bij_enumerate) | |
| 66 | ultimately have "bij_betw (the_inv_into UNIV (enumerate (f`S)) \<circ> f) S UNIV" | |
| 67 | by (rule bij_betw_trans) | |
| 68 | then show thesis .. | |
| 69 | qed | |
| 70 | ||
| 71 | lemma countable_enum_cases: | |
| 72 | assumes "countable S" | |
| 73 |   obtains (finite) f :: "'a \<Rightarrow> nat" where "finite S" "bij_betw f S {..<card S}"
 | |
| 74 | | (infinite) f :: "'a \<Rightarrow> nat" where "infinite S" "bij_betw f S UNIV" | |
| 75 | using ex_bij_betw_finite_nat[of S] countableE_infinite `countable S` | |
| 76 | by (cases "finite S") (auto simp add: atLeast0LessThan) | |
| 77 | ||
| 78 | definition to_nat_on :: "'a set \<Rightarrow> 'a \<Rightarrow> nat" where | |
| 79 |   "to_nat_on S = (SOME f. if finite S then bij_betw f S {..< card S} else bij_betw f S UNIV)"
 | |
| 80 | ||
| 81 | 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 | 82 | "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 | 83 | |
| 84 | lemma to_nat_on_finite: "finite S \<Longrightarrow> bij_betw (to_nat_on S) S {..< card S}"
 | |
| 85 | using ex_bij_betw_finite_nat unfolding to_nat_on_def | |
| 86 |   by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S {..<card S}"]) (auto simp add: atLeast0LessThan)
 | |
| 87 | ||
| 88 | lemma to_nat_on_infinite: "countable S \<Longrightarrow> infinite S \<Longrightarrow> bij_betw (to_nat_on S) S UNIV" | |
| 89 | using countableE_infinite unfolding to_nat_on_def | |
| 90 | by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S UNIV"]) auto | |
| 91 | ||
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 92 | 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 | 93 | 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 | 94 | 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 | 95 | apply (subst bij_betw_cong) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 96 | apply (split split_if) | 
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 97 | 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 | 98 | 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 | 99 | 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 | 100 | done | 
| 50134 | 101 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 102 | 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 | 103 | 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 | 104 | 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 | 105 | by (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_infinite) | 
| 50134 | 106 | |
| 107 | lemma inj_on_to_nat_on[intro]: "countable A \<Longrightarrow> inj_on (to_nat_on A) A" | |
| 108 | using to_nat_on_infinite[of A] to_nat_on_finite[of A] | |
| 109 | by (cases "finite A") (auto simp: bij_betw_def) | |
| 110 | ||
| 111 | lemma to_nat_on_inj[simp]: | |
| 112 | "countable A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> to_nat_on A a = to_nat_on A b \<longleftrightarrow> a = b" | |
| 113 | using inj_on_to_nat_on[of A] by (auto dest: inj_onD) | |
| 114 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 115 | 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 | 116 | by (auto simp: from_nat_into_def intro!: inv_into_f_f) | 
| 117 | ||
| 118 | lemma subset_range_from_nat_into: "countable A \<Longrightarrow> A \<subseteq> range (from_nat_into A)" | |
| 119 | by (auto intro: from_nat_into_to_nat_on[symmetric]) | |
| 120 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 121 | 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 | 122 | 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 | 123 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 124 | 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 | 125 | 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 | 126 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 127 | 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 | 128 | 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 | 129 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 130 | 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 | 131 | 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 | 132 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 133 | 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 | 134 | 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 | 135 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 136 | 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 | 137 | 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 | 138 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 139 | 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 | 140 | "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 | 141 | 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 | 142 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 143 | lemma from_nat_into_inj: | 
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 144 | "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 | 145 | 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 | 146 | by (subst to_nat_on_inj[symmetric, of A]) auto | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 147 | |
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 148 | lemma from_nat_into_inj_infinite[simp]: | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 149 | "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 | 150 | 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 | 151 | |
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 152 | lemma eq_from_nat_into_iff: | 
| 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 153 | "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 | 154 | by auto | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 155 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 156 | 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 | 157 | 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 | 158 | |
| 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 159 | lemma from_nat_into_inject[simp]: | 
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 160 |   "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 | 161 | 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 | 162 | |
| 50148 
b8cff6a8fda2
Countable_Set: tuned lemma names; more generic lemmas
 hoelzl parents: 
50144diff
changeset | 163 | 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 | 164 | 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 | 165 | |
| 50134 | 166 | subsection {* Closure properties of countability *}
 | 
| 167 | ||
| 168 | lemma countable_SIGMA[intro, simp]: | |
| 169 | "countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (A i)) \<Longrightarrow> countable (SIGMA i : I. A i)" | |
| 170 | by (intro countableI'[of "\<lambda>(i, a). (to_nat_on I i, to_nat_on (A i) a)"]) (auto simp: inj_on_def) | |
| 171 | ||
| 172 | lemma countable_image[intro, simp]: assumes A: "countable A" shows "countable (f`A)" | |
| 173 | proof - | |
| 174 | from A guess g by (rule countableE) | |
| 175 | moreover have "inj_on (inv_into A f) (f`A)" "inv_into A f ` f ` A \<subseteq> A" | |
| 176 | by (auto intro: inj_on_inv_into inv_into_into) | |
| 177 | ultimately show ?thesis | |
| 178 | by (blast dest: comp_inj_on subset_inj_on intro: countableI) | |
| 179 | qed | |
| 180 | ||
| 181 | lemma countable_UN[intro, simp]: | |
| 182 | fixes I :: "'i set" and A :: "'i => 'a set" | |
| 183 | assumes I: "countable I" | |
| 184 | assumes A: "\<And>i. i \<in> I \<Longrightarrow> countable (A i)" | |
| 185 | shows "countable (\<Union>i\<in>I. A i)" | |
| 186 | proof - | |
| 187 | have "(\<Union>i\<in>I. A i) = snd ` (SIGMA i : I. A i)" by (auto simp: image_iff) | |
| 188 | then show ?thesis by (simp add: assms) | |
| 189 | qed | |
| 190 | ||
| 191 | lemma countable_Un[intro]: "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A \<union> B)" | |
| 192 |   by (rule countable_UN[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B", simplified])
 | |
| 193 | (simp split: bool.split) | |
| 194 | ||
| 195 | lemma countable_Un_iff[simp]: "countable (A \<union> B) \<longleftrightarrow> countable A \<and> countable B" | |
| 196 | by (metis countable_Un countable_subset inf_sup_ord(3,4)) | |
| 197 | ||
| 198 | lemma countable_Plus[intro, simp]: | |
| 199 | "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A <+> B)" | |
| 200 | by (simp add: Plus_def) | |
| 201 | ||
| 202 | lemma countable_empty[intro, simp]: "countable {}"
 | |
| 203 | by (blast intro: countable_finite) | |
| 204 | ||
| 205 | lemma countable_insert[intro, simp]: "countable A \<Longrightarrow> countable (insert a A)" | |
| 206 |   using countable_Un[of "{a}" A] by (auto simp: countable_finite)
 | |
| 207 | ||
| 208 | lemma countable_Int1[intro, simp]: "countable A \<Longrightarrow> countable (A \<inter> B)" | |
| 209 | by (force intro: countable_subset) | |
| 210 | ||
| 211 | lemma countable_Int2[intro, simp]: "countable B \<Longrightarrow> countable (A \<inter> B)" | |
| 212 | by (blast intro: countable_subset) | |
| 213 | ||
| 214 | lemma countable_INT[intro, simp]: "i \<in> I \<Longrightarrow> countable (A i) \<Longrightarrow> countable (\<Inter>i\<in>I. A i)" | |
| 215 | by (blast intro: countable_subset) | |
| 216 | ||
| 217 | lemma countable_Diff[intro, simp]: "countable A \<Longrightarrow> countable (A - B)" | |
| 218 | by (blast intro: countable_subset) | |
| 219 | ||
| 220 | lemma countable_vimage: "B \<subseteq> range f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B" | |
| 221 | by (metis Int_absorb2 assms countable_image image_vimage_eq) | |
| 222 | ||
| 223 | lemma surj_countable_vimage: "surj f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B" | |
| 224 | by (metis countable_vimage top_greatest) | |
| 225 | ||
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50134diff
changeset | 226 | 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 | 227 | 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 | 228 | |
| 50134 | 229 | lemma countable_lists[intro, simp]: | 
| 230 | assumes A: "countable A" shows "countable (lists A)" | |
| 231 | proof - | |
| 232 | have "countable (lists (range (from_nat_into A)))" | |
| 233 | by (auto simp: lists_image) | |
| 234 | with A show ?thesis | |
| 235 | by (auto dest: subset_range_from_nat_into countable_subset lists_mono) | |
| 236 | qed | |
| 237 | ||
| 50245 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 238 | lemma Collect_finite_eq_lists: "Collect finite = set ` lists UNIV" | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 239 | using finite_list by auto | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 240 | |
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 241 | lemma countable_Collect_finite: "countable (Collect (finite::'a::countable set\<Rightarrow>bool))" | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 242 | by (simp add: Collect_finite_eq_lists) | 
| 
dea9363887a6
based countable topological basis on Countable_Set
 immler parents: 
50148diff
changeset | 243 | |
| 50936 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 244 | lemma countable_rat: "countable \<rat>" | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 245 | unfolding Rats_def by auto | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 246 | |
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 247 | 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 | 248 | using finite_list by (auto simp: lists_eq_set) | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 249 | |
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 250 | lemma countable_Collect_finite_subset: | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 251 |   "countable T \<Longrightarrow> countable {A. finite A \<and> A \<subseteq> T}"
 | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 252 | unfolding Collect_finite_subset_eq_lists by auto | 
| 
b28f258ebc1a
countablility of finite subsets and rational numbers
 hoelzl parents: 
50245diff
changeset | 253 | |
| 50134 | 254 | subsection {* Misc lemmas *}
 | 
| 255 | ||
| 256 | lemma countable_all: | |
| 257 | assumes S: "countable S" | |
| 258 | 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))" | |
| 259 | using S[THEN subset_range_from_nat_into] by auto | |
| 260 | ||
| 261 | end |