| author | wenzelm | 
| Fri, 03 Oct 2008 21:06:39 +0200 | |
| changeset 28491 | c5420429a5aa | 
| parent 27487 | c8a6ce181805 | 
| child 28823 | dcbef866c9e2 | 
| permissions | -rw-r--r-- | 
| 26169 | 1 | (* Title: HOL/Library/Countable.thy | 
| 2 | ID: $Id$ | |
| 26350 | 3 | Author: Alexander Krauss, TU Muenchen | 
| 26169 | 4 | *) | 
| 5 | ||
| 6 | header {* Encoding (almost) everything into natural numbers *}
 | |
| 7 | ||
| 8 | theory Countable | |
| 27487 | 9 | imports Plain "~~/src/HOL/List" "~~/src/HOL/Hilbert_Choice" | 
| 26169 | 10 | begin | 
| 11 | ||
| 12 | subsection {* The class of countable types *}
 | |
| 13 | ||
| 14 | class countable = itself + | |
| 15 | assumes ex_inj: "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat" | |
| 16 | ||
| 17 | lemma countable_classI: | |
| 18 | fixes f :: "'a \<Rightarrow> nat" | |
| 19 | assumes "\<And>x y. f x = f y \<Longrightarrow> x = y" | |
| 20 |   shows "OFCLASS('a, countable_class)"
 | |
| 21 | proof (intro_classes, rule exI) | |
| 22 | show "inj f" | |
| 23 | by (rule injI [OF assms]) assumption | |
| 24 | qed | |
| 25 | ||
| 26 | ||
| 26585 | 27 | subsection {* Conversion functions *}
 | 
| 26169 | 28 | |
| 29 | definition to_nat :: "'a\<Colon>countable \<Rightarrow> nat" where | |
| 30 | "to_nat = (SOME f. inj f)" | |
| 31 | ||
| 32 | definition from_nat :: "nat \<Rightarrow> 'a\<Colon>countable" where | |
| 33 | "from_nat = inv (to_nat \<Colon> 'a \<Rightarrow> nat)" | |
| 34 | ||
| 35 | lemma inj_to_nat [simp]: "inj to_nat" | |
| 36 | by (rule exE_some [OF ex_inj]) (simp add: to_nat_def) | |
| 37 | ||
| 38 | lemma to_nat_split [simp]: "to_nat x = to_nat y \<longleftrightarrow> x = y" | |
| 39 | using injD [OF inj_to_nat] by auto | |
| 40 | ||
| 41 | lemma from_nat_to_nat [simp]: | |
| 42 | "from_nat (to_nat x) = x" | |
| 43 | by (simp add: from_nat_def) | |
| 44 | ||
| 45 | ||
| 46 | subsection {* Countable types *}
 | |
| 47 | ||
| 48 | instance nat :: countable | |
| 49 | by (rule countable_classI [of "id"]) simp | |
| 50 | ||
| 51 | subclass (in finite) countable | |
| 52 | proof unfold_locales | |
| 53 | have "finite (UNIV\<Colon>'a set)" by (rule finite_UNIV) | |
| 54 | with finite_conv_nat_seg_image [of UNIV] | |
| 55 | obtain n and f :: "nat \<Rightarrow> 'a" | |
| 56 |     where "UNIV = f ` {i. i < n}" by auto
 | |
| 57 | then have "surj f" unfolding surj_def by auto | |
| 58 | then have "inj (inv f)" by (rule surj_imp_inj_inv) | |
| 59 | then show "\<exists>to_nat \<Colon> 'a \<Rightarrow> nat. inj to_nat" by (rule exI[of inj]) | |
| 60 | qed | |
| 61 | ||
| 62 | text {* Pairs *}
 | |
| 63 | ||
| 64 | primrec sum :: "nat \<Rightarrow> nat" | |
| 65 | where | |
| 66 | "sum 0 = 0" | |
| 67 | | "sum (Suc n) = Suc n + sum n" | |
| 68 | ||
| 69 | lemma sum_arith: "sum n = n * Suc n div 2" | |
| 70 | by (induct n) auto | |
| 71 | ||
| 72 | lemma sum_mono: "n \<ge> m \<Longrightarrow> sum n \<ge> sum m" | |
| 73 | by (induct n m rule: diff_induct) auto | |
| 74 | ||
| 75 | definition | |
| 76 | "pair_encode = (\<lambda>(m, n). sum (m + n) + m)" | |
| 77 | ||
| 78 | lemma inj_pair_cencode: "inj pair_encode" | |
| 79 | unfolding pair_encode_def | |
| 80 | proof (rule injI, simp only: split_paired_all split_conv) | |
| 81 | fix a b c d | |
| 82 | assume eq: "sum (a + b) + a = sum (c + d) + c" | |
| 83 | have "a + b = c + d \<or> a + b \<ge> Suc (c + d) \<or> c + d \<ge> Suc (a + b)" by arith | |
| 84 | then | |
| 85 | show "(a, b) = (c, d)" | |
| 86 | proof (elim disjE) | |
| 87 | assume sumeq: "a + b = c + d" | |
| 88 | then have "a = c" using eq by auto | |
| 89 | moreover from sumeq this have "b = d" by auto | |
| 90 | ultimately show ?thesis by simp | |
| 91 | next | |
| 92 | assume "a + b \<ge> Suc (c + d)" | |
| 93 | from sum_mono[OF this] eq | |
| 94 | show ?thesis by auto | |
| 95 | next | |
| 96 | assume "c + d \<ge> Suc (a + b)" | |
| 97 | from sum_mono[OF this] eq | |
| 98 | show ?thesis by auto | |
| 99 | qed | |
| 100 | qed | |
| 101 | ||
| 102 | instance "*" :: (countable, countable) countable | |
| 103 | by (rule countable_classI [of "\<lambda>(x, y). pair_encode (to_nat x, to_nat y)"]) | |
| 104 | (auto dest: injD [OF inj_pair_cencode] injD [OF inj_to_nat]) | |
| 105 | ||
| 106 | ||
| 107 | text {* Sums *}
 | |
| 108 | ||
| 109 | instance "+":: (countable, countable) countable | |
| 110 | by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a) | |
| 111 | | Inr b \<Rightarrow> to_nat (True, to_nat b))"]) | |
| 112 | (auto split:sum.splits) | |
| 113 | ||
| 114 | ||
| 115 | text {* Integers *}
 | |
| 116 | ||
| 117 | lemma int_cases: "(i::int) = 0 \<or> i < 0 \<or> i > 0" | |
| 118 | by presburger | |
| 119 | ||
| 120 | lemma int_pos_neg_zero: | |
| 121 | obtains (zero) "(z::int) = 0" "sgn z = 0" "abs z = 0" | |
| 122 | | (pos) n where "z = of_nat n" "sgn z = 1" "abs z = of_nat n" | |
| 123 | | (neg) n where "z = - (of_nat n)" "sgn z = -1" "abs z = of_nat n" | |
| 26580 
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
 krauss parents: 
26350diff
changeset | 124 | apply atomize_elim | 
| 26169 | 125 | apply (insert int_cases[of z]) | 
| 126 | apply (auto simp:zsgn_def) | |
| 127 | apply (rule_tac x="nat (-z)" in exI, simp) | |
| 128 | apply (rule_tac x="nat z" in exI, simp) | |
| 129 | done | |
| 130 | ||
| 131 | instance int :: countable | |
| 132 | proof (rule countable_classI [of "(\<lambda>i. to_nat (nat (sgn i + 1), nat (abs i)))"], | |
| 133 | auto dest: injD [OF inj_to_nat]) | |
| 134 | fix x y | |
| 135 | assume a: "nat (sgn x + 1) = nat (sgn y + 1)" "nat (abs x) = nat (abs y)" | |
| 136 | show "x = y" | |
| 137 | proof (cases rule: int_pos_neg_zero[of x]) | |
| 138 | case zero | |
| 139 | with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto | |
| 140 | next | |
| 141 | case (pos n) | |
| 142 | with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto | |
| 143 | next | |
| 144 | case (neg n) | |
| 145 | with a show "x = y" by (cases rule: int_pos_neg_zero[of y]) auto | |
| 146 | qed | |
| 147 | qed | |
| 148 | ||
| 149 | ||
| 150 | text {* Options *}
 | |
| 151 | ||
| 152 | instance option :: (countable) countable | |
| 153 | by (rule countable_classI[of "\<lambda>x. case x of None \<Rightarrow> 0 | |
| 154 | | Some y \<Rightarrow> Suc (to_nat y)"]) | |
| 155 | (auto split:option.splits) | |
| 156 | ||
| 157 | ||
| 158 | text {* Lists *}
 | |
| 159 | ||
| 160 | lemma from_nat_to_nat_map [simp]: "map from_nat (map to_nat xs) = xs" | |
| 161 | by (simp add: comp_def map_compose [symmetric]) | |
| 162 | ||
| 163 | primrec | |
| 164 | list_encode :: "'a\<Colon>countable list \<Rightarrow> nat" | |
| 165 | where | |
| 166 | "list_encode [] = 0" | |
| 167 | | "list_encode (x#xs) = Suc (to_nat (x, list_encode xs))" | |
| 168 | ||
| 169 | instance list :: (countable) countable | |
| 170 | proof (rule countable_classI [of "list_encode"]) | |
| 171 | fix xs ys :: "'a list" | |
| 172 | assume cenc: "list_encode xs = list_encode ys" | |
| 173 | then show "xs = ys" | |
| 174 | proof (induct xs arbitrary: ys) | |
| 175 | case (Nil ys) | |
| 176 | with cenc show ?case by (cases ys, auto) | |
| 177 | next | |
| 178 | case (Cons x xs' ys) | |
| 179 | thus ?case by (cases ys) auto | |
| 180 | qed | |
| 181 | qed | |
| 182 | ||
| 26243 | 183 | |
| 184 | text {* Functions *}
 | |
| 185 | ||
| 186 | instance "fun" :: (finite, countable) countable | |
| 187 | proof | |
| 188 | obtain xs :: "'a list" where xs: "set xs = UNIV" | |
| 189 | using finite_list [OF finite_UNIV] .. | |
| 190 |   show "\<exists>to_nat::('a \<Rightarrow> 'b) \<Rightarrow> nat. inj to_nat"
 | |
| 191 | proof | |
| 192 | show "inj (\<lambda>f. to_nat (map f xs))" | |
| 193 | by (rule injI, simp add: xs expand_fun_eq) | |
| 194 | qed | |
| 195 | qed | |
| 196 | ||
| 26169 | 197 | end |