# HG changeset patch # User haftmann # Date 1204144867 -3600 # Node ID 73027318f9ba09ec6b17533cd0326234e80286c5 # Parent 3bd9ac4e0b971d16d1c19ac2861d59760da78df7 added theory for countable types diff -r 3bd9ac4e0b97 -r 73027318f9ba src/HOL/Library/Countable.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Countable.thy Wed Feb 27 21:41:07 2008 +0100 @@ -0,0 +1,183 @@ +(* Title: HOL/Library/Countable.thy + ID: $Id$ + Author: Tobias Nipkow +*) + +header {* Encoding (almost) everything into natural numbers *} + +theory Countable +imports Finite_Set List Hilbert_Choice +begin + +subsection {* The class of countable types *} + +class countable = itself + + assumes ex_inj: "\to_nat \ 'a \ nat. inj to_nat" + +lemma countable_classI: + fixes f :: "'a \ nat" + assumes "\x y. f x = f y \ x = y" + shows "OFCLASS('a, countable_class)" +proof (intro_classes, rule exI) + show "inj f" + by (rule injI [OF assms]) assumption +qed + + +subsection {* Converion functions *} + +definition to_nat :: "'a\countable \ nat" where + "to_nat = (SOME f. inj f)" + +definition from_nat :: "nat \ 'a\countable" where + "from_nat = inv (to_nat \ 'a \ nat)" + +lemma inj_to_nat [simp]: "inj to_nat" + by (rule exE_some [OF ex_inj]) (simp add: to_nat_def) + +lemma to_nat_split [simp]: "to_nat x = to_nat y \ x = y" + using injD [OF inj_to_nat] by auto + +lemma from_nat_to_nat [simp]: + "from_nat (to_nat x) = x" + by (simp add: from_nat_def) + + +subsection {* Countable types *} + +instance nat :: countable + by (rule countable_classI [of "id"]) simp + +subclass (in finite) countable +proof unfold_locales + have "finite (UNIV\'a set)" by (rule finite_UNIV) + with finite_conv_nat_seg_image [of UNIV] + obtain n and f :: "nat \ 'a" + where "UNIV = f ` {i. i < n}" by auto + then have "surj f" unfolding surj_def by auto + then have "inj (inv f)" by (rule surj_imp_inj_inv) + then show "\to_nat \ 'a \ nat. inj to_nat" by (rule exI[of inj]) +qed + +text {* Pairs *} + +primrec sum :: "nat \ nat" +where + "sum 0 = 0" +| "sum (Suc n) = Suc n + sum n" + +lemma sum_arith: "sum n = n * Suc n div 2" + by (induct n) auto + +lemma sum_mono: "n \ m \ sum n \ sum m" + by (induct n m rule: diff_induct) auto + +definition + "pair_encode = (\(m, n). sum (m + n) + m)" + +lemma inj_pair_cencode: "inj pair_encode" + unfolding pair_encode_def +proof (rule injI, simp only: split_paired_all split_conv) + fix a b c d + assume eq: "sum (a + b) + a = sum (c + d) + c" + have "a + b = c + d \ a + b \ Suc (c + d) \ c + d \ Suc (a + b)" by arith + then + show "(a, b) = (c, d)" + proof (elim disjE) + assume sumeq: "a + b = c + d" + then have "a = c" using eq by auto + moreover from sumeq this have "b = d" by auto + ultimately show ?thesis by simp + next + assume "a + b \ Suc (c + d)" + from sum_mono[OF this] eq + show ?thesis by auto + next + assume "c + d \ Suc (a + b)" + from sum_mono[OF this] eq + show ?thesis by auto + qed +qed + +instance "*" :: (countable, countable) countable +by (rule countable_classI [of "\(x, y). pair_encode (to_nat x, to_nat y)"]) + (auto dest: injD [OF inj_pair_cencode] injD [OF inj_to_nat]) + + +text {* Sums *} + +instance "+":: (countable, countable) countable + by (rule countable_classI [of "(\x. case x of Inl a \ to_nat (False, to_nat a) + | Inr b \ to_nat (True, to_nat b))"]) + (auto split:sum.splits) + + +text {* Integers *} + +lemma int_cases: "(i::int) = 0 \ i < 0 \ i > 0" +by presburger + +lemma int_pos_neg_zero: + obtains (zero) "(z::int) = 0" "sgn z = 0" "abs z = 0" + | (pos) n where "z = of_nat n" "sgn z = 1" "abs z = of_nat n" + | (neg) n where "z = - (of_nat n)" "sgn z = -1" "abs z = of_nat n" +apply elim_to_cases +apply (insert int_cases[of z]) +apply (auto simp:zsgn_def) +apply (rule_tac x="nat (-z)" in exI, simp) +apply (rule_tac x="nat z" in exI, simp) +done + +instance int :: countable +proof (rule countable_classI [of "(\i. to_nat (nat (sgn i + 1), nat (abs i)))"], + auto dest: injD [OF inj_to_nat]) + fix x y + assume a: "nat (sgn x + 1) = nat (sgn y + 1)" "nat (abs x) = nat (abs y)" + show "x = y" + proof (cases rule: int_pos_neg_zero[of x]) + case zero + with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto + next + case (pos n) + with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto + next + case (neg n) + with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto + qed +qed + + +text {* Options *} + +instance option :: (countable) countable +by (rule countable_classI[of "\x. case x of None \ 0 + | Some y \ Suc (to_nat y)"]) + (auto split:option.splits) + + +text {* Lists *} + +lemma from_nat_to_nat_map [simp]: "map from_nat (map to_nat xs) = xs" + by (simp add: comp_def map_compose [symmetric]) + +primrec + list_encode :: "'a\countable list \ nat" +where + "list_encode [] = 0" +| "list_encode (x#xs) = Suc (to_nat (x, list_encode xs))" + +instance list :: (countable) countable +proof (rule countable_classI [of "list_encode"]) + fix xs ys :: "'a list" + assume cenc: "list_encode xs = list_encode ys" + then show "xs = ys" + proof (induct xs arbitrary: ys) + case (Nil ys) + with cenc show ?case by (cases ys, auto) + next + case (Cons x xs' ys) + thus ?case by (cases ys) auto + qed +qed + +end