src/HOL/Integ/Integ.thy
author paulson
Mon, 23 Sep 1996 18:12:45 +0200
changeset 2009 9023e474d22a
parent 1559 9ba0906aa60d
child 2215 ebf910e7ec87
permissions -rw-r--r--
Addition of le_refl to default simpset/claset
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     1
(*  Title:      Integ.thy
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     3
    Authors:    Riccardo Mattolini, Dip. Sistemi e Informatica
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     4
                Lawrence C Paulson, Cambridge University Computer Laboratory
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     5
    Copyright   1994 Universita' di Firenze
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     6
    Copyright   1993  University of Cambridge
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     7
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     8
The integers as equivalence classes over nat*nat.
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     9
*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    10
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    11
Integ = Equiv + Arith +
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    12
constdefs
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    13
  intrel      :: "((nat * nat) * (nat * nat)) set"
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    14
  "intrel == {p. ? x1 y1 x2 y2. p=((x1::nat,y1),(x2,y2)) & x1+y2 = x2+y1}"
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    15
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
    16
typedef (Integ)
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
    17
  int = "{x::(nat*nat).True}/intrel"            (Equiv.quotient_def)
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    18
994
b5e3fa9664fe replaced 'arities' by 'instance';
wenzelm
parents: 972
diff changeset
    19
instance
b5e3fa9664fe replaced 'arities' by 'instance';
wenzelm
parents: 972
diff changeset
    20
  int :: {ord, plus, times, minus}
b5e3fa9664fe replaced 'arities' by 'instance';
wenzelm
parents: 972
diff changeset
    21
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    22
constdefs
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    23
1374
5e407f2a3323 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    24
  zNat        :: nat set
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    25
  "zNat == {x::nat. True}"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    26
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    27
  znat        :: nat => int                                  ("$# _" [80] 80)
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    28
  "$# m == Abs_Integ(intrel ^^ {(m,0)})"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    29
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    30
  zminus      :: int => int                                  ("$~ _" [80] 80)
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    31
  "$~ Z == Abs_Integ(UN p:Rep_Integ(Z). split (%x y. intrel^^{(y,x)}) p)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    32
1374
5e407f2a3323 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    33
  znegative   :: int => bool
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    34
  "znegative(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    35
1374
5e407f2a3323 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    36
  zmagnitude  :: int => int
1559
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    37
  "zmagnitude(Z) == Abs_Integ(UN p:Rep_Integ(Z).
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    38
                              split (%x y. intrel^^{((y-x) + (x-y),0)}) p)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    39
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    40
  zdiv        :: [int,int]=>int                              (infixl 70)
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    41
  "Z1 zdiv Z2 ==   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    42
      Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split (%x1 y1.   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    43
          split (%x2 y2. intrel^^{((x1-y1)div(x2-y2)+(y1-x1)div(y2-x2),   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    44
          (x1-y1)div(y2-x2)+(y1-x1)div(x2-y2))}) p2) p1)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    45
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    46
  zmod        :: [int,int]=>int                              (infixl 70)
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    47
  "Z1 zmod Z2 ==   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    48
      Abs_Integ(UN p1:Rep_Integ(Z1).UN p2:Rep_Integ(Z2).split (%x1 y1.   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    49
          split (%x2 y2. intrel^^{((x1-y1)mod((x2-y2)+(y2-x2)),   
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    50
          (x1-y1)mod((x2-y2)+(x2-y2)))}) p2) p1)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    51
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    52
  zpred       :: int=>int
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    53
  "zpred(Z) == Z - $# Suc(0)"
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    54
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    55
  zsuc        :: int=>int
9ba0906aa60d added constdefs section
clasohm
parents: 1476
diff changeset
    56
  "zsuc(Z) == Z + $# Suc(0)"
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    57
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    58
defs
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    59
  zadd_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    60
   "Z1 + Z2 == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    61
       Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2).   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    62
           split (%x1 y1. split (%x2 y2. intrel^^{(x1+x2, y1+y2)}) p2) p1)"
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    63
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    64
  zdiff_def "Z1 - Z2 == Z1 + zminus(Z2)"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    65
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    66
  zless_def "Z1<Z2 == znegative(Z1 - Z2)"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    67
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    68
  zle_def   "Z1 <= (Z2::int) == ~(Z2 < Z1)"
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    69
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    70
  zmult_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    71
   "Z1 * Z2 == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    72
       Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split (%x1 y1.   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 994
diff changeset
    73
           split (%x2 y2. intrel^^{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)"
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    74
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    75
end