1 (* Title: HOL/Library/NatPair.thy |
|
2 ID: $Id$ |
|
3 Author: Stefan Richter |
|
4 Copyright 2003 Technische Universitaet Muenchen |
|
5 *) |
|
6 |
|
7 header {* Pairs of Natural Numbers *} |
|
8 |
|
9 theory NatPair |
|
10 imports Main |
|
11 begin |
|
12 |
|
13 text{* |
|
14 A bijection between @{text "\<nat>\<twosuperior>"} and @{text \<nat>}. Definition |
|
15 and proofs are from \cite[page 85]{Oberschelp:1993}. |
|
16 *} |
|
17 |
|
18 definition nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" where |
|
19 "nat2_to_nat pair = (let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n)" |
|
20 definition nat_to_nat2:: "nat \<Rightarrow> (nat * nat)" where |
|
21 "nat_to_nat2 = inv nat2_to_nat" |
|
22 |
|
23 lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)" |
|
24 proof (cases "2 dvd a") |
|
25 case True |
|
26 then show ?thesis by (rule dvd_mult2) |
|
27 next |
|
28 case False |
|
29 then have "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0) |
|
30 then have "Suc a mod 2 = 0" by (simp add: mod_Suc) |
|
31 then have "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0) |
|
32 then show ?thesis by (rule dvd_mult) |
|
33 qed |
|
34 |
|
35 lemma |
|
36 assumes eq: "nat2_to_nat (u,v) = nat2_to_nat (x,y)" |
|
37 shows nat2_to_nat_help: "u+v \<le> x+y" |
|
38 proof (rule classical) |
|
39 assume "\<not> ?thesis" |
|
40 then have contrapos: "x+y < u+v" |
|
41 by simp |
|
42 have "nat2_to_nat (x,y) < (x+y) * Suc (x+y) div 2 + Suc (x + y)" |
|
43 by (unfold nat2_to_nat_def) (simp add: Let_def) |
|
44 also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2" |
|
45 by (simp only: div_mult_self1_is_m) |
|
46 also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2 |
|
47 + ((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2" |
|
48 proof - |
|
49 have "2 dvd (x+y)*Suc(x+y)" |
|
50 by (rule dvd2_a_x_suc_a) |
|
51 then have "(x+y)*Suc(x+y) mod 2 = 0" |
|
52 by (simp only: dvd_eq_mod_eq_0) |
|
53 also |
|
54 have "2 * Suc(x+y) mod 2 = 0" |
|
55 by (rule mod_mult_self1_is_0) |
|
56 ultimately have |
|
57 "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0" |
|
58 by simp |
|
59 then show ?thesis |
|
60 by simp |
|
61 qed |
|
62 also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2" |
|
63 by (rule div_add1_eq [symmetric]) |
|
64 also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2" |
|
65 by (simp only: add_mult_distrib [symmetric]) |
|
66 also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2" |
|
67 by (simp only: mult_le_mono div_le_mono) |
|
68 also have "\<dots> \<le> nat2_to_nat (u,v)" |
|
69 by (unfold nat2_to_nat_def) (simp add: Let_def) |
|
70 finally show ?thesis |
|
71 by (simp only: eq) |
|
72 qed |
|
73 |
|
74 theorem nat2_to_nat_inj: "inj nat2_to_nat" |
|
75 proof - |
|
76 { |
|
77 fix u v x y |
|
78 assume eq1: "nat2_to_nat (u,v) = nat2_to_nat (x,y)" |
|
79 then have "u+v \<le> x+y" by (rule nat2_to_nat_help) |
|
80 also from eq1 [symmetric] have "x+y \<le> u+v" |
|
81 by (rule nat2_to_nat_help) |
|
82 finally have eq2: "u+v = x+y" . |
|
83 with eq1 have ux: "u=x" |
|
84 by (simp add: nat2_to_nat_def Let_def) |
|
85 with eq2 have vy: "v=y" by simp |
|
86 with ux have "(u,v) = (x,y)" by simp |
|
87 } |
|
88 then have "\<And>x y. nat2_to_nat x = nat2_to_nat y \<Longrightarrow> x=y" by fast |
|
89 then show ?thesis unfolding inj_on_def by simp |
|
90 qed |
|
91 |
|
92 lemma nat_to_nat2_surj: "surj nat_to_nat2" |
|
93 by (simp only: nat_to_nat2_def nat2_to_nat_inj inj_imp_surj_inv) |
|
94 |
|
95 |
|
96 lemma gauss_sum_nat_upto: "2 * (\<Sum>i\<le>n::nat. i) = n * (n + 1)" |
|
97 using gauss_sum[where 'a = nat] |
|
98 by (simp add:atLeast0AtMost setsum_shift_lb_Suc0_0 numeral_2_eq_2) |
|
99 |
|
100 lemma nat2_to_nat_surj: "surj nat2_to_nat" |
|
101 proof (unfold surj_def) |
|
102 { |
|
103 fix z::nat |
|
104 def r \<equiv> "Max {r. (\<Sum>i\<le>r. i) \<le> z}" |
|
105 def x \<equiv> "z - (\<Sum>i\<le>r. i)" |
|
106 |
|
107 hence "finite {r. (\<Sum>i\<le>r. i) \<le> z}" |
|
108 by (simp add: lessThan_Suc_atMost[symmetric] lessThan_Suc finite_less_ub) |
|
109 also have "0 \<in> {r. (\<Sum>i\<le>r. i) \<le> z}" by simp |
|
110 hence "{r::nat. (\<Sum>i\<le>r. i) \<le> z} \<noteq> {}" by fast |
|
111 ultimately have a: "r \<in> {r. (\<Sum>i\<le>r. i) \<le> z} \<and> (\<forall>s \<in> {r. (\<Sum>i\<le>r. i) \<le> z}. s \<le> r)" |
|
112 by (simp add: r_def del:mem_Collect_eq) |
|
113 { |
|
114 assume "r<x" |
|
115 hence "r+1\<le>x" by simp |
|
116 hence "(\<Sum>i\<le>r. i)+(r+1)\<le>z" using x_def by arith |
|
117 hence "(r+1) \<in> {r. (\<Sum>i\<le>r. i) \<le> z}" by simp |
|
118 with a have "(r+1)\<le>r" by simp |
|
119 } |
|
120 hence b: "x\<le>r" by force |
|
121 |
|
122 def y \<equiv> "r-x" |
|
123 have "2*z=2*(\<Sum>i\<le>r. i)+2*x" using x_def a by simp arith |
|
124 also have "\<dots> = r * (r+1) + 2*x" using gauss_sum_nat_upto by simp |
|
125 also have "\<dots> = (x+y)*(x+y+1)+2*x" using y_def b by simp |
|
126 also { have "2 dvd ((x+y)*(x+y+1))" using dvd2_a_x_suc_a by simp } |
|
127 hence "\<dots> = 2 * nat2_to_nat(x,y)" |
|
128 using nat2_to_nat_def by (simp add: Let_def dvd_mult_div_cancel) |
|
129 finally have "z=nat2_to_nat (x, y)" by simp |
|
130 } |
|
131 thus "\<forall>y. \<exists>x. y = nat2_to_nat x" by fast |
|
132 qed |
|
133 |
|
134 end |
|