src/HOL/Library/Countable_Set.thy
author Fabian Huch <huch@in.tum.de>
Thu, 30 Nov 2023 13:44:51 +0100
changeset 79084 dd689c4ab688
parent 78200 264f2b69d09c
permissions -rw-r--r--
consistent hosts ordering;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Library/Countable_Set.thy
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     3
    Author:     Andrei Popescu
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     4
*)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     5
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
     6
section \<open>Countable sets\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     7
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
     8
theory Countable_Set
51542
738598beeb26 tuned imports;
wenzelm
parents: 50936
diff changeset
     9
imports Countable Infinite_Set
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    10
begin
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    11
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    12
subsection \<open>Predicate for countable sets\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    13
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    14
definition countable :: "'a set \<Rightarrow> bool" where
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    15
  "countable S \<longleftrightarrow> (\<exists>f::'a \<Rightarrow> nat. inj_on f S)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    16
78200
264f2b69d09c New and generalised analysis lemmas
paulson <lp15@cam.ac.uk>
parents: 77138
diff 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: 77138
diff 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: 77138
diff changeset
    19
   
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    20
lemma countableE:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    21
  assumes S: "countable S" obtains f :: "'a \<Rightarrow> nat" where "inj_on f S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    22
  using S by (auto simp: countable_def)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    23
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    24
lemma countableI: "inj_on (f::'a \<Rightarrow> nat) S \<Longrightarrow> countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    25
  by (auto simp: countable_def)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    26
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    27
lemma countableI': "inj_on (f::'a \<Rightarrow> 'b::countable) S \<Longrightarrow> countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    28
  using comp_inj_on[of f S to_nat] by (auto intro: countableI)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    29
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    30
lemma countableE_bij:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    31
  assumes S: "countable S" obtains f :: "nat \<Rightarrow> 'a" and C :: "nat set" where "bij_betw f C S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    32
  using S by (blast elim: countableE dest: inj_on_imp_bij_betw bij_betw_inv)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    33
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    34
lemma countableI_bij: "bij_betw f (C::nat set) S \<Longrightarrow> countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    35
  by (blast intro: countableI bij_betw_inv_into bij_betw_imp_inj_on)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    36
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    37
lemma countable_finite: "finite S \<Longrightarrow> countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    38
  by (blast dest: finite_imp_inj_to_nat_seg countableI)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    39
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    40
lemma countableI_bij1: "bij_betw f A B \<Longrightarrow> countable A \<Longrightarrow> countable B"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    41
  by (blast elim: countableE_bij intro: bij_betw_trans countableI_bij)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    42
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    43
lemma countableI_bij2: "bij_betw f B A \<Longrightarrow> countable A \<Longrightarrow> countable B"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    44
  by (blast elim: countableE_bij intro: bij_betw_trans bij_betw_inv_into countableI_bij)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    45
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    46
lemma countable_iff_bij[simp]: "bij_betw f A B \<Longrightarrow> countable A \<longleftrightarrow> countable B"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    47
  by (blast intro: countableI_bij1 countableI_bij2)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    48
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    49
lemma countable_subset: "A \<subseteq> B \<Longrightarrow> countable B \<Longrightarrow> countable A"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    50
  by (auto simp: countable_def intro: subset_inj_on)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    51
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    52
lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    53
  using countableI[of to_nat A] by auto
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    54
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    55
subsection \<open>Enumerate a countable set\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    56
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    57
lemma countableE_infinite:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    58
  assumes "countable S" "infinite S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    59
  obtains e :: "'a \<Rightarrow> nat" where "bij_betw e S UNIV"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    60
proof -
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
    61
  obtain f :: "'a \<Rightarrow> nat" where "inj_on f S"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    62
    using \<open>countable S\<close> by (rule countableE)
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    63
  then have "bij_betw f S (f`S)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    64
    unfolding bij_betw_def by simp
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    65
  moreover
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    66
  from \<open>inj_on f S\<close> \<open>infinite S\<close> have inf_fS: "infinite (f`S)"
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    67
    by (auto dest: finite_imageD)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    68
  then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    69
    by (intro bij_betw_the_inv_into bij_enumerate)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    70
  ultimately have "bij_betw (the_inv_into UNIV (enumerate (f`S)) \<circ> f) S UNIV"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    71
    by (rule bij_betw_trans)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    72
  then show thesis ..
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    73
qed
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    74
68860
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 67613
diff changeset
    75
lemma countable_infiniteE':
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 67613
diff changeset
    76
  assumes "countable A" "infinite A"
f443ec10447d Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents: 67613
diff 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: 74438
diff 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: 67613
diff changeset
    79
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    80
lemma countable_enum_cases:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    81
  assumes "countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    82
  obtains (finite) f :: "'a \<Rightarrow> nat" where "finite S" "bij_betw f S {..<card S}"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    83
        | (infinite) f :: "'a \<Rightarrow> nat" where "infinite S" "bij_betw f S UNIV"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
    84
  using ex_bij_betw_finite_nat[of S] countableE_infinite \<open>countable S\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    85
  by (cases "finite S") (auto simp add: atLeast0LessThan)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    86
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    87
definition to_nat_on :: "'a set \<Rightarrow> 'a \<Rightarrow> nat" where
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    88
  "to_nat_on S = (SOME f. if finite S then bij_betw f S {..< card S} else bij_betw f S UNIV)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    89
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    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: 50134
diff 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
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    92
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    93
lemma to_nat_on_finite: "finite S \<Longrightarrow> bij_betw (to_nat_on S) S {..< card S}"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    94
  using ex_bij_betw_finite_nat unfolding to_nat_on_def
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    95
  by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S {..<card S}"]) (auto simp add: atLeast0LessThan)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    96
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    97
lemma to_nat_on_infinite: "countable S \<Longrightarrow> infinite S \<Longrightarrow> bij_betw (to_nat_on S) S UNIV"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    98
  using countableE_infinite unfolding to_nat_on_def
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
    99
  by (intro someI2_ex[where Q="\<lambda>f. bij_betw f S UNIV"]) auto
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   100
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   104
  apply (subst bij_betw_cong)
62390
842917225d56 more canonical names
nipkow
parents: 62370
diff changeset
   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: 50134
diff 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: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   109
  done
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   110
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   114
  by (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_infinite)
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   115
74438
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff 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: 74325
diff 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: 74325
diff changeset
   118
\<close>
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff 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: 74325
diff changeset
   120
begin
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff changeset
   121
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff 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: 74325
diff changeset
   132
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff changeset
   133
end
5827b91ef30e new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
paulson <lp15@cam.ac.uk>
parents: 74325
diff changeset
   134
63127
360d9997fac9 new theorem
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   135
lemma countable_as_injective_image:
360d9997fac9 new theorem
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   136
  assumes "countable A" "infinite A"
360d9997fac9 new theorem
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   137
  obtains f :: "nat \<Rightarrow> 'a" where "A = range f" "inj f"
360d9997fac9 new theorem
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   138
by (metis bij_betw_def bij_betw_from_nat_into [OF assms])
360d9997fac9 new theorem
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   139
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   140
lemma inj_on_to_nat_on[intro]: "countable A \<Longrightarrow> inj_on (to_nat_on A) A"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   141
  using to_nat_on_infinite[of A] to_nat_on_finite[of A]
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   142
  by (cases "finite A") (auto simp: bij_betw_def)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   143
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   144
lemma to_nat_on_inj[simp]:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   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"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   146
  using inj_on_to_nat_on[of A] by (auto dest: inj_onD)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   147
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   149
  by (auto simp: from_nat_into_def intro!: inv_into_f_f)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   150
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   151
lemma subset_range_from_nat_into: "countable A \<Longrightarrow> A \<subseteq> range (from_nat_into A)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   152
  by (auto intro: from_nat_into_to_nat_on[symmetric])
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   153
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   156
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   159
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50134
diff 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: 50134
diff changeset
   162
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   165
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   168
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   171
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   175
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff changeset
   176
lemma from_nat_into_inj:
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50144
diff 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: 50144
diff changeset
   179
  by (subst to_nat_on_inj[symmetric, of A]) auto
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff changeset
   180
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff changeset
   181
lemma from_nat_into_inj_infinite[simp]:
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50144
diff 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: 50144
diff changeset
   184
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff changeset
   185
lemma eq_from_nat_into_iff:
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50144
diff changeset
   187
  by auto
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff changeset
   188
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   191
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff changeset
   192
lemma from_nat_into_inject[simp]:
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50144
diff 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: 50134
diff changeset
   195
50148
b8cff6a8fda2 Countable_Set: tuned lemma names; more generic lemmas
hoelzl
parents: 50144
diff 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: 50144
diff 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: 50134
diff changeset
   198
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   199
subsection \<open>Closure properties of countability\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   200
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   201
lemma countable_SIGMA[intro, simp]:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   202
  "countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (A i)) \<Longrightarrow> countable (SIGMA i : I. A i)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   203
  by (intro countableI'[of "\<lambda>(i, a). (to_nat_on I i, to_nat_on (A i) a)"]) (auto simp: inj_on_def)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   204
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
   205
lemma countable_image[intro, simp]:
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
   206
  assumes "countable A"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
   207
  shows "countable (f`A)"
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   208
proof -
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
   209
  obtain g :: "'a \<Rightarrow> nat" where "inj_on g A"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 51542
diff changeset
   210
    using assms by (rule countableE)
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   211
  moreover have "inj_on (inv_into A f) (f`A)" "inv_into A f ` f ` A \<subseteq> A"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   212
    by (auto intro: inj_on_inv_into inv_into_into)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   213
  ultimately show ?thesis
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   214
    by (blast dest: comp_inj_on subset_inj_on intro: countableI)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   215
qed
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   216
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   217
lemma countable_image_inj_on: "countable (f ` A) \<Longrightarrow> inj_on f A \<Longrightarrow> countable A"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   218
  by (metis countable_image the_inv_into_onto)
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   219
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   220
lemma countable_image_inj_Int_vimage:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff 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: 69661
diff changeset
   223
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   224
lemma countable_image_inj_gen:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   226
  using countable_image_inj_Int_vimage
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   227
  by (auto simp: vimage_def Collect_conj_eq)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   228
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   229
lemma countable_image_inj_eq:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   231
  using countable_image_inj_on by blast
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   232
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   233
lemma countable_image_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff 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: 69661
diff changeset
   236
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   237
lemma countable_UN[intro, simp]:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   238
  fixes I :: "'i set" and A :: "'i => 'a set"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   239
  assumes I: "countable I"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   240
  assumes A: "\<And>i. i \<in> I \<Longrightarrow> countable (A i)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   241
  shows "countable (\<Union>i\<in>I. A i)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   242
proof -
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   243
  have "(\<Union>i\<in>I. A i) = snd ` (SIGMA i : I. A i)" by (auto simp: image_iff)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   244
  then show ?thesis by (simp add: assms)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   245
qed
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   246
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   247
lemma countable_Un[intro]: "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A \<union> B)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   248
  by (rule countable_UN[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B", simplified])
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   249
     (simp split: bool.split)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   250
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   251
lemma countable_Un_iff[simp]: "countable (A \<union> B) \<longleftrightarrow> countable A \<and> countable B"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   252
  by (metis countable_Un countable_subset inf_sup_ord(3,4))
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   253
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   254
lemma countable_Plus[intro, simp]:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   255
  "countable A \<Longrightarrow> countable B \<Longrightarrow> countable (A <+> B)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   256
  by (simp add: Plus_def)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   257
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   258
lemma countable_empty[intro, simp]: "countable {}"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   259
  by (blast intro: countable_finite)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   260
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   261
lemma countable_insert[intro, simp]: "countable A \<Longrightarrow> countable (insert a A)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   262
  using countable_Un[of "{a}" A] by (auto simp: countable_finite)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   263
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   264
lemma countable_Int1[intro, simp]: "countable A \<Longrightarrow> countable (A \<inter> B)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   265
  by (force intro: countable_subset)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   266
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   267
lemma countable_Int2[intro, simp]: "countable B \<Longrightarrow> countable (A \<inter> B)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   268
  by (blast intro: countable_subset)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   269
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   270
lemma countable_INT[intro, simp]: "i \<in> I \<Longrightarrow> countable (A i) \<Longrightarrow> countable (\<Inter>i\<in>I. A i)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   271
  by (blast intro: countable_subset)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   272
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   273
lemma countable_Diff[intro, simp]: "countable A \<Longrightarrow> countable (A - B)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   274
  by (blast intro: countable_subset)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   275
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   276
lemma countable_insert_eq [simp]: "countable (insert x A) = countable A"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   277
    by auto (metis Diff_insert_absorb countable_Diff insert_absorb)
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   278
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   279
lemma countable_vimage: "B \<subseteq> range f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 62648
diff changeset
   280
  by (metis Int_absorb2 countable_image image_vimage_eq)
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   281
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   282
lemma surj_countable_vimage: "surj f \<Longrightarrow> countable (f -` B) \<Longrightarrow> countable B"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   283
  by (metis countable_vimage top_greatest)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   284
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50134
diff 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: 50134
diff 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: 50134
diff changeset
   287
54410
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   288
lemma countable_Image:
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   289
  assumes "\<And>y. y \<in> Y \<Longrightarrow> countable (X `` {y})"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   290
  assumes "countable Y"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   291
  shows "countable (X `` Y)"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   292
proof -
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   293
  have "countable (X `` (\<Union>y\<in>Y. {y}))"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   294
    unfolding Image_UN by (intro countable_UN assms)
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   295
  then show ?thesis by simp
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   296
qed
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   297
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   298
lemma countable_relpow:
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   299
  fixes X :: "'a rel"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff 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: 53381
diff changeset
   301
  assumes Y: "countable Y"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   302
  shows "countable ((X ^^ i) `` Y)"
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff 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: 53381
diff changeset
   304
60058
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   305
lemma countable_funpow:
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   306
  fixes f :: "'a set \<Rightarrow> 'a set"
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   307
  assumes "\<And>A. countable A \<Longrightarrow> countable (f A)"
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   308
  and "countable A"
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   309
  shows "countable ((f ^^ n) A)"
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   310
by(induction n)(simp_all add: assms)
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   311
54410
0a578fb7fb73 countability of the image of a reflexive transitive closure
hoelzl
parents: 53381
diff changeset
   312
lemma countable_rtrancl:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   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: 53381
diff 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: 53381
diff changeset
   315
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   316
lemma countable_lists[intro, simp]:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   317
  assumes A: "countable A" shows "countable (lists A)"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   318
proof -
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   319
  have "countable (lists (range (from_nat_into A)))"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   320
    by (auto simp: lists_image)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   321
  with A show ?thesis
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   322
    by (auto dest: subset_range_from_nat_into countable_subset lists_mono)
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   323
qed
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   324
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   325
lemma Collect_finite_eq_lists: "Collect finite = set ` lists UNIV"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   326
  using finite_list by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   327
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   328
lemma countable_Collect_finite: "countable (Collect (finite::'a::countable set\<Rightarrow>bool))"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   329
  by (simp add: Collect_finite_eq_lists)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50148
diff changeset
   330
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63649
diff changeset
   331
lemma countable_int: "countable \<int>"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63649
diff changeset
   332
  unfolding Ints_def by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63649
diff changeset
   333
50936
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   334
lemma countable_rat: "countable \<rat>"
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   335
  unfolding Rats_def by auto
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   336
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff 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: 50245
diff changeset
   338
  using finite_list by (auto simp: lists_eq_set)
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   339
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   340
lemma countable_Collect_finite_subset:
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   341
  "countable T \<Longrightarrow> countable {A. finite A \<and> A \<subseteq> T}"
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   342
  unfolding Collect_finite_subset_eq_lists by auto
b28f258ebc1a countablility of finite subsets and rational numbers
hoelzl
parents: 50245
diff changeset
   343
71848
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff changeset
   345
  using countable_Collect_finite_subset
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff changeset
   347
60058
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   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: 70178
diff changeset
   349
  by (cases x) auto
60058
f17bb06599f6 add lemmas
Andreas Lochbihler
parents: 58881
diff changeset
   350
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   351
subsection \<open>Misc lemmas\<close>
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   352
63301
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   353
lemma countable_subset_image:
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   354
   "countable B \<and> B \<subseteq> (f ` A) \<longleftrightarrow> (\<exists>A'. countable A' \<and> A' \<subseteq> A \<and> (B = f ` A'))"
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   355
   (is "?lhs = ?rhs")
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   356
proof
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   357
  assume ?lhs
63649
e690d6f2185b tuned proofs;
wenzelm
parents: 63301
diff changeset
   358
  show ?rhs
e690d6f2185b tuned proofs;
wenzelm
parents: 63301
diff changeset
   359
    by (rule exI [where x="inv_into A f ` B"])
e690d6f2185b tuned proofs;
wenzelm
parents: 63301
diff changeset
   360
      (use \<open>?lhs\<close> in \<open>auto simp: f_inv_into_f subset_iff image_inv_into_cancel inv_into_into\<close>)
63301
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   361
next
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   362
  assume ?rhs
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   363
  then show ?lhs by force
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   364
qed
d3c87eb0bad2 new results about topology
paulson <lp15@cam.ac.uk>
parents: 63127
diff changeset
   365
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   366
lemma ex_subset_image_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   368
  by (auto simp: subset_image_inj)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   369
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   370
lemma all_subset_image_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   372
  by (metis subset_image_inj)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   373
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   374
lemma ex_countable_subset_image_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff 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: 69661
diff changeset
   377
  by (metis countable_image_inj_eq subset_image_inj)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   378
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   379
lemma all_countable_subset_image_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   381
  by (metis countable_image_inj_eq subset_image_inj)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   382
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   383
lemma ex_countable_subset_image:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   385
  by (metis countable_subset_image)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   386
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   387
lemma all_countable_subset_image:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff changeset
   389
  by (metis countable_subset_image)
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   390
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   391
lemma countable_image_eq:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff 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: 69661
diff changeset
   394
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
   395
lemma countable_image_eq_inj:
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 69661
diff 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: 69661
diff 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: 69661
diff changeset
   398
62648
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   399
lemma infinite_countable_subset':
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   400
  assumes X: "infinite X" shows "\<exists>C\<subseteq>X. countable C \<and> infinite C"
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   401
proof -
74325
8d0c2d74ad63 tuned proofs --- eliminated 'guess';
wenzelm
parents: 71848
diff changeset
   402
  obtain f :: "nat \<Rightarrow> 'a" where "inj f" "range f \<subseteq> X"
8d0c2d74ad63 tuned proofs --- eliminated 'guess';
wenzelm
parents: 71848
diff changeset
   403
    using infinite_countable_subset [OF X] by blast
62648
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   404
  then show ?thesis
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   405
    by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite)
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   406
qed
ee48e0b4f669 more stuff for extended nonnegative real numbers
hoelzl
parents: 62390
diff changeset
   407
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   408
lemma countable_all:
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   409
  assumes S: "countable S"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   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))"
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   411
  using S[THEN subset_range_from_nat_into] by auto
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   412
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 54410
diff changeset
   413
lemma finite_sequence_to_countable_set:
69661
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   414
  assumes "countable X"
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   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"
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   416
proof -
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   417
  show thesis
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 54410
diff changeset
   418
    apply (rule that[of "\<lambda>i. if X = {} then {} else from_nat_into X ` {..i}"])
69661
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   419
       apply (auto simp add: image_iff intro: from_nat_into split: if_splits)
a03a63b81f44 tuned proofs
haftmann
parents: 68860
diff changeset
   420
    using assms from_nat_into_surj by (fastforce cong: image_cong)
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 54410
diff changeset
   421
qed
e7fd64f82876 add various lemmas
hoelzl
parents: 54410
diff changeset
   422
62370
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   423
lemma transfer_countable[transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64008
diff changeset
   424
  "bi_unique R \<Longrightarrow> rel_fun (rel_set R) (=) countable countable"
62370
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   425
  by (rule rel_funI, erule (1) bi_unique_rel_set_lemma)
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   426
     (auto dest: countable_image_inj_on)
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   427
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60303
diff changeset
   428
subsection \<open>Uncountable\<close>
57234
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   429
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   430
abbreviation uncountable where
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   431
  "uncountable A \<equiv> \<not> countable A"
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   432
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff 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: 57025
diff 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: 57025
diff 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: 57025
diff changeset
   436
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff 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: 57025
diff changeset
   438
  unfolding bij_betw_def by (metis countable_image)
62370
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   439
57234
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   440
lemma uncountable_infinite: "uncountable A \<Longrightarrow> infinite A"
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   441
  by (metis countable_finite)
62370
4a35e3945cab add transfer rule for countable
hoelzl
parents: 60500
diff changeset
   442
57234
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   443
lemma uncountable_minus_countable:
596a499318ab clean up ContNotDenum; add lemmas by Jeremy Avigad and Luke Serafin
hoelzl
parents: 57025
diff changeset
   444
  "uncountable A \<Longrightarrow> countable B \<Longrightarrow> uncountable (A - B)"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 62648
diff changeset
   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: 57025
diff changeset
   446
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   447
lemma countable_Diff_eq [simp]: "countable (A - {x}) = countable A"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   448
  by (meson countable_Diff countable_empty countable_insert uncountable_minus_countable)
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60058
diff changeset
   449
71848
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff 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: 70178
diff changeset
   452
lemma infinite_infinite_partition:
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   453
  assumes "infinite A"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   454
  obtains C :: "nat \<Rightarrow> 'a set" 
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff changeset
   456
proof -
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff changeset
   458
    using assms infinite_countable_subset by blast
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff changeset
   460
  show thesis
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   461
  proof
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff 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: 70178
diff changeset
   464
    show "(\<Union>i. ?C i) \<subseteq> A"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff 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: 70178
diff 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: 70178
diff 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: 70178
diff changeset
   468
    then show "\<And>i. infinite (?C i)"
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   469
      using that by auto
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   470
  qed
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   471
qed
3c7852327787 A few new theorems, plus some tidying up
paulson <lp15@cam.ac.uk>
parents: 70178
diff changeset
   472
50134
13211e07d931 add Countable_Set theory
hoelzl
parents:
diff changeset
   473
end