217
|
1 |
(* Title: Integ.thy
|
|
2 |
ID: $Id$
|
|
3 |
Authors: Riccardo Mattolini, Dip. Sistemi e Informatica
|
|
4 |
Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
5 |
Copyright 1994 Universita' di Firenze
|
|
6 |
Copyright 1993 University of Cambridge
|
|
7 |
|
|
8 |
The integers as equivalence classes over nat*nat.
|
|
9 |
*)
|
|
10 |
|
|
11 |
Integ = Equiv + Arith +
|
|
12 |
consts
|
|
13 |
intrel :: "((nat * nat) * (nat * nat)) set"
|
|
14 |
|
|
15 |
defs
|
|
16 |
intrel_def
|
|
17 |
"intrel == {p. ? x1 y1 x2 y2. p=<<x1::nat,y1>,<x2,y2>> & x1+y2 = x2+y1}"
|
|
18 |
|
|
19 |
subtype (Integ)
|
237
|
20 |
int = "{x::(nat*nat).True}/intrel" (Equiv.quotient_def)
|
217
|
21 |
|
237
|
22 |
instance
|
|
23 |
int :: {ord, plus, times, minus}
|
|
24 |
|
217
|
25 |
consts
|
|
26 |
zNat :: "nat set"
|
|
27 |
znat :: "nat => int" ("$# _" [80] 80)
|
|
28 |
zminus :: "int => int" ("$~ _" [80] 80)
|
|
29 |
znegative :: "int => bool"
|
|
30 |
zmagnitude :: "int => int"
|
|
31 |
zdiv,zmod :: "[int,int]=>int" (infixl 70)
|
|
32 |
zpred,zsuc :: "int=>int"
|
|
33 |
|
|
34 |
defs
|
|
35 |
zNat_def "zNat == {x::nat. True}"
|
|
36 |
|
|
37 |
znat_def "$# m == Abs_Integ(intrel ^^ {<m,0>})"
|
|
38 |
|
|
39 |
zminus_def
|
|
40 |
"$~ Z == Abs_Integ(UN p:Rep_Integ(Z). split(%x y. intrel^^{<y,x>},p))"
|
|
41 |
|
|
42 |
znegative_def
|
|
43 |
"znegative(Z) == EX x y. x<y & <x,y::nat>:Rep_Integ(Z)"
|
|
44 |
|
|
45 |
zmagnitude_def
|
|
46 |
"zmagnitude(Z) == Abs_Integ(UN p:Rep_Integ(Z).split(%x y. intrel^^{<(y-x) + (x-y),0>},p))"
|
|
47 |
|
|
48 |
zadd_def
|
|
49 |
"Z1 + Z2 == \
|
|
50 |
\ Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). \
|
|
51 |
\ split(%x1 y1. split(%x2 y2. intrel^^{<x1+x2, y1+y2>},p2),p1))"
|
|
52 |
|
|
53 |
zdiff_def "Z1 - Z2 == Z1 + zminus(Z2)"
|
|
54 |
|
|
55 |
zless_def "Z1<Z2 == znegative(Z1 - Z2)"
|
|
56 |
|
|
57 |
zle_def "Z1 <= (Z2::int) == ~(Z2 < Z1)"
|
|
58 |
|
|
59 |
zmult_def
|
|
60 |
"Z1 * Z2 == \
|
|
61 |
\ Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split(%x1 y1. \
|
|
62 |
\ split(%x2 y2. intrel^^{<x1*x2 + y1*y2, x1*y2 + y1*x2>},p2),p1))"
|
|
63 |
|
|
64 |
zdiv_def
|
|
65 |
"Z1 zdiv Z2 == \
|
|
66 |
\ Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split(%x1 y1. \
|
|
67 |
\ split(%x2 y2. intrel^^{<(x1-y1)div(x2-y2)+(y1-x1)div(y2-x2), \
|
|
68 |
\ (x1-y1)div(y2-x2)+(y1-x1)div(x2-y2)>},p2),p1))"
|
|
69 |
|
|
70 |
zmod_def
|
|
71 |
"Z1 zmod Z2 == \
|
|
72 |
\ Abs_Integ(UN p1:Rep_Integ(Z1).UN p2:Rep_Integ(Z2).split(%x1 y1. \
|
|
73 |
\ split(%x2 y2. intrel^^{<(x1-y1)mod((x2-y2)+(y2-x2)), \
|
|
74 |
\ (x1-y1)mod((x2-y2)+(x2-y2))>},p2),p1))"
|
|
75 |
|
|
76 |
zsuc_def "zsuc(Z) == Z + $# Suc(0)"
|
|
77 |
|
|
78 |
zpred_def "zpred(Z) == Z - $# Suc(0)"
|
|
79 |
end
|