haftmann@26169: (* Title: HOL/Library/Countable.thy haftmann@26350: Author: Alexander Krauss, TU Muenchen haftmann@26169: *) haftmann@26169: haftmann@26169: header {* Encoding (almost) everything into natural numbers *} haftmann@26169: haftmann@26169: theory Countable huffman@35700: imports Main Rat Nat_Bijection haftmann@26169: begin haftmann@26169: haftmann@26169: subsection {* The class of countable types *} haftmann@26169: haftmann@29797: class countable = haftmann@26169: assumes ex_inj: "\to_nat \ 'a \ nat. inj to_nat" haftmann@26169: haftmann@26169: lemma countable_classI: haftmann@26169: fixes f :: "'a \ nat" haftmann@26169: assumes "\x y. f x = f y \ x = y" haftmann@26169: shows "OFCLASS('a, countable_class)" haftmann@26169: proof (intro_classes, rule exI) haftmann@26169: show "inj f" haftmann@26169: by (rule injI [OF assms]) assumption haftmann@26169: qed haftmann@26169: haftmann@26169: huffman@26585: subsection {* Conversion functions *} haftmann@26169: haftmann@26169: definition to_nat :: "'a\countable \ nat" where haftmann@26169: "to_nat = (SOME f. inj f)" haftmann@26169: haftmann@26169: definition from_nat :: "nat \ 'a\countable" where haftmann@26169: "from_nat = inv (to_nat \ 'a \ nat)" haftmann@26169: haftmann@26169: lemma inj_to_nat [simp]: "inj to_nat" haftmann@26169: by (rule exE_some [OF ex_inj]) (simp add: to_nat_def) haftmann@26169: huffman@29910: lemma surj_from_nat [simp]: "surj from_nat" huffman@29910: unfolding from_nat_def by (simp add: inj_imp_surj_inv) huffman@29910: haftmann@26169: lemma to_nat_split [simp]: "to_nat x = to_nat y \ x = y" haftmann@26169: using injD [OF inj_to_nat] by auto haftmann@26169: haftmann@26169: lemma from_nat_to_nat [simp]: haftmann@26169: "from_nat (to_nat x) = x" haftmann@26169: by (simp add: from_nat_def) haftmann@26169: haftmann@26169: haftmann@26169: subsection {* Countable types *} haftmann@26169: haftmann@26169: instance nat :: countable huffman@35700: by (rule countable_classI [of "id"]) simp haftmann@26169: haftmann@26169: subclass (in finite) countable haftmann@28823: proof haftmann@26169: have "finite (UNIV\'a set)" by (rule finite_UNIV) nipkow@31992: with finite_conv_nat_seg_image [of "UNIV::'a set"] haftmann@26169: obtain n and f :: "nat \ 'a" haftmann@26169: where "UNIV = f ` {i. i < n}" by auto haftmann@26169: then have "surj f" unfolding surj_def by auto haftmann@26169: then have "inj (inv f)" by (rule surj_imp_inj_inv) haftmann@26169: then show "\to_nat \ 'a \ nat. inj to_nat" by (rule exI[of inj]) haftmann@26169: qed haftmann@26169: haftmann@26169: text {* Pairs *} haftmann@26169: haftmann@26169: instance "*" :: (countable, countable) countable huffman@35700: by (rule countable_classI [of "\(x, y). prod_encode (to_nat x, to_nat y)"]) huffman@35700: (auto simp add: prod_encode_eq) haftmann@26169: haftmann@26169: haftmann@26169: text {* Sums *} haftmann@26169: haftmann@26169: instance "+":: (countable, countable) countable haftmann@26169: by (rule countable_classI [of "(\x. case x of Inl a \ to_nat (False, to_nat a) haftmann@26169: | Inr b \ to_nat (True, to_nat b))"]) huffman@35700: (simp split: sum.split_asm) haftmann@26169: haftmann@26169: haftmann@26169: text {* Integers *} haftmann@26169: haftmann@26169: instance int :: countable huffman@35700: by (rule countable_classI [of "int_encode"]) huffman@35700: (simp add: int_encode_eq) haftmann@26169: haftmann@26169: haftmann@26169: text {* Options *} haftmann@26169: haftmann@26169: instance option :: (countable) countable huffman@35700: by (rule countable_classI [of "option_case 0 (Suc \ to_nat)"]) huffman@35700: (simp split: option.split_asm) haftmann@26169: haftmann@26169: haftmann@26169: text {* Lists *} haftmann@26169: haftmann@26169: instance list :: (countable) countable huffman@35700: by (rule countable_classI [of "list_encode \ map to_nat"]) huffman@35700: (simp add: list_encode_eq) haftmann@26169: huffman@26243: huffman@26243: text {* Functions *} huffman@26243: huffman@26243: instance "fun" :: (finite, countable) countable huffman@26243: proof huffman@26243: obtain xs :: "'a list" where xs: "set xs = UNIV" huffman@26243: using finite_list [OF finite_UNIV] .. huffman@26243: show "\to_nat::('a \ 'b) \ nat. inj to_nat" huffman@26243: proof huffman@26243: show "inj (\f. to_nat (map f xs))" huffman@26243: by (rule injI, simp add: xs expand_fun_eq) huffman@26243: qed huffman@26243: qed huffman@26243: huffman@29880: huffman@29880: subsection {* The Rationals are Countably Infinite *} huffman@29880: huffman@29880: definition nat_to_rat_surj :: "nat \ rat" where huffman@35700: "nat_to_rat_surj n = (let (a,b) = prod_decode n huffman@35700: in Fract (int_decode a) (int_decode b))" huffman@29880: huffman@29880: lemma surj_nat_to_rat_surj: "surj nat_to_rat_surj" huffman@29880: unfolding surj_def huffman@29880: proof huffman@29880: fix r::rat huffman@29880: show "\n. r = nat_to_rat_surj n" haftmann@35374: proof (cases r) haftmann@35374: fix i j assume [simp]: "r = Fract i j" and "j > 0" huffman@35700: have "r = (let m = int_encode i; n = int_encode j huffman@35700: in nat_to_rat_surj(prod_encode (m,n)))" huffman@35700: by (simp add: Let_def nat_to_rat_surj_def) huffman@29880: thus "\n. r = nat_to_rat_surj n" by(auto simp:Let_def) huffman@29880: qed huffman@29880: qed huffman@29880: huffman@29880: lemma Rats_eq_range_nat_to_rat_surj: "\ = range nat_to_rat_surj" huffman@29880: by (simp add: Rats_def surj_nat_to_rat_surj surj_range) huffman@29880: huffman@29880: context field_char_0 huffman@29880: begin huffman@29880: huffman@29880: lemma Rats_eq_range_of_rat_o_nat_to_rat_surj: huffman@29880: "\ = range (of_rat o nat_to_rat_surj)" huffman@29880: using surj_nat_to_rat_surj huffman@29880: by (auto simp: Rats_def image_def surj_def) huffman@29880: (blast intro: arg_cong[where f = of_rat]) huffman@29880: huffman@29880: lemma surj_of_rat_nat_to_rat_surj: huffman@29880: "r\\ \ \n. r = of_rat(nat_to_rat_surj n)" huffman@29880: by(simp add: Rats_eq_range_of_rat_o_nat_to_rat_surj image_def) huffman@29880: haftmann@26169: end huffman@29880: huffman@29880: instance rat :: countable huffman@29880: proof huffman@29880: show "\to_nat::rat \ nat. inj to_nat" huffman@29880: proof huffman@29880: have "surj nat_to_rat_surj" huffman@29880: by (rule surj_nat_to_rat_surj) huffman@29880: then show "inj (inv nat_to_rat_surj)" huffman@29880: by (rule surj_imp_inj_inv) huffman@29880: qed huffman@29880: qed huffman@29880: huffman@29880: end