src/HOL/Library/Countable.thy
author haftmann
Mon, 05 Jul 2010 15:12:20 +0200
changeset 37712 44b27ea94a16
parent 37678 0040bafffdef
child 39428 f967a16dfcdd
permissions -rw-r--r--
tuned proof
haftmann@26169
     1
(*  Title:      HOL/Library/Countable.thy
haftmann@26350
     2
    Author:     Alexander Krauss, TU Muenchen
haftmann@26169
     3
*)
haftmann@26169
     4
haftmann@26169
     5
header {* Encoding (almost) everything into natural numbers *}
haftmann@26169
     6
haftmann@26169
     7
theory Countable
huffman@35700
     8
imports Main Rat Nat_Bijection
haftmann@26169
     9
begin
haftmann@26169
    10
haftmann@26169
    11
subsection {* The class of countable types *}
haftmann@26169
    12
haftmann@29734
    13
class countable =
haftmann@26169
    14
  assumes ex_inj: "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat"
haftmann@26169
    15
haftmann@26169
    16
lemma countable_classI:
haftmann@26169
    17
  fixes f :: "'a \<Rightarrow> nat"
haftmann@26169
    18
  assumes "\<And>x y. f x = f y \<Longrightarrow> x = y"
haftmann@26169
    19
  shows "OFCLASS('a, countable_class)"
haftmann@26169
    20
proof (intro_classes, rule exI)
haftmann@26169
    21
  show "inj f"
haftmann@26169
    22
    by (rule injI [OF assms]) assumption
haftmann@26169
    23
qed
haftmann@26169
    24
haftmann@26169
    25
huffman@26585
    26
subsection {* Conversion functions *}
haftmann@26169
    27
haftmann@26169
    28
definition to_nat :: "'a\<Colon>countable \<Rightarrow> nat" where
haftmann@26169
    29
  "to_nat = (SOME f. inj f)"
haftmann@26169
    30
haftmann@26169
    31
definition from_nat :: "nat \<Rightarrow> 'a\<Colon>countable" where
haftmann@26169
    32
  "from_nat = inv (to_nat \<Colon> 'a \<Rightarrow> nat)"
haftmann@26169
    33
haftmann@26169
    34
lemma inj_to_nat [simp]: "inj to_nat"
haftmann@26169
    35
  by (rule exE_some [OF ex_inj]) (simp add: to_nat_def)
haftmann@26169
    36
huffman@29847
    37
lemma surj_from_nat [simp]: "surj from_nat"
huffman@29847
    38
  unfolding from_nat_def by (simp add: inj_imp_surj_inv)
huffman@29847
    39
haftmann@26169
    40
lemma to_nat_split [simp]: "to_nat x = to_nat y \<longleftrightarrow> x = y"
haftmann@26169
    41
  using injD [OF inj_to_nat] by auto
haftmann@26169
    42
haftmann@26169
    43
lemma from_nat_to_nat [simp]:
haftmann@26169
    44
  "from_nat (to_nat x) = x"
haftmann@26169
    45
  by (simp add: from_nat_def)
haftmann@26169
    46
haftmann@26169
    47
haftmann@26169
    48
subsection {* Countable types *}
haftmann@26169
    49
haftmann@26169
    50
instance nat :: countable
huffman@35700
    51
  by (rule countable_classI [of "id"]) simp
haftmann@26169
    52
haftmann@26169
    53
subclass (in finite) countable
haftmann@28823
    54
proof
haftmann@26169
    55
  have "finite (UNIV\<Colon>'a set)" by (rule finite_UNIV)
nipkow@31992
    56
  with finite_conv_nat_seg_image [of "UNIV::'a set"]
haftmann@26169
    57
  obtain n and f :: "nat \<Rightarrow> 'a" 
haftmann@26169
    58
    where "UNIV = f ` {i. i < n}" by auto
haftmann@26169
    59
  then have "surj f" unfolding surj_def by auto
haftmann@26169
    60
  then have "inj (inv f)" by (rule surj_imp_inj_inv)
haftmann@26169
    61
  then show "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat" by (rule exI[of inj])
haftmann@26169
    62
qed
haftmann@26169
    63
haftmann@26169
    64
text {* Pairs *}
haftmann@26169
    65
haftmann@37678
    66
instance prod :: (countable, countable) countable
huffman@35700
    67
  by (rule countable_classI [of "\<lambda>(x, y). prod_encode (to_nat x, to_nat y)"])
huffman@35700
    68
    (auto simp add: prod_encode_eq)
haftmann@26169
    69
haftmann@26169
    70
haftmann@26169
    71
text {* Sums *}
haftmann@26169
    72
haftmann@37678
    73
instance sum :: (countable, countable) countable
haftmann@26169
    74
  by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a)
haftmann@26169
    75
                                     | Inr b \<Rightarrow> to_nat (True, to_nat b))"])
huffman@35700
    76
    (simp split: sum.split_asm)
haftmann@26169
    77
haftmann@26169
    78
haftmann@26169
    79
text {* Integers *}
haftmann@26169
    80
haftmann@26169
    81
instance int :: countable
huffman@35700
    82
  by (rule countable_classI [of "int_encode"])
huffman@35700
    83
    (simp add: int_encode_eq)
haftmann@26169
    84
haftmann@26169
    85
haftmann@26169
    86
text {* Options *}
haftmann@26169
    87
haftmann@26169
    88
instance option :: (countable) countable
huffman@35700
    89
  by (rule countable_classI [of "option_case 0 (Suc \<circ> to_nat)"])
huffman@35700
    90
    (simp split: option.split_asm)
haftmann@26169
    91
haftmann@26169
    92
haftmann@26169
    93
text {* Lists *}
haftmann@26169
    94
haftmann@26169
    95
instance list :: (countable) countable
huffman@35700
    96
  by (rule countable_classI [of "list_encode \<circ> map to_nat"])
huffman@35700
    97
    (simp add: list_encode_eq)
haftmann@26169
    98
huffman@26243
    99
haftmann@37652
   100
text {* Further *}
haftmann@37652
   101
haftmann@37652
   102
instance String.literal :: countable
haftmann@37652
   103
  by (rule countable_classI [of "String.literal_case to_nat"])
haftmann@37652
   104
   (auto split: String.literal.splits)
haftmann@37652
   105
haftmann@37652
   106
instantiation typerep :: countable
haftmann@37652
   107
begin
haftmann@37652
   108
haftmann@37652
   109
fun to_nat_typerep :: "typerep \<Rightarrow> nat" where
haftmann@37652
   110
  "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))"
haftmann@37652
   111
haftmann@37652
   112
instance proof (rule countable_classI)
haftmann@37712
   113
  fix t t' :: typerep and ts ts' :: "typerep list"
haftmann@37712
   114
  assume "to_nat_typerep t = to_nat_typerep t'"
haftmann@37712
   115
  moreover have "to_nat_typerep t = to_nat_typerep t' \<Longrightarrow> t = t'"
haftmann@37712
   116
    and "map to_nat_typerep ts = map to_nat_typerep ts' \<Longrightarrow> ts = ts'"
haftmann@37712
   117
  proof (induct t and ts arbitrary: t' and ts' rule: typerep.inducts)
haftmann@37712
   118
    case (Typerep c ts t')
haftmann@37712
   119
    then obtain c' ts' where t': "t' = Typerep.Typerep c' ts'" by (cases t') auto
haftmann@37712
   120
    with Typerep have "c = c'" and "ts = ts'" by simp_all
haftmann@37712
   121
    with t' show "Typerep.Typerep c ts = t'" by simp
haftmann@37652
   122
  next
haftmann@37652
   123
    case Nil_typerep then show ?case by simp
haftmann@37652
   124
  next
haftmann@37652
   125
    case (Cons_typerep t ts) then show ?case by auto
haftmann@37652
   126
  qed
haftmann@37652
   127
  ultimately show "t = t'" by simp
haftmann@37652
   128
qed
haftmann@37652
   129
haftmann@37652
   130
end
haftmann@37652
   131
haftmann@37652
   132
huffman@26243
   133
text {* Functions *}
huffman@26243
   134
huffman@26243
   135
instance "fun" :: (finite, countable) countable
huffman@26243
   136
proof
huffman@26243
   137
  obtain xs :: "'a list" where xs: "set xs = UNIV"
huffman@26243
   138
    using finite_list [OF finite_UNIV] ..
huffman@26243
   139
  show "\<exists>to_nat::('a \<Rightarrow> 'b) \<Rightarrow> nat. inj to_nat"
huffman@26243
   140
  proof
huffman@26243
   141
    show "inj (\<lambda>f. to_nat (map f xs))"
huffman@26243
   142
      by (rule injI, simp add: xs expand_fun_eq)
huffman@26243
   143
  qed
huffman@26243
   144
qed
huffman@26243
   145
huffman@29817
   146
huffman@29817
   147
subsection {* The Rationals are Countably Infinite *}
huffman@29817
   148
huffman@29817
   149
definition nat_to_rat_surj :: "nat \<Rightarrow> rat" where
huffman@35700
   150
"nat_to_rat_surj n = (let (a,b) = prod_decode n
huffman@35700
   151
                      in Fract (int_decode a) (int_decode b))"
huffman@29817
   152
huffman@29817
   153
lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj"
huffman@29817
   154
unfolding surj_def
huffman@29817
   155
proof
huffman@29817
   156
  fix r::rat
huffman@29817
   157
  show "\<exists>n. r = nat_to_rat_surj n"
haftmann@35374
   158
  proof (cases r)
haftmann@35374
   159
    fix i j assume [simp]: "r = Fract i j" and "j > 0"
huffman@35700
   160
    have "r = (let m = int_encode i; n = int_encode j
huffman@35700
   161
               in nat_to_rat_surj(prod_encode (m,n)))"
huffman@35700
   162
      by (simp add: Let_def nat_to_rat_surj_def)
huffman@29817
   163
    thus "\<exists>n. r = nat_to_rat_surj n" by(auto simp:Let_def)
huffman@29817
   164
  qed
huffman@29817
   165
qed
huffman@29817
   166
huffman@29817
   167
lemma Rats_eq_range_nat_to_rat_surj: "\<rat> = range nat_to_rat_surj"
huffman@29817
   168
by (simp add: Rats_def surj_nat_to_rat_surj surj_range)
huffman@29817
   169
huffman@29817
   170
context field_char_0
huffman@29817
   171
begin
huffman@29817
   172
huffman@29817
   173
lemma Rats_eq_range_of_rat_o_nat_to_rat_surj:
huffman@29817
   174
  "\<rat> = range (of_rat o nat_to_rat_surj)"
huffman@29817
   175
using surj_nat_to_rat_surj
huffman@29817
   176
by (auto simp: Rats_def image_def surj_def)
huffman@29817
   177
   (blast intro: arg_cong[where f = of_rat])
huffman@29817
   178
huffman@29817
   179
lemma surj_of_rat_nat_to_rat_surj:
huffman@29817
   180
  "r\<in>\<rat> \<Longrightarrow> \<exists>n. r = of_rat(nat_to_rat_surj n)"
huffman@29817
   181
by(simp add: Rats_eq_range_of_rat_o_nat_to_rat_surj image_def)
huffman@29817
   182
haftmann@26169
   183
end
huffman@29817
   184
huffman@29817
   185
instance rat :: countable
huffman@29817
   186
proof
huffman@29817
   187
  show "\<exists>to_nat::rat \<Rightarrow> nat. inj to_nat"
huffman@29817
   188
  proof
huffman@29817
   189
    have "surj nat_to_rat_surj"
huffman@29817
   190
      by (rule surj_nat_to_rat_surj)
huffman@29817
   191
    then show "inj (inv nat_to_rat_surj)"
huffman@29817
   192
      by (rule surj_imp_inj_inv)
huffman@29817
   193
  qed
huffman@29817
   194
qed
huffman@29817
   195
huffman@29817
   196
end