src/ZF/Integ/Int.thy
author paulson
Fri, 25 Sep 1998 13:18:07 +0200
changeset 5561 426c1e330903
child 9333 5cacc383157a
permissions -rw-r--r--
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
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
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     9
Int = EquivClass + Arith +
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    10
consts
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    11
    intrel,int::      i
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    12
    int_of      ::      i=>i            ("$# _" [80] 80)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    13
    zminus      ::      i=>i            ("$~ _" [80] 80)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    14
    znegative   ::      i=>o
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    15
    zmagnitude  ::      i=>i
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    16
    "$*"        ::      [i,i]=>i      (infixl 70)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    17
    "$'/"       ::      [i,i]=>i      (infixl 70) 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    18
    "$'/'/"     ::      [i,i]=>i      (infixl 70)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    19
    "$+"        ::      [i,i]=>i      (infixl 65)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    20
    "$-"        ::      [i,i]=>i      (infixl 65)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    21
    "$<"        ::      [i,i]=>o        (infixl 50)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    22
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    23
defs
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    24
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    25
    intrel_def
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    26
     "intrel == {p:(nat*nat)*(nat*nat).                 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    27
        EX x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    28
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    29
    int_def   "int == (nat*nat)/intrel"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    30
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    31
    int_of_def  "$# m == intrel `` {<m,0>}"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    32
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    33
    zminus_def  "$~ Z == UN <x,y>:Z. intrel``{<y,x>}"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    34
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    35
    znegative_def
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    36
        "znegative(Z) == EX x y. x<y & y:nat & <x,y>:Z"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    37
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    38
    zmagnitude_def
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    39
        "zmagnitude(Z) ==
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    40
	 THE m. m : nat & ((~ znegative(Z) & Z = $# m) |
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    41
	                   (znegative(Z) & $~ Z = $# m))"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    42
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    43
    (*Cannot use UN<x1,y2> here or in zmult because of the form of congruent2.
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    44
      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
    45
      better.*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    46
    zadd_def
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    47
     "Z1 $+ Z2 == 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    48
       UN z1:Z1. UN z2:Z2. let <x1,y1>=z1; <x2,y2>=z2                 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    49
                           in intrel``{<x1#+x2, y1#+y2>}"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    50
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    51
    zdiff_def   "Z1 $- Z2 == Z1 $+ zminus(Z2)"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    52
    zless_def   "Z1 $< Z2 == znegative(Z1 $- Z2)"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    53
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    54
    (*This illustrates the primitive form of definitions (no patterns)*)
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    55
    zmult_def
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    56
     "Z1 $* Z2 == 
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    57
       UN p1:Z1. UN p2:Z2.  split(%x1 y1. split(%x2 y2.        
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    58
                   intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    59
    
426c1e330903 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    60
 end