| author | paulson | 
| Fri, 24 May 2002 16:55:28 +0200 | |
| changeset 13178 | bc54319f6875 | 
| parent 12114 | a8e860c86252 | 
| child 13560 | d9651081578b | 
| permissions | -rw-r--r-- | 
| 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 | 9 | Int = EquivClass + ArithSimp + | 
| 9496 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 10 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 11 | constdefs | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 12 | intrel :: i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 13 |     "intrel == {p:(nat*nat)*(nat*nat).                 
 | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 16 | int :: i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 17 | "int == (nat*nat)//intrel" | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 18 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 20 |     "$# m == intrel `` {<natify(m), 0>}"
 | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 21 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 22 | intify :: i=>i (*coercion from ANYTHING to int*) | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 25 | raw_zminus :: i=>i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 27 | |
| 9548 | 28 |   zminus :: i=>i                                 ("$- _" [80] 80)
 | 
| 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: 
9333diff
changeset | 31 | znegative :: i=>o | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9548diff
changeset | 33 | |
| 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
9548diff
changeset | 34 | iszero :: i=>o | 
| 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
9548diff
changeset | 35 | "iszero(z) == z = $# 0" | 
| 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
9548diff
changeset | 36 | |
| 9883 | 37 | raw_nat_of :: i => i | 
| 38 | "raw_nat_of(z) == if znegative(z) then 0 | |
| 39 | else (THE m. m: nat & z = int_of(m))" | |
| 40 | ||
| 41 | nat_of :: i => i | |
| 42 | "nat_of(z) == raw_nat_of (intify(z))" | |
| 43 | ||
| 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: 
9333diff
changeset | 45 | zmagnitude :: i=>i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 46 | "zmagnitude(z) == | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 47 | THE m. m : nat & ((~ znegative(z) & z = $# m) | | 
| 9548 | 48 | (znegative(z) & $- z = $# m))" | 
| 9496 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 49 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 50 | raw_zmult :: [i,i]=>i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 54 | "raw_zmult(z1,z2) == | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
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: 
9333diff
changeset | 57 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 58 | zmult :: [i,i]=>i (infixl "$*" 70) | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 59 | "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 60 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 61 | raw_zadd :: [i,i]=>i | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 62 | "raw_zadd (z1, z2) == | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
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: 
9333diff
changeset | 65 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 66 | zadd :: [i,i]=>i (infixl "$+" 65) | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 67 | "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 68 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 69 | zdiff :: [i,i]=>i (infixl "$-" 65) | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 70 | "z1 $- z2 == z1 $+ zminus(z2)" | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 71 | |
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 72 | zless :: [i,i]=>o (infixl "$<" 50) | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 73 | "z1 $< z2 == znegative(z1 $- z2)" | 
| 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 74 | |
| 9548 | 75 | zle :: [i,i]=>o (infixl "$<=" 50) | 
| 9570 
e16e168984e1
installation of cancellation simprocs for the integers
 paulson parents: 
9548diff
changeset | 76 | "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)" | 
| 9548 | 77 | |
| 9654 
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
 paulson parents: 
9578diff
changeset | 78 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
11321diff
changeset | 79 | syntax (xsymbols) | 
| 11321 | 80 | "zmult" :: [i,i] => i (infixl "$\\<times>" 70) | 
| 81 | "zle" :: [i,i] => o (infixl "$\\<le>" 50) (*less than / equals*) | |
| 9654 
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
 paulson parents: 
9578diff
changeset | 82 | |
| 9964 | 83 | syntax (HTML output) | 
| 11321 | 84 | "zmult" :: [i,i] => i (infixl "$\\<times>" 70) | 
| 9496 
07e33cac5d9c
coercion "intify" to remove type constraints from integer algebraic laws
 paulson parents: 
9333diff
changeset | 85 | end |