src/HOL/Library/Nat_Int_Bij.thy
author nipkow
Fri, 18 Sep 2009 14:40:24 +0200
changeset 32608 c0056c2c1d17
parent 32338 e73eb2db4727
child 32822 45fa9254ddc8
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28098
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Nat_Int_Bij.thy
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     2
    Author:     Stefan Richter, Tobias Nipkow
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     3
*)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     4
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     5
header{* Bijections $\mathbb{N}\to\mathbb{N}^2$ and $\mathbb{N}\to\mathbb{Z}$*}
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     6
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     7
theory Nat_Int_Bij
30663
0b6aff7451b2 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 29888
diff changeset
     8
imports Main
28098
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
     9
begin
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    10
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    11
subsection{*  A bijection between @{text "\<nat>"} and @{text "\<nat>\<twosuperior>"} *}
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    12
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    13
text{* Definition and proofs are from \cite[page 85]{Oberschelp:1993}. *}
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    14
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    15
definition nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" where
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    16
"nat2_to_nat pair = (let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    17
definition nat_to_nat2::  "nat \<Rightarrow> (nat * nat)" where
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    18
"nat_to_nat2 = inv nat2_to_nat"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    19
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    20
lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    21
proof (cases "2 dvd a")
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    22
  case True
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    23
  then show ?thesis by (rule dvd_mult2)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    24
next
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    25
  case False
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    26
  then have "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    27
  then have "Suc a mod 2 = 0" by (simp add: mod_Suc)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    28
  then have "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    29
  then show ?thesis by (rule dvd_mult)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    30
qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    31
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    32
lemma
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    33
  assumes eq: "nat2_to_nat (u,v) = nat2_to_nat (x,y)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    34
  shows nat2_to_nat_help: "u+v \<le> x+y"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    35
proof (rule classical)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    36
  assume "\<not> ?thesis"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    37
  then have contrapos: "x+y < u+v"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    38
    by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    39
  have "nat2_to_nat (x,y) < (x+y) * Suc (x+y) div 2 + Suc (x + y)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    40
    by (unfold nat2_to_nat_def) (simp add: Let_def)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    41
  also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    42
    by (simp only: div_mult_self1_is_m)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    43
  also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    44
    + ((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    45
  proof -
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    46
    have "2 dvd (x+y)*Suc(x+y)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    47
      by (rule dvd2_a_x_suc_a)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    48
    then have "(x+y)*Suc(x+y) mod 2 = 0"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    49
      by (simp only: dvd_eq_mod_eq_0)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    50
    also
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    51
    have "2 * Suc(x+y) mod 2 = 0"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    52
      by (rule mod_mult_self1_is_0)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    53
    ultimately have
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    54
      "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    55
      by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    56
    then show ?thesis
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    57
      by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    58
  qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    59
  also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    60
    by (rule div_add1_eq [symmetric])
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    61
  also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    62
    by (simp only: add_mult_distrib [symmetric])
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    63
  also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    64
    by (simp only: mult_le_mono div_le_mono)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    65
  also have "\<dots> \<le> nat2_to_nat (u,v)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    66
    by (unfold nat2_to_nat_def) (simp add: Let_def)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    67
  finally show ?thesis
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    68
    by (simp only: eq)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    69
qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    70
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    71
theorem nat2_to_nat_inj: "inj nat2_to_nat"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    72
proof -
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    73
  {
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    74
    fix u v x y
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    75
    assume eq1: "nat2_to_nat (u,v) = nat2_to_nat (x,y)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    76
    then have "u+v \<le> x+y" by (rule nat2_to_nat_help)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    77
    also from eq1 [symmetric] have "x+y \<le> u+v"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    78
      by (rule nat2_to_nat_help)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    79
    finally have eq2: "u+v = x+y" .
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    80
    with eq1 have ux: "u=x"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    81
      by (simp add: nat2_to_nat_def Let_def)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    82
    with eq2 have vy: "v=y" by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    83
    with ux have "(u,v) = (x,y)" by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    84
  }
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    85
  then have "\<And>x y. nat2_to_nat x = nat2_to_nat y \<Longrightarrow> x=y" by fast
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    86
  then show ?thesis unfolding inj_on_def by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    87
qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    88
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    89
lemma nat_to_nat2_surj: "surj nat_to_nat2"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    90
by (simp only: nat_to_nat2_def nat2_to_nat_inj inj_imp_surj_inv)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    91
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    92
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    93
lemma gauss_sum_nat_upto: "2 * (\<Sum>i\<le>n::nat. i) = n * (n + 1)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    94
using gauss_sum[where 'a = nat]
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    95
by (simp add:atLeast0AtMost setsum_shift_lb_Suc0_0 numeral_2_eq_2)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    96
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    97
lemma nat2_to_nat_surj: "surj nat2_to_nat"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    98
proof (unfold surj_def)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
    99
  {
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   100
    fix z::nat 
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   101
    def r \<equiv> "Max {r. (\<Sum>i\<le>r. i) \<le> z}" 
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   102
    def x \<equiv> "z - (\<Sum>i\<le>r. i)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   103
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   104
    hence "finite  {r. (\<Sum>i\<le>r. i) \<le> z}"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   105
      by (simp add: lessThan_Suc_atMost[symmetric] lessThan_Suc finite_less_ub)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   106
    also have "0 \<in> {r. (\<Sum>i\<le>r. i) \<le> z}"  by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   107
    hence "{r::nat. (\<Sum>i\<le>r. i) \<le> z} \<noteq> {}"  by fast
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   108
    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)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   109
      by (simp add: r_def del:mem_Collect_eq)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   110
    {
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   111
      assume "r<x"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   112
      hence "r+1\<le>x"  by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   113
      hence "(\<Sum>i\<le>r. i)+(r+1)\<le>z"  using x_def by arith
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   114
      hence "(r+1) \<in>  {r. (\<Sum>i\<le>r. i) \<le> z}"  by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   115
      with a have "(r+1)\<le>r"  by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   116
    }
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   117
    hence b: "x\<le>r"  by force
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   118
    
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   119
    def y \<equiv> "r-x"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   120
    have "2*z=2*(\<Sum>i\<le>r. i)+2*x"  using x_def a by simp arith
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   121
    also have "\<dots> = r * (r+1) + 2*x"   using gauss_sum_nat_upto by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   122
    also have "\<dots> = (x+y)*(x+y+1)+2*x" using y_def b by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   123
    also { have "2 dvd ((x+y)*(x+y+1))"	using dvd2_a_x_suc_a by simp }
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   124
    hence "\<dots> = 2 * nat2_to_nat(x,y)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   125
      using nat2_to_nat_def by (simp add: Let_def dvd_mult_div_cancel)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   126
    finally have "z=nat2_to_nat (x, y)"  by simp
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   127
  }
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   128
  thus "\<forall>y. \<exists>x. y = nat2_to_nat x"  by fast
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   129
qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   130
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   131
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   132
subsection{*  A bijection between @{text "\<nat>"} and @{text "\<int>"} *}
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   133
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   134
definition nat_to_int_bij :: "nat \<Rightarrow> int" where
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   135
"nat_to_int_bij n = (if 2 dvd n then int(n div 2) else -int(Suc n div 2))"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   136
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   137
definition int_to_nat_bij :: "int \<Rightarrow> nat" where
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   138
"int_to_nat_bij i = (if 0<=i then 2*nat(i) else 2*nat(-i) - 1)"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   139
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   140
lemma  i2n_n2i_id: "int_to_nat_bij (nat_to_int_bij n) = n"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   141
by (simp add: int_to_nat_bij_def nat_to_int_bij_def) presburger
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   142
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   143
lemma n2i_i2n_id: "nat_to_int_bij(int_to_nat_bij i) = i"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   144
proof -
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   145
  have "ALL m n::nat. m>0 \<longrightarrow> 2 * m - Suc 0 \<noteq> 2 * n" by presburger
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   146
  thus ?thesis
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   147
    by(simp add: nat_to_int_bij_def int_to_nat_bij_def, simp add:dvd_def)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   148
qed
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   149
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   150
lemma inv_nat_to_int_bij: "inv nat_to_int_bij = int_to_nat_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   151
by (simp add: i2n_n2i_id inv_equality n2i_i2n_id)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   152
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   153
lemma inv_int_to_nat_bij: "inv int_to_nat_bij = nat_to_int_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   154
by (simp add: i2n_n2i_id inv_equality n2i_i2n_id)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   155
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   156
lemma surj_nat_to_int_bij: "surj nat_to_int_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   157
by (blast intro: n2i_i2n_id surjI)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   158
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   159
lemma surj_int_to_nat_bij: "surj int_to_nat_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   160
by (blast intro: i2n_n2i_id surjI)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   161
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   162
lemma inj_nat_to_int_bij: "inj nat_to_int_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   163
by(simp add:inv_int_to_nat_bij[symmetric] surj_int_to_nat_bij surj_imp_inj_inv)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   164
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   165
lemma inj_int_to_nat_bij: "inj int_to_nat_bij"
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   166
by(simp add:inv_nat_to_int_bij[symmetric] surj_nat_to_int_bij surj_imp_inj_inv)
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   167
32338
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   168
lemma bij_nat_to_int_bij: "bij nat_to_int_bij"
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   169
by(simp add:bij_def inj_nat_to_int_bij surj_nat_to_int_bij)
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   170
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   171
lemma bij_int_to_nat_bij: "bij int_to_nat_bij"
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   172
by(simp add:bij_def inj_int_to_nat_bij surj_int_to_nat_bij)
e73eb2db4727 added bij lemmas
nipkow
parents: 30663
diff changeset
   173
28098
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   174
c92850d2d16c Replaced Library/NatPair by Nat_Int_Bij.
nipkow
parents:
diff changeset
   175
end