| author | paulson |
| Mon, 23 Sep 1996 18:12:45 +0200 | |
| changeset 2009 | 9023e474d22a |
| parent 1559 | 9ba0906aa60d |
| child 2215 | ebf910e7ec87 |
| 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$ |
| 1476 | 3 |
Authors: Riccardo Mattolini, Dip. Sistemi e Informatica |
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 | 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 | 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 | 16 |
typedef (Integ) |
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 | 19 |
instance |
20 |
int :: {ord, plus, times, minus}
|
|
21 |
||
| 1559 | 22 |
constdefs |
23 |
||
| 1374 | 24 |
zNat :: nat set |
| 1559 | 25 |
"zNat == {x::nat. True}"
|
26 |
||
27 |
znat :: nat => int ("$# _" [80] 80)
|
|
28 |
"$# m == Abs_Integ(intrel ^^ {(m,0)})"
|
|
29 |
||
30 |
zminus :: int => int ("$~ _" [80] 80)
|
|
31 |
"$~ Z == Abs_Integ(UN p:Rep_Integ(Z). split (%x y. intrel^^{(y,x)}) p)"
|
|
32 |
||
| 1374 | 33 |
znegative :: int => bool |
| 1559 | 34 |
"znegative(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)" |
35 |
||
| 1374 | 36 |
zmagnitude :: int => int |
| 1559 | 37 |
"zmagnitude(Z) == Abs_Integ(UN p:Rep_Integ(Z). |
38 |
split (%x y. intrel^^{((y-x) + (x-y),0)}) p)"
|
|
39 |
||
40 |
zdiv :: [int,int]=>int (infixl 70) |
|
41 |
"Z1 zdiv Z2 == |
|
42 |
Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split (%x1 y1. |
|
43 |
split (%x2 y2. intrel^^{((x1-y1)div(x2-y2)+(y1-x1)div(y2-x2),
|
|
44 |
(x1-y1)div(y2-x2)+(y1-x1)div(x2-y2))}) p2) p1)" |
|
45 |
||
46 |
zmod :: [int,int]=>int (infixl 70) |
|
47 |
"Z1 zmod Z2 == |
|
48 |
Abs_Integ(UN p1:Rep_Integ(Z1).UN p2:Rep_Integ(Z2).split (%x1 y1. |
|
49 |
split (%x2 y2. intrel^^{((x1-y1)mod((x2-y2)+(y2-x2)),
|
|
50 |
(x1-y1)mod((x2-y2)+(x2-y2)))}) p2) p1)" |
|
51 |
||
52 |
zpred :: int=>int |
|
53 |
"zpred(Z) == Z - $# Suc(0)" |
|
54 |
||
55 |
zsuc :: int=>int |
|
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 | 60 |
"Z1 + Z2 == |
61 |
Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). |
|
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 | 71 |
"Z1 * Z2 == |
72 |
Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split (%x1 y1. |
|
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 |