| author | wenzelm | 
| Sun, 26 Nov 2006 18:07:31 +0100 | |
| changeset 21533 | bada5ac1216a | 
| parent 21404 | eb85850d3eb7 | 
| child 23394 | 474ff28210c0 | 
| 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  | 
|
| 15131 | 9  | 
theory NatPair  | 
| 15140 | 10  | 
imports Main  | 
| 15131 | 11  | 
begin  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
12  | 
|
| 14706 | 13  | 
text{*
 | 
| 19736 | 14  | 
  An injective function from @{text "\<nat>\<twosuperior>"} to @{text \<nat>}.  Definition
 | 
15  | 
  and proofs are from \cite[page 85]{Oberschelp:1993}.
 | 
|
| 14706 | 16  | 
*}  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
17  | 
|
| 19736 | 18  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
19736 
diff
changeset
 | 
19  | 
nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" where  | 
| 19736 | 20  | 
"nat2_to_nat pair = (let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n)"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 14706 | 22  | 
lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)"  | 
23  | 
proof (cases "2 dvd a")  | 
|
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
24  | 
case True  | 
| 19736 | 25  | 
then show ?thesis by (rule dvd_mult2)  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
26  | 
next  | 
| 14706 | 27  | 
case False  | 
| 19736 | 28  | 
then have "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0)  | 
29  | 
then have "Suc a mod 2 = 0" by (simp add: mod_Suc)  | 
|
30  | 
then have "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0)  | 
|
31  | 
then show ?thesis by (rule dvd_mult)  | 
|
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
32  | 
qed  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
33  | 
|
| 14706 | 34  | 
lemma  | 
35  | 
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
 | 
36  | 
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
 | 
37  | 
proof (rule classical)  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
38  | 
assume "\<not> ?thesis"  | 
| 19736 | 39  | 
then have contrapos: "x+y < u+v"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
40  | 
by simp  | 
| 14706 | 41  | 
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
 | 
42  | 
by (unfold nat2_to_nat_def) (simp add: Let_def)  | 
| 14706 | 43  | 
also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2"  | 
44  | 
by (simp only: div_mult_self1_is_m)  | 
|
45  | 
also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2  | 
|
46  | 
+ ((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
 | 
47  | 
proof -  | 
| 14706 | 48  | 
have "2 dvd (x+y)*Suc(x+y)"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
49  | 
by (rule dvd2_a_x_suc_a)  | 
| 19736 | 50  | 
then have "(x+y)*Suc(x+y) mod 2 = 0"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
51  | 
by (simp only: dvd_eq_mod_eq_0)  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
52  | 
also  | 
| 14706 | 53  | 
have "2 * Suc(x+y) mod 2 = 0"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
54  | 
by (rule mod_mult_self1_is_0)  | 
| 14706 | 55  | 
ultimately have  | 
56  | 
"((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0"  | 
|
57  | 
by simp  | 
|
| 19736 | 58  | 
then show ?thesis  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
59  | 
by simp  | 
| 14706 | 60  | 
qed  | 
61  | 
also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2"  | 
|
62  | 
by (rule div_add1_eq [symmetric])  | 
|
63  | 
also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2"  | 
|
64  | 
by (simp only: add_mult_distrib [symmetric])  | 
|
65  | 
also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2"  | 
|
66  | 
by (simp only: mult_le_mono div_le_mono)  | 
|
67  | 
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
 | 
68  | 
by (unfold nat2_to_nat_def) (simp add: Let_def)  | 
| 14706 | 69  | 
finally show ?thesis  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
70  | 
by (simp only: eq)  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
71  | 
qed  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
72  | 
|
| 14706 | 73  | 
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
 | 
74  | 
proof -  | 
| 14706 | 75  | 
  {
 | 
76  | 
fix u v x y assume "nat2_to_nat (u,v) = nat2_to_nat (x,y)"  | 
|
| 19736 | 77  | 
then have "u+v \<le> x+y" by (rule nat2_to_nat_help)  | 
| 14706 | 78  | 
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
 | 
79  | 
by (rule nat2_to_nat_help)  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
80  | 
finally have eq: "u+v = x+y" .  | 
| 14706 | 81  | 
with prems have ux: "u=x"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
82  | 
by (simp add: nat2_to_nat_def Let_def)  | 
| 14706 | 83  | 
with eq have vy: "v=y"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
84  | 
by simp  | 
| 14706 | 85  | 
with ux have "(u,v) = (x,y)"  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
86  | 
by simp  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
87  | 
}  | 
| 19736 | 88  | 
then have "\<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
 | 
89  | 
by fast  | 
| 19736 | 90  | 
then show ?thesis  | 
| 
14127
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
91  | 
by (unfold inj_on_def) simp  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
92  | 
qed  | 
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
40a4768c8e0b
new theory NatPair of the injection from nat*nat -> nat
 
paulson 
parents:  
diff
changeset
 | 
94  | 
end  |