haftmann@26169: (* Title: HOL/Library/Countable.thy haftmann@26350: Author: Alexander Krauss, TU Muenchen huffman@43534: Author: Brian Huffman, Portland State University 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: hoelzl@43992: lemma inj_on_to_nat[simp, intro]: "inj_on to_nat S" hoelzl@43992: using inj_to_nat by (auto simp: inj_on_def) hoelzl@43992: 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@37678: instance prod :: (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@37678: instance sum :: (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: haftmann@37652: text {* Further *} haftmann@37652: haftmann@37652: instance String.literal :: countable bulwahn@39250: by (rule countable_classI [of "to_nat o explode"]) bulwahn@39250: (auto simp add: explode_inject) haftmann@37652: 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))" nipkow@39302: by (rule injI, simp add: xs fun_eq_iff) 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" hoelzl@40702: by (simp add: Rats_def surj_nat_to_rat_surj) 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@43534: huffman@43534: subsection {* Automatically proving countability of datatypes *} huffman@43534: huffman@43534: inductive finite_item :: "'a Datatype.item \ bool" where huffman@43534: undefined: "finite_item undefined" huffman@43534: | In0: "finite_item x \ finite_item (Datatype.In0 x)" huffman@43534: | In1: "finite_item x \ finite_item (Datatype.In1 x)" huffman@43534: | Leaf: "finite_item (Datatype.Leaf a)" huffman@43534: | Scons: "\finite_item x; finite_item y\ \ finite_item (Datatype.Scons x y)" huffman@43534: huffman@43534: function huffman@43534: nth_item :: "nat \ ('a::countable) Datatype.item" huffman@43534: where huffman@43534: "nth_item 0 = undefined" huffman@43534: | "nth_item (Suc n) = huffman@43534: (case sum_decode n of huffman@43534: Inl i \ huffman@43534: (case sum_decode i of huffman@43534: Inl j \ Datatype.In0 (nth_item j) huffman@43534: | Inr j \ Datatype.In1 (nth_item j)) huffman@43534: | Inr i \ huffman@43534: (case sum_decode i of huffman@43534: Inl j \ Datatype.Leaf (from_nat j) huffman@43534: | Inr j \ huffman@43534: (case prod_decode j of huffman@43534: (a, b) \ Datatype.Scons (nth_item a) (nth_item b))))" huffman@43534: by pat_completeness auto huffman@43534: huffman@43534: lemma le_sum_encode_Inl: "x \ y \ x \ sum_encode (Inl y)" huffman@43534: unfolding sum_encode_def by simp huffman@43534: huffman@43534: lemma le_sum_encode_Inr: "x \ y \ x \ sum_encode (Inr y)" huffman@43534: unfolding sum_encode_def by simp huffman@43534: huffman@43534: termination huffman@43534: by (relation "measure id") huffman@43534: (auto simp add: sum_encode_eq [symmetric] prod_encode_eq [symmetric] huffman@43534: le_imp_less_Suc le_sum_encode_Inl le_sum_encode_Inr huffman@43534: le_prod_encode_1 le_prod_encode_2) huffman@43534: huffman@43534: lemma nth_item_covers: "finite_item x \ \n. nth_item n = x" huffman@43534: proof (induct set: finite_item) huffman@43534: case undefined huffman@43534: have "nth_item 0 = undefined" by simp huffman@43534: thus ?case .. huffman@43534: next huffman@43534: case (In0 x) huffman@43534: then obtain n where "nth_item n = x" by fast huffman@43534: hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inl n))))) huffman@43534: = Datatype.In0 x" by simp huffman@43534: thus ?case .. huffman@43534: next huffman@43534: case (In1 x) huffman@43534: then obtain n where "nth_item n = x" by fast huffman@43534: hence "nth_item (Suc (sum_encode (Inl (sum_encode (Inr n))))) huffman@43534: = Datatype.In1 x" by simp huffman@43534: thus ?case .. huffman@43534: next huffman@43534: case (Leaf a) huffman@43534: have "nth_item (Suc (sum_encode (Inr (sum_encode (Inl (to_nat a)))))) huffman@43534: = Datatype.Leaf a" by simp huffman@43534: thus ?case .. huffman@43534: next huffman@43534: case (Scons x y) huffman@43534: then obtain i j where "nth_item i = x" and "nth_item j = y" by fast huffman@43534: hence "nth_item huffman@43534: (Suc (sum_encode (Inr (sum_encode (Inr (prod_encode (i, j))))))) huffman@43534: = Datatype.Scons x y" by simp huffman@43534: thus ?case .. huffman@43534: qed huffman@43534: huffman@43534: theorem countable_datatype: huffman@43534: fixes Rep :: "'b \ ('a::countable) Datatype.item" huffman@43534: fixes Abs :: "('a::countable) Datatype.item \ 'b" huffman@43534: fixes rep_set :: "('a::countable) Datatype.item \ bool" huffman@43534: assumes type: "type_definition Rep Abs (Collect rep_set)" huffman@43534: assumes finite_item: "\x. rep_set x \ finite_item x" huffman@43534: shows "OFCLASS('b, countable_class)" huffman@43534: proof huffman@43534: def f \ "\y. LEAST n. nth_item n = Rep y" huffman@43534: { huffman@43534: fix y :: 'b huffman@43534: have "rep_set (Rep y)" huffman@43534: using type_definition.Rep [OF type] by simp huffman@43534: hence "finite_item (Rep y)" huffman@43534: by (rule finite_item) huffman@43534: hence "\n. nth_item n = Rep y" huffman@43534: by (rule nth_item_covers) huffman@43534: hence "nth_item (f y) = Rep y" huffman@43534: unfolding f_def by (rule LeastI_ex) huffman@43534: hence "Abs (nth_item (f y)) = y" huffman@43534: using type_definition.Rep_inverse [OF type] by simp huffman@43534: } huffman@43534: hence "inj f" huffman@43534: by (rule inj_on_inverseI) huffman@43534: thus "\f::'b \ nat. inj f" huffman@43534: by - (rule exI) huffman@43534: qed huffman@43534: huffman@43534: method_setup countable_datatype = {* huffman@43534: let huffman@43534: fun countable_tac ctxt = huffman@43534: SUBGOAL (fn (goal, i) => huffman@43534: let huffman@43534: val ty_name = huffman@43534: (case goal of wenzelm@46998: (_ $ Const (@{const_name TYPE}, Type (@{type_name itself}, [Type (n, _)]))) => n huffman@43534: | _ => raise Match) huffman@43534: val typedef_info = hd (Typedef.get_info ctxt ty_name) huffman@43534: val typedef_thm = #type_definition (snd typedef_info) huffman@43534: val pred_name = huffman@43534: (case HOLogic.dest_Trueprop (concl_of typedef_thm) of huffman@43534: (typedef $ rep $ abs $ (collect $ Const (n, _))) => n huffman@43534: | _ => raise Match) huffman@43534: val induct_info = Inductive.the_inductive ctxt pred_name huffman@43534: val pred_names = #names (fst induct_info) huffman@43534: val induct_thms = #inducts (snd induct_info) huffman@43534: val alist = pred_names ~~ induct_thms huffman@43534: val induct_thm = the (AList.lookup (op =) alist pred_name) huffman@43534: val rules = @{thms finite_item.intros} huffman@43534: in huffman@43534: SOLVED' (fn i => EVERY huffman@43534: [rtac @{thm countable_datatype} i, huffman@43534: rtac typedef_thm i, huffman@43534: etac induct_thm i, huffman@43534: REPEAT (resolve_tac rules i ORELSE atac i)]) 1 huffman@43534: end) huffman@43534: in huffman@43534: Scan.succeed (fn ctxt => SIMPLE_METHOD' (countable_tac ctxt)) huffman@29880: end huffman@43534: *} "prove countable class instances for datatypes" huffman@43534: huffman@43534: hide_const (open) finite_item nth_item huffman@43534: huffman@43534: huffman@43534: subsection {* Countable datatypes *} huffman@43534: huffman@43534: instance typerep :: countable huffman@43534: by countable_datatype huffman@43534: huffman@43534: end