paulson@14127: (* Title: HOL/Library/NatPair.thy paulson@14127: ID: $Id$ paulson@14127: Author: Christophe Tabacznyj and Lawrence C Paulson paulson@14127: Copyright 1996 University of Cambridge paulson@14127: *) paulson@14127: paulson@14127: header {* paulson@14127: \title{Pairs of Natural Numbers} paulson@14127: \author{Stefan Richter} paulson@14127: *} paulson@14127: paulson@14127: theory NatPair = Main: paulson@14127: paulson@14127: text{*An injective function from $\mathbf{N}^2$ to paulson@14127: $\mathbf{N}$. Definition and proofs are from paulson@14127: Arnold Oberschelp. Rekursionstheorie. BI-Wissenschafts-Verlag, 1993 paulson@14127: (page 85). *} paulson@14127: paulson@14127: constdefs paulson@14127: nat2_to_nat:: "(nat * nat) \ nat" paulson@14127: "nat2_to_nat pair \ let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n" paulson@14127: paulson@14127: lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)" paulson@14127: proof (cases "2 dvd a") paulson@14127: case True paulson@14127: thus ?thesis by (rule dvd_mult2) paulson@14127: next paulson@14127: case False paulson@14127: hence "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0) paulson@14127: hence "Suc a mod 2 = 0" by (simp add: mod_Suc) paulson@14127: hence "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0) paulson@14127: thus ?thesis by (rule dvd_mult) paulson@14127: qed paulson@14127: paulson@14127: lemma assumes eq: "nat2_to_nat (u,v) = nat2_to_nat (x,y)" paulson@14127: shows nat2_to_nat_help: "u+v \ x+y" paulson@14127: proof (rule classical) paulson@14127: assume "\ ?thesis" paulson@14127: hence contrapos: "x+y < u+v" paulson@14127: by simp paulson@14127: have "nat2_to_nat (x,y) < (x+y) * Suc (x+y) div 2 + Suc (x + y)" paulson@14127: by (unfold nat2_to_nat_def) (simp add: Let_def) paulson@14127: also have "\ = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2" paulson@14127: by (simp only: div_mult_self1_is_m) paulson@14127: also have "\ = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2 paulson@14127: + ((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2" paulson@14127: proof - paulson@14127: have "2 dvd (x+y)*Suc(x+y)" paulson@14127: by (rule dvd2_a_x_suc_a) paulson@14127: hence "(x+y)*Suc(x+y) mod 2 = 0" paulson@14127: by (simp only: dvd_eq_mod_eq_0) paulson@14127: also paulson@14127: have "2 * Suc(x+y) mod 2 = 0" paulson@14127: by (rule mod_mult_self1_is_0) paulson@14127: ultimately have paulson@14127: "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0" paulson@14127: by simp paulson@14127: thus ?thesis paulson@14127: by simp paulson@14127: qed paulson@14127: also have "\ = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2" paulson@14127: by (rule div_add1_eq[THEN sym]) paulson@14127: also have "\ = ((x+y+2)*Suc(x+y)) div 2" paulson@14127: by (simp only: add_mult_distrib[THEN sym]) paulson@14127: also from contrapos have "\ \ ((Suc(u+v))*(u+v)) div 2" paulson@14127: by (simp only: mult_le_mono div_le_mono) paulson@14127: also have "\ \ nat2_to_nat (u,v)" paulson@14127: by (unfold nat2_to_nat_def) (simp add: Let_def) paulson@14127: finally show ?thesis paulson@14127: by (simp only: eq) paulson@14127: qed paulson@14127: paulson@14127: lemma nat2_to_nat_inj: "inj nat2_to_nat" paulson@14127: proof - paulson@14127: {fix u v x y assume "nat2_to_nat (u,v) = nat2_to_nat (x,y)" paulson@14127: hence "u+v \ x+y" by (rule nat2_to_nat_help) paulson@14127: also from prems[THEN sym] have "x+y \ u+v" paulson@14127: by (rule nat2_to_nat_help) paulson@14127: finally have eq: "u+v = x+y" . paulson@14127: with prems have ux: "u=x" paulson@14127: by (simp add: nat2_to_nat_def Let_def) paulson@14127: with eq have vy: "v=y" paulson@14127: by simp paulson@14127: with ux have "(u,v) = (x,y)" paulson@14127: by simp paulson@14127: } paulson@14127: hence "\x y. nat2_to_nat x = nat2_to_nat y \ x=y" paulson@14127: by fast paulson@14127: thus ?thesis paulson@14127: by (unfold inj_on_def) simp paulson@14127: qed paulson@14127: paulson@14127: end