| author | nipkow | 
| Mon, 16 Aug 2004 14:21:54 +0200 | |
| changeset 15130 | dc6be28d7f4e | 
| parent 14706 | 71590b7733b7 | 
| child 15131 | c69542757a4d | 
| permissions | -rw-r--r-- | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 1 | (* Title: HOL/Library/NatPair.thy | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 14414 | 3 | Author: Stefan Richter | 
| 4 | Copyright 2003 Technische Universitaet Muenchen | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 5 | *) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 6 | |
| 14706 | 7 | header {* Pairs of Natural Numbers *}
 | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 8 | |
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 9 | theory NatPair = Main: | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 10 | |
| 14706 | 11 | text{*
 | 
| 12 |   An injective function from @{text "\<nat>\<twosuperior>"} to @{text
 | |
| 13 | \<nat>}. Definition and proofs are from \cite[page | |
| 14 |   85]{Oberschelp:1993}.
 | |
| 15 | *} | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 16 | |
| 14706 | 17 | constdefs | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 18 | nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 19 | "nat2_to_nat pair \<equiv> let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n" | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 20 | |
| 14706 | 21 | lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)" | 
| 22 | proof (cases "2 dvd a") | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 23 | case True | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 24 | thus ?thesis by (rule dvd_mult2) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 25 | next | 
| 14706 | 26 | case False | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 27 | hence "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0) | 
| 14706 | 28 | hence "Suc a mod 2 = 0" by (simp add: mod_Suc) | 
| 29 | hence "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0) | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 30 | thus ?thesis by (rule dvd_mult) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 31 | qed | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 32 | |
| 14706 | 33 | lemma | 
| 34 | assumes eq: "nat2_to_nat (u,v) = nat2_to_nat (x,y)" | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 35 | shows nat2_to_nat_help: "u+v \<le> x+y" | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 36 | proof (rule classical) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 37 | assume "\<not> ?thesis" | 
| 14706 | 38 | hence contrapos: "x+y < u+v" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 39 | by simp | 
| 14706 | 40 | have "nat2_to_nat (x,y) < (x+y) * Suc (x+y) div 2 + Suc (x + y)" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 41 | by (unfold nat2_to_nat_def) (simp add: Let_def) | 
| 14706 | 42 | also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2" | 
| 43 | by (simp only: div_mult_self1_is_m) | |
| 44 | also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2 | |
| 45 | + ((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2" | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 46 | proof - | 
| 14706 | 47 | have "2 dvd (x+y)*Suc(x+y)" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 48 | by (rule dvd2_a_x_suc_a) | 
| 14706 | 49 | hence "(x+y)*Suc(x+y) mod 2 = 0" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 50 | by (simp only: dvd_eq_mod_eq_0) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 51 | also | 
| 14706 | 52 | have "2 * Suc(x+y) mod 2 = 0" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 53 | by (rule mod_mult_self1_is_0) | 
| 14706 | 54 | ultimately have | 
| 55 | "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0" | |
| 56 | by simp | |
| 57 | thus ?thesis | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 58 | by simp | 
| 14706 | 59 | qed | 
| 60 | also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2" | |
| 61 | by (rule div_add1_eq [symmetric]) | |
| 62 | also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2" | |
| 63 | by (simp only: add_mult_distrib [symmetric]) | |
| 64 | also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2" | |
| 65 | by (simp only: mult_le_mono div_le_mono) | |
| 66 | also have "\<dots> \<le> nat2_to_nat (u,v)" | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 67 | by (unfold nat2_to_nat_def) (simp add: Let_def) | 
| 14706 | 68 | finally show ?thesis | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 69 | by (simp only: eq) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 70 | qed | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 71 | |
| 14706 | 72 | theorem nat2_to_nat_inj: "inj nat2_to_nat" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 73 | proof - | 
| 14706 | 74 |   {
 | 
| 75 | fix u v x y assume "nat2_to_nat (u,v) = nat2_to_nat (x,y)" | |
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 76 | hence "u+v \<le> x+y" by (rule nat2_to_nat_help) | 
| 14706 | 77 | also from prems [symmetric] have "x+y \<le> u+v" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 78 | by (rule nat2_to_nat_help) | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 79 | finally have eq: "u+v = x+y" . | 
| 14706 | 80 | with prems have ux: "u=x" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 81 | by (simp add: nat2_to_nat_def Let_def) | 
| 14706 | 82 | with eq have vy: "v=y" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 83 | by simp | 
| 14706 | 84 | with ux have "(u,v) = (x,y)" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 85 | by simp | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 86 | } | 
| 14706 | 87 | hence "\<And>x y. nat2_to_nat x = nat2_to_nat y \<Longrightarrow> x=y" | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 88 | by fast | 
| 14706 | 89 | thus ?thesis | 
| 14127 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 90 | by (unfold inj_on_def) simp | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 91 | qed | 
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 92 | |
| 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 paulson parents: diff
changeset | 93 | end |