| author | paulson <lp15@cam.ac.uk> | 
| Tue, 28 Apr 2015 16:23:05 +0100 | |
| changeset 60149 | 9b0825a00b1a | 
| parent 59506 | 4af607652318 | 
| child 60352 | d46de31a50c4 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Library/Nat_Bijection.thy | 
| 35700 | 2 | Author: Brian Huffman | 
| 3 | Author: Florian Haftmann | |
| 4 | Author: Stefan Richter | |
| 5 | Author: Tobias Nipkow | |
| 6 | Author: Alexander Krauss | |
| 7 | *) | |
| 8 | ||
| 58881 | 9 | section {* Bijections between natural numbers and other types *}
 | 
| 35700 | 10 | |
| 11 | theory Nat_Bijection | |
| 58770 | 12 | imports Main | 
| 35700 | 13 | begin | 
| 14 | ||
| 15 | subsection {* Type @{typ "nat \<times> nat"} *}
 | |
| 16 | ||
| 17 | text "Triangle numbers: 0, 1, 3, 6, 10, 15, ..." | |
| 18 | ||
| 19 | definition | |
| 20 | triangle :: "nat \<Rightarrow> nat" | |
| 21 | where | |
| 22 | "triangle n = n * Suc n div 2" | |
| 23 | ||
| 24 | lemma triangle_0 [simp]: "triangle 0 = 0" | |
| 25 | unfolding triangle_def by simp | |
| 26 | ||
| 27 | lemma triangle_Suc [simp]: "triangle (Suc n) = triangle n + Suc n" | |
| 28 | unfolding triangle_def by simp | |
| 29 | ||
| 30 | definition | |
| 31 | prod_encode :: "nat \<times> nat \<Rightarrow> nat" | |
| 32 | where | |
| 33 | "prod_encode = (\<lambda>(m, n). triangle (m + n) + m)" | |
| 34 | ||
| 35 | text {* In this auxiliary function, @{term "triangle k + m"} is an invariant. *}
 | |
| 36 | ||
| 37 | fun | |
| 38 | prod_decode_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" | |
| 39 | where | |
| 40 | "prod_decode_aux k m = | |
| 41 | (if m \<le> k then (m, k - m) else prod_decode_aux (Suc k) (m - Suc k))" | |
| 42 | ||
| 43 | declare prod_decode_aux.simps [simp del] | |
| 44 | ||
| 45 | definition | |
| 46 | prod_decode :: "nat \<Rightarrow> nat \<times> nat" | |
| 47 | where | |
| 48 | "prod_decode = prod_decode_aux 0" | |
| 49 | ||
| 50 | lemma prod_encode_prod_decode_aux: | |
| 51 | "prod_encode (prod_decode_aux k m) = triangle k + m" | |
| 52 | apply (induct k m rule: prod_decode_aux.induct) | |
| 53 | apply (subst prod_decode_aux.simps) | |
| 54 | apply (simp add: prod_encode_def) | |
| 55 | done | |
| 56 | ||
| 57 | lemma prod_decode_inverse [simp]: "prod_encode (prod_decode n) = n" | |
| 58 | unfolding prod_decode_def by (simp add: prod_encode_prod_decode_aux) | |
| 59 | ||
| 60 | lemma prod_decode_triangle_add: | |
| 61 | "prod_decode (triangle k + m) = prod_decode_aux k m" | |
| 62 | apply (induct k arbitrary: m) | |
| 63 | apply (simp add: prod_decode_def) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51489diff
changeset | 64 | apply (simp only: triangle_Suc add.assoc) | 
| 35700 | 65 | apply (subst prod_decode_aux.simps, simp) | 
| 66 | done | |
| 67 | ||
| 68 | lemma prod_encode_inverse [simp]: "prod_decode (prod_encode x) = x" | |
| 69 | unfolding prod_encode_def | |
| 70 | apply (induct x) | |
| 71 | apply (simp add: prod_decode_triangle_add) | |
| 72 | apply (subst prod_decode_aux.simps, simp) | |
| 73 | done | |
| 74 | ||
| 75 | lemma inj_prod_encode: "inj_on prod_encode A" | |
| 76 | by (rule inj_on_inverseI, rule prod_encode_inverse) | |
| 77 | ||
| 78 | lemma inj_prod_decode: "inj_on prod_decode A" | |
| 79 | by (rule inj_on_inverseI, rule prod_decode_inverse) | |
| 80 | ||
| 81 | lemma surj_prod_encode: "surj prod_encode" | |
| 82 | by (rule surjI, rule prod_decode_inverse) | |
| 83 | ||
| 84 | lemma surj_prod_decode: "surj prod_decode" | |
| 85 | by (rule surjI, rule prod_encode_inverse) | |
| 86 | ||
| 87 | lemma bij_prod_encode: "bij prod_encode" | |
| 88 | by (rule bijI [OF inj_prod_encode surj_prod_encode]) | |
| 89 | ||
| 90 | lemma bij_prod_decode: "bij prod_decode" | |
| 91 | by (rule bijI [OF inj_prod_decode surj_prod_decode]) | |
| 92 | ||
| 93 | lemma prod_encode_eq: "prod_encode x = prod_encode y \<longleftrightarrow> x = y" | |
| 94 | by (rule inj_prod_encode [THEN inj_eq]) | |
| 95 | ||
| 96 | lemma prod_decode_eq: "prod_decode x = prod_decode y \<longleftrightarrow> x = y" | |
| 97 | by (rule inj_prod_decode [THEN inj_eq]) | |
| 98 | ||
| 99 | text {* Ordering properties *}
 | |
| 100 | ||
| 101 | lemma le_prod_encode_1: "a \<le> prod_encode (a, b)" | |
| 102 | unfolding prod_encode_def by simp | |
| 103 | ||
| 104 | lemma le_prod_encode_2: "b \<le> prod_encode (a, b)" | |
| 105 | unfolding prod_encode_def by (induct b, simp_all) | |
| 106 | ||
| 107 | ||
| 108 | subsection {* Type @{typ "nat + nat"} *}
 | |
| 109 | ||
| 110 | definition | |
| 111 | sum_encode :: "nat + nat \<Rightarrow> nat" | |
| 112 | where | |
| 113 | "sum_encode x = (case x of Inl a \<Rightarrow> 2 * a | Inr b \<Rightarrow> Suc (2 * b))" | |
| 114 | ||
| 115 | definition | |
| 116 | sum_decode :: "nat \<Rightarrow> nat + nat" | |
| 117 | where | |
| 118 | "sum_decode n = (if even n then Inl (n div 2) else Inr (n div 2))" | |
| 119 | ||
| 120 | lemma sum_encode_inverse [simp]: "sum_decode (sum_encode x) = x" | |
| 121 | unfolding sum_decode_def sum_encode_def | |
| 122 | by (induct x) simp_all | |
| 123 | ||
| 124 | lemma sum_decode_inverse [simp]: "sum_encode (sum_decode n) = n" | |
| 58834 | 125 | by (simp add: even_two_times_div_two sum_decode_def sum_encode_def) | 
| 35700 | 126 | |
| 127 | lemma inj_sum_encode: "inj_on sum_encode A" | |
| 128 | by (rule inj_on_inverseI, rule sum_encode_inverse) | |
| 129 | ||
| 130 | lemma inj_sum_decode: "inj_on sum_decode A" | |
| 131 | by (rule inj_on_inverseI, rule sum_decode_inverse) | |
| 132 | ||
| 133 | lemma surj_sum_encode: "surj sum_encode" | |
| 134 | by (rule surjI, rule sum_decode_inverse) | |
| 135 | ||
| 136 | lemma surj_sum_decode: "surj sum_decode" | |
| 137 | by (rule surjI, rule sum_encode_inverse) | |
| 138 | ||
| 139 | lemma bij_sum_encode: "bij sum_encode" | |
| 140 | by (rule bijI [OF inj_sum_encode surj_sum_encode]) | |
| 141 | ||
| 142 | lemma bij_sum_decode: "bij sum_decode" | |
| 143 | by (rule bijI [OF inj_sum_decode surj_sum_decode]) | |
| 144 | ||
| 145 | lemma sum_encode_eq: "sum_encode x = sum_encode y \<longleftrightarrow> x = y" | |
| 146 | by (rule inj_sum_encode [THEN inj_eq]) | |
| 147 | ||
| 148 | lemma sum_decode_eq: "sum_decode x = sum_decode y \<longleftrightarrow> x = y" | |
| 149 | by (rule inj_sum_decode [THEN inj_eq]) | |
| 150 | ||
| 151 | ||
| 152 | subsection {* Type @{typ "int"} *}
 | |
| 153 | ||
| 154 | definition | |
| 155 | int_encode :: "int \<Rightarrow> nat" | |
| 156 | where | |
| 157 | "int_encode i = sum_encode (if 0 \<le> i then Inl (nat i) else Inr (nat (- i - 1)))" | |
| 158 | ||
| 159 | definition | |
| 160 | int_decode :: "nat \<Rightarrow> int" | |
| 161 | where | |
| 162 | "int_decode n = (case sum_decode n of Inl a \<Rightarrow> int a | Inr b \<Rightarrow> - int b - 1)" | |
| 163 | ||
| 164 | lemma int_encode_inverse [simp]: "int_decode (int_encode x) = x" | |
| 165 | unfolding int_decode_def int_encode_def by simp | |
| 166 | ||
| 167 | lemma int_decode_inverse [simp]: "int_encode (int_decode n) = n" | |
| 168 | unfolding int_decode_def int_encode_def using sum_decode_inverse [of n] | |
| 169 | by (cases "sum_decode n", simp_all) | |
| 170 | ||
| 171 | lemma inj_int_encode: "inj_on int_encode A" | |
| 172 | by (rule inj_on_inverseI, rule int_encode_inverse) | |
| 173 | ||
| 174 | lemma inj_int_decode: "inj_on int_decode A" | |
| 175 | by (rule inj_on_inverseI, rule int_decode_inverse) | |
| 176 | ||
| 177 | lemma surj_int_encode: "surj int_encode" | |
| 178 | by (rule surjI, rule int_decode_inverse) | |
| 179 | ||
| 180 | lemma surj_int_decode: "surj int_decode" | |
| 181 | by (rule surjI, rule int_encode_inverse) | |
| 182 | ||
| 183 | lemma bij_int_encode: "bij int_encode" | |
| 184 | by (rule bijI [OF inj_int_encode surj_int_encode]) | |
| 185 | ||
| 186 | lemma bij_int_decode: "bij int_decode" | |
| 187 | by (rule bijI [OF inj_int_decode surj_int_decode]) | |
| 188 | ||
| 189 | lemma int_encode_eq: "int_encode x = int_encode y \<longleftrightarrow> x = y" | |
| 190 | by (rule inj_int_encode [THEN inj_eq]) | |
| 191 | ||
| 192 | lemma int_decode_eq: "int_decode x = int_decode y \<longleftrightarrow> x = y" | |
| 193 | by (rule inj_int_decode [THEN inj_eq]) | |
| 194 | ||
| 195 | ||
| 196 | subsection {* Type @{typ "nat list"} *}
 | |
| 197 | ||
| 198 | fun | |
| 199 | list_encode :: "nat list \<Rightarrow> nat" | |
| 200 | where | |
| 201 | "list_encode [] = 0" | |
| 202 | | "list_encode (x # xs) = Suc (prod_encode (x, list_encode xs))" | |
| 203 | ||
| 204 | function | |
| 205 | list_decode :: "nat \<Rightarrow> nat list" | |
| 206 | where | |
| 207 | "list_decode 0 = []" | |
| 208 | | "list_decode (Suc n) = (case prod_decode n of (x, y) \<Rightarrow> x # list_decode y)" | |
| 209 | by pat_completeness auto | |
| 210 | ||
| 211 | termination list_decode | |
| 212 | apply (relation "measure id", simp_all) | |
| 213 | apply (drule arg_cong [where f="prod_encode"]) | |
| 37591 | 214 | apply (drule sym) | 
| 35700 | 215 | apply (simp add: le_imp_less_Suc le_prod_encode_2) | 
| 216 | done | |
| 217 | ||
| 218 | lemma list_encode_inverse [simp]: "list_decode (list_encode x) = x" | |
| 219 | by (induct x rule: list_encode.induct) simp_all | |
| 220 | ||
| 221 | lemma list_decode_inverse [simp]: "list_encode (list_decode n) = n" | |
| 222 | apply (induct n rule: list_decode.induct, simp) | |
| 223 | apply (simp split: prod.split) | |
| 224 | apply (simp add: prod_decode_eq [symmetric]) | |
| 225 | done | |
| 226 | ||
| 227 | lemma inj_list_encode: "inj_on list_encode A" | |
| 228 | by (rule inj_on_inverseI, rule list_encode_inverse) | |
| 229 | ||
| 230 | lemma inj_list_decode: "inj_on list_decode A" | |
| 231 | by (rule inj_on_inverseI, rule list_decode_inverse) | |
| 232 | ||
| 233 | lemma surj_list_encode: "surj list_encode" | |
| 234 | by (rule surjI, rule list_decode_inverse) | |
| 235 | ||
| 236 | lemma surj_list_decode: "surj list_decode" | |
| 237 | by (rule surjI, rule list_encode_inverse) | |
| 238 | ||
| 239 | lemma bij_list_encode: "bij list_encode" | |
| 240 | by (rule bijI [OF inj_list_encode surj_list_encode]) | |
| 241 | ||
| 242 | lemma bij_list_decode: "bij list_decode" | |
| 243 | by (rule bijI [OF inj_list_decode surj_list_decode]) | |
| 244 | ||
| 245 | lemma list_encode_eq: "list_encode x = list_encode y \<longleftrightarrow> x = y" | |
| 246 | by (rule inj_list_encode [THEN inj_eq]) | |
| 247 | ||
| 248 | lemma list_decode_eq: "list_decode x = list_decode y \<longleftrightarrow> x = y" | |
| 249 | by (rule inj_list_decode [THEN inj_eq]) | |
| 250 | ||
| 251 | ||
| 252 | subsection {* Finite sets of naturals *}
 | |
| 253 | ||
| 254 | subsubsection {* Preliminaries *}
 | |
| 255 | ||
| 256 | lemma finite_vimage_Suc_iff: "finite (Suc -` F) \<longleftrightarrow> finite F" | |
| 257 | apply (safe intro!: finite_vimageI inj_Suc) | |
| 258 | apply (rule finite_subset [where B="insert 0 (Suc ` Suc -` F)"]) | |
| 259 | apply (rule subsetI, case_tac x, simp, simp) | |
| 260 | apply (rule finite_insert [THEN iffD2]) | |
| 261 | apply (erule finite_imageI) | |
| 262 | done | |
| 263 | ||
| 264 | lemma vimage_Suc_insert_0: "Suc -` insert 0 A = Suc -` A" | |
| 265 | by auto | |
| 266 | ||
| 267 | lemma vimage_Suc_insert_Suc: | |
| 268 | "Suc -` insert (Suc n) A = insert n (Suc -` A)" | |
| 269 | by auto | |
| 270 | ||
| 271 | lemma div2_even_ext_nat: | |
| 58834 | 272 | fixes x y :: nat | 
| 273 | assumes "x div 2 = y div 2" | |
| 274 | and "even x \<longleftrightarrow> even y" | |
| 275 | shows "x = y" | |
| 276 | proof - | |
| 277 | from `even x \<longleftrightarrow> even y` have "x mod 2 = y mod 2" | |
| 278 | by (simp only: even_iff_mod_2_eq_zero) auto | |
| 279 | with assms have "x div 2 * 2 + x mod 2 = y div 2 * 2 + y mod 2" | |
| 280 | by simp | |
| 281 | then show ?thesis | |
| 282 | by simp | |
| 283 | qed | |
| 35700 | 284 | |
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
57512diff
changeset | 285 | |
| 35700 | 286 | subsubsection {* From sets to naturals *}
 | 
| 287 | ||
| 288 | definition | |
| 289 | set_encode :: "nat set \<Rightarrow> nat" | |
| 290 | where | |
| 291 | "set_encode = setsum (op ^ 2)" | |
| 292 | ||
| 293 | lemma set_encode_empty [simp]: "set_encode {} = 0"
 | |
| 294 | by (simp add: set_encode_def) | |
| 295 | ||
| 59506 
4af607652318
Not a simprule, as it complicates proofs
 paulson <lp15@cam.ac.uk> parents: 
58881diff
changeset | 296 | lemma set_encode_inf: "~ finite A \<Longrightarrow> set_encode A = 0" | 
| 
4af607652318
Not a simprule, as it complicates proofs
 paulson <lp15@cam.ac.uk> parents: 
58881diff
changeset | 297 | by (simp add: set_encode_def) | 
| 
4af607652318
Not a simprule, as it complicates proofs
 paulson <lp15@cam.ac.uk> parents: 
58881diff
changeset | 298 | |
| 35700 | 299 | lemma set_encode_insert [simp]: | 
| 300 | "\<lbrakk>finite A; n \<notin> A\<rbrakk> \<Longrightarrow> set_encode (insert n A) = 2^n + set_encode A" | |
| 301 | by (simp add: set_encode_def) | |
| 302 | ||
| 303 | lemma even_set_encode_iff: "finite A \<Longrightarrow> even (set_encode A) \<longleftrightarrow> 0 \<notin> A" | |
| 304 | unfolding set_encode_def by (induct set: finite, auto) | |
| 305 | ||
| 306 | lemma set_encode_vimage_Suc: "set_encode (Suc -` A) = set_encode A div 2" | |
| 307 | apply (cases "finite A") | |
| 308 | apply (erule finite_induct, simp) | |
| 309 | apply (case_tac x) | |
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
57512diff
changeset | 310 | apply (simp add: even_set_encode_iff vimage_Suc_insert_0) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51489diff
changeset | 311 | apply (simp add: finite_vimageI add.commute vimage_Suc_insert_Suc) | 
| 35700 | 312 | apply (simp add: set_encode_def finite_vimage_Suc_iff) | 
| 313 | done | |
| 314 | ||
| 315 | lemmas set_encode_div_2 = set_encode_vimage_Suc [symmetric] | |
| 316 | ||
| 317 | subsubsection {* From naturals to sets *}
 | |
| 318 | ||
| 319 | definition | |
| 320 | set_decode :: "nat \<Rightarrow> nat set" | |
| 321 | where | |
| 322 |   "set_decode x = {n. odd (x div 2 ^ n)}"
 | |
| 323 | ||
| 324 | lemma set_decode_0 [simp]: "0 \<in> set_decode x \<longleftrightarrow> odd x" | |
| 325 | by (simp add: set_decode_def) | |
| 326 | ||
| 327 | lemma set_decode_Suc [simp]: | |
| 328 | "Suc n \<in> set_decode x \<longleftrightarrow> n \<in> set_decode (x div 2)" | |
| 329 | by (simp add: set_decode_def div_mult2_eq) | |
| 330 | ||
| 331 | lemma set_decode_zero [simp]: "set_decode 0 = {}"
 | |
| 332 | by (simp add: set_decode_def) | |
| 333 | ||
| 334 | lemma set_decode_div_2: "set_decode (x div 2) = Suc -` set_decode x" | |
| 335 | by auto | |
| 336 | ||
| 337 | lemma set_decode_plus_power_2: | |
| 338 | "n \<notin> set_decode z \<Longrightarrow> set_decode (2 ^ n + z) = insert n (set_decode z)" | |
| 339 | apply (induct n arbitrary: z, simp_all) | |
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
57512diff
changeset | 340 | apply (rule set_eqI, induct_tac x, simp, simp) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51489diff
changeset | 341 | apply (rule set_eqI, induct_tac x, simp, simp add: add.commute) | 
| 35700 | 342 | done | 
| 343 | ||
| 344 | lemma finite_set_decode [simp]: "finite (set_decode n)" | |
| 345 | apply (induct n rule: nat_less_induct) | |
| 346 | apply (case_tac "n = 0", simp) | |
| 347 | apply (drule_tac x="n div 2" in spec, simp) | |
| 348 | apply (simp add: set_decode_div_2) | |
| 349 | apply (simp add: finite_vimage_Suc_iff) | |
| 350 | done | |
| 351 | ||
| 352 | subsubsection {* Proof of isomorphism *}
 | |
| 353 | ||
| 354 | lemma set_decode_inverse [simp]: "set_encode (set_decode n) = n" | |
| 355 | apply (induct n rule: nat_less_induct) | |
| 356 | apply (case_tac "n = 0", simp) | |
| 357 | apply (drule_tac x="n div 2" in spec, simp) | |
| 358 | apply (simp add: set_decode_div_2 set_encode_vimage_Suc) | |
| 359 | apply (erule div2_even_ext_nat) | |
| 360 | apply (simp add: even_set_encode_iff) | |
| 361 | done | |
| 362 | ||
| 363 | lemma set_encode_inverse [simp]: "finite A \<Longrightarrow> set_decode (set_encode A) = A" | |
| 364 | apply (erule finite_induct, simp_all) | |
| 365 | apply (simp add: set_decode_plus_power_2) | |
| 366 | done | |
| 367 | ||
| 368 | lemma inj_on_set_encode: "inj_on set_encode (Collect finite)" | |
| 369 | by (rule inj_on_inverseI [where g="set_decode"], simp) | |
| 370 | ||
| 371 | lemma set_encode_eq: | |
| 372 | "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> set_encode A = set_encode B \<longleftrightarrow> A = B" | |
| 373 | by (rule iffI, simp add: inj_onD [OF inj_on_set_encode], simp) | |
| 374 | ||
| 51414 | 375 | lemma subset_decode_imp_le: assumes "set_decode m \<subseteq> set_decode n" shows "m \<le> n" | 
| 376 | proof - | |
| 377 | have "n = m + set_encode (set_decode n - set_decode m)" | |
| 378 | proof - | |
| 379 | obtain A B where "m = set_encode A" "finite A" | |
| 380 | "n = set_encode B" "finite B" | |
| 381 | by (metis finite_set_decode set_decode_inverse) | |
| 382 | thus ?thesis using assms | |
| 383 | apply auto | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
51489diff
changeset | 384 | apply (simp add: set_encode_def add.commute setsum.subset_diff) | 
| 51414 | 385 | done | 
| 386 | qed | |
| 387 | thus ?thesis | |
| 388 | by (metis le_add1) | |
| 389 | qed | |
| 390 | ||
| 35700 | 391 | end | 
| 51489 | 392 |