src/HOL/Library/NatPair.thy
author haftmann
Thu, 19 Jul 2007 21:47:39 +0200
changeset 23854 688a8a7bcd4e
parent 23394 474ff28210c0
child 25594 43c718438f9f
permissions -rw-r--r--
uniform naming conventions for CG theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
3fd75e96145d corrected authors
nipkow
parents: 14127
diff changeset
     3
    Author:     Stefan Richter
3fd75e96145d corrected authors
nipkow
parents: 14127
diff changeset
     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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
     7
header {* Pairs of Natural Numbers *}
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
     9
theory NatPair
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports Main
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    11
begin
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    12
14706
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    13
text{*
19736
wenzelm
parents: 15140
diff changeset
    14
  An injective function from @{text "\<nat>\<twosuperior>"} to @{text \<nat>}.  Definition
wenzelm
parents: 15140
diff changeset
    15
  and proofs are from \cite[page 85]{Oberschelp:1993}.
14706
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    16
*}
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    17
19736
wenzelm
parents: 15140
diff changeset
    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
wenzelm
parents: 15140
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    22
lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
wenzelm
parents: 15140
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    27
  case False
19736
wenzelm
parents: 15140
diff changeset
    28
  then have "Suc (a mod 2) = 2" by (simp add: dvd_eq_mod_eq_0)
wenzelm
parents: 15140
diff changeset
    29
  then have "Suc a mod 2 = 0" by (simp add: mod_Suc)
wenzelm
parents: 15140
diff changeset
    30
  then have "2 dvd Suc a" by (simp only:dvd_eq_mod_eq_0)
wenzelm
parents: 15140
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    34
lemma
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
wenzelm
parents: 15140
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    43
  also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    44
    by (simp only: div_mult_self1_is_m)
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    45
  also have "\<dots> = (x+y)*Suc(x+y) div 2 + 2 * Suc(x+y) div 2
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
wenzelm
parents: 15140
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    55
    ultimately have
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    56
      "((x+y)*Suc(x+y) mod 2 + 2 * Suc(x+y) mod 2) div 2 = 0"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    57
      by simp
19736
wenzelm
parents: 15140
diff changeset
    58
    then show ?thesis
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    59
      by simp
14706
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    60
  qed
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    61
  also have "\<dots> = ((x+y)*Suc(x+y) + 2*Suc(x+y)) div 2"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    62
    by (rule div_add1_eq [symmetric])
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    63
  also have "\<dots> = ((x+y+2)*Suc(x+y)) div 2"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    64
    by (simp only: add_mult_distrib [symmetric])
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    65
  also from contrapos have "\<dots> \<le> ((Suc(u+v))*(u+v)) div 2"
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    66
    by (simp only: mult_le_mono div_le_mono)
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    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
71590b7733b7 tuned document;
wenzelm
parents: 14414
diff changeset
    75
  {
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    76
    fix u v x y
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    77
    assume eq1: "nat2_to_nat (u,v) = nat2_to_nat (x,y)"
19736
wenzelm
parents: 15140
diff changeset
    78
    then have "u+v \<le> x+y" by (rule nat2_to_nat_help)
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    79
    also from eq1 [symmetric] have "x+y \<le> u+v"
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    80
      by (rule nat2_to_nat_help)
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    81
    finally have eq2: "u+v = x+y" .
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    82
    with eq1 have ux: "u=x"
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    83
      by (simp add: nat2_to_nat_def Let_def)
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    84
    with eq2 have vy: "v=y" by simp
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    85
    with ux have "(u,v) = (x,y)" by simp
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    86
  }
23394
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    87
  then have "\<And>x y. nat2_to_nat x = nat2_to_nat y \<Longrightarrow> x=y" by fast
474ff28210c0 tuned proofs;
wenzelm
parents: 21404
diff changeset
    88
  then show ?thesis unfolding inj_on_def by simp
14127
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    89
qed
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    90
40a4768c8e0b new theory NatPair of the injection from nat*nat -> nat
paulson
parents:
diff changeset
    91
end