src/ZF/Integ/Int.thy
author paulson
Thu, 07 Sep 2000 17:36:37 +0200
changeset 9883 c1c8647af477
parent 9654 9754ba005b64
child 9964 7966a2902266
permissions -rw-r--r--
a number of new theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     1
(*  Title:      ZF/Integ/Int.thy
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     2
    ID:         $Id$
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     5
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     6
The integers as equivalence classes over nat*nat.
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     7
*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     8
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
     9
Int = EquivClass + ArithSimp +
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    10
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    11
constdefs
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    12
  intrel :: i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    13
    "intrel == {p:(nat*nat)*(nat*nat).                 
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    14
                EX x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    15
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    16
  int :: i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    17
    "int == (nat*nat)//intrel"  
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    18
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    19
  int_of :: i=>i (*coercion from nat to int*)    ("$# _" [80] 80)
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    20
    "$# m == intrel `` {<natify(m), 0>}"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    21
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    22
  intify :: i=>i (*coercion from ANYTHING to int*) 
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    23
    "intify(m) == if m : int then m else $#0"
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    24
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    25
  raw_zminus :: i=>i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    26
    "raw_zminus(z) == UN <x,y>: z. intrel``{<y,x>}"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    27
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
    28
  zminus :: i=>i                                 ("$- _" [80] 80)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
    29
    "$- z == raw_zminus (intify(z))"
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    30
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    31
  znegative   ::      i=>o
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    32
    "znegative(z) == EX x y. x<y & y:nat & <x,y>:z"
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    33
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    34
  iszero      ::      i=>o
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    35
    "iszero(z) == z = $# 0"
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    36
    
9883
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    37
  raw_nat_of  :: i => i
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    38
  "raw_nat_of(z) == if znegative(z) then 0
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    39
                    else (THE m. m: nat & z = int_of(m))"
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    40
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    41
  nat_of  :: i => i
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    42
  "nat_of(z) == raw_nat_of (intify(z))"
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    43
c1c8647af477 a number of new theorems
paulson
parents: 9654
diff changeset
    44
  (*could be replaced by an absolute value function from int to int?*)
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    45
  zmagnitude  ::      i=>i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    46
    "zmagnitude(z) ==
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    47
     THE m. m : nat & ((~ znegative(z) & z = $# m) |
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
    48
		       (znegative(z) & $- z = $# m))"
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    49
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    50
  raw_zmult   ::      [i,i]=>i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    51
    (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    52
      Perhaps a "curried" or even polymorphic congruent predicate would be
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    53
      better.*)
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    54
     "raw_zmult(z1,z2) == 
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    55
       UN p1:z1. UN p2:z2.  split(%x1 y1. split(%x2 y2.        
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    56
                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    57
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    58
  zmult       ::      [i,i]=>i      (infixl "$*" 70)
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    59
     "z1 $* z2 == raw_zmult (intify(z1),intify(z2))"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    60
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    61
  raw_zadd    ::      [i,i]=>i
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    62
     "raw_zadd (z1, z2) == 
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    63
       UN z1:z1. UN z2:z2. let <x1,y1>=z1; <x2,y2>=z2                 
5561
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    64
                           in intrel``{<x1#+x2, y1#+y2>}"
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    65
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    66
  zadd        ::      [i,i]=>i      (infixl "$+" 65)
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    67
     "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    68
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    69
  zdiff        ::      [i,i]=>i      (infixl "$-" 65)
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    70
     "z1 $- z2 == z1 $+ zminus(z2)"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    71
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    72
  zless        ::      [i,i]=>o      (infixl "$<" 50)
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    73
     "z1 $< z2 == znegative(z1 $- z2)"
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    74
  
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
    75
  zle          ::      [i,i]=>o      (infixl "$<=" 50)
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    76
     "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)"
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9496
diff changeset
    77
  
9654
9754ba005b64 X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents: 9578
diff changeset
    78
9754ba005b64 X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents: 9578
diff changeset
    79
syntax (symbols)
9754ba005b64 X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents: 9578
diff changeset
    80
  "zmult"     :: [i,i] => i          (infixr "$\\<times>" 70)
9754ba005b64 X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents: 9578
diff changeset
    81
  "zle"       :: [i,i] => o          (infixl "$\\<le>" 50)  (*less than or equals*)
9754ba005b64 X-symbols for ordinal, cardinal, integer arithmetic
paulson
parents: 9578
diff changeset
    82
9496
07e33cac5d9c coercion "intify" to remove type constraints from integer algebraic laws
paulson
parents: 9333
diff changeset
    83
end