author | paulson |
Thu, 20 Nov 1997 11:03:26 +0100 | |
changeset 4242 | 97601cf26262 |
parent 2618 | 15451c558a32 |
child 5491 | 22f8331cdf47 |
permissions | -rw-r--r-- |
1476 | 1 |
(* Title: Integ.thy |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
2 |
ID: $Id$ |
2224
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
2215 | 4 |
Copyright 1996 University of Cambridge |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
5 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
6 |
The integers as equivalence classes over nat*nat. |
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 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
9 |
Integ = Equiv + Arith + |
1559 | 10 |
constdefs |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
11 |
intrel :: "((nat * nat) * (nat * nat)) set" |
1559 | 12 |
"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
|
13 |
|
1476 | 14 |
typedef (Integ) |
15 |
int = "{x::(nat*nat).True}/intrel" (Equiv.quotient_def) |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
16 |
|
994 | 17 |
instance |
18 |
int :: {ord, plus, times, minus} |
|
19 |
||
1559 | 20 |
constdefs |
21 |
||
22 |
znat :: nat => int ("$# _" [80] 80) |
|
23 |
"$# m == Abs_Integ(intrel ^^ {(m,0)})" |
|
24 |
||
25 |
zminus :: int => int ("$~ _" [80] 80) |
|
26 |
"$~ Z == Abs_Integ(UN p:Rep_Integ(Z). split (%x y. intrel^^{(y,x)}) p)" |
|
27 |
||
1374 | 28 |
znegative :: int => bool |
1559 | 29 |
"znegative(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)" |
30 |
||
1374 | 31 |
zmagnitude :: int => int |
1559 | 32 |
"zmagnitude(Z) == Abs_Integ(UN p:Rep_Integ(Z). |
33 |
split (%x y. intrel^^{((y-x) + (x-y),0)}) p)" |
|
34 |
||
35 |
zpred :: int=>int |
|
36 |
"zpred(Z) == Z - $# Suc(0)" |
|
37 |
||
38 |
zsuc :: int=>int |
|
39 |
"zsuc(Z) == Z + $# Suc(0)" |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
40 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
41 |
defs |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
42 |
zadd_def |
1151 | 43 |
"Z1 + Z2 == |
44 |
Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). |
|
45 |
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
|
46 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
47 |
zdiff_def "Z1 - Z2 == Z1 + zminus(Z2)" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
48 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
49 |
zless_def "Z1<Z2 == znegative(Z1 - Z2)" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
50 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
51 |
zle_def "Z1 <= (Z2::int) == ~(Z2 < Z1)" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
52 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
53 |
zmult_def |
1151 | 54 |
"Z1 * Z2 == |
55 |
Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split (%x1 y1. |
|
56 |
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
|
57 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
58 |
end |