author | wenzelm |
Fri, 05 Oct 2001 21:52:39 +0200 | |
changeset 11701 | 3d51fbf81c17 |
parent 11451 | 8abfb4f7bd02 |
child 11713 | 883d559b0b8c |
permissions | -rw-r--r-- |
5508 | 1 |
(* Title: IntDef.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
The integers as equivalence classes over nat*nat. |
|
7 |
*) |
|
8 |
||
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
10834
diff
changeset
|
9 |
IntDef = Equiv + NatArith + |
5508 | 10 |
constdefs |
11 |
intrel :: "((nat * nat) * (nat * nat)) set" |
|
9392 | 12 |
"intrel == {p. EX x1 y1 x2 y2. p=((x1::nat,y1),(x2,y2)) & x1+y2 = x2+y1}" |
5508 | 13 |
|
14 |
typedef (Integ) |
|
9392 | 15 |
int = "UNIV//intrel" (Equiv.quotient_def) |
5508 | 16 |
|
17 |
instance |
|
8937 | 18 |
int :: {ord, zero, plus, times, minus} |
5508 | 19 |
|
20 |
defs |
|
21 |
zminus_def |
|
10834 | 22 |
"- Z == Abs_Integ(UN (x,y):Rep_Integ(Z). intrel``{(y,x)})" |
5508 | 23 |
|
24 |
constdefs |
|
25 |
||
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
26 |
int :: nat => int |
10834 | 27 |
"int m == Abs_Integ(intrel `` {(m,0)})" |
5508 | 28 |
|
5540 | 29 |
neg :: int => bool |
30 |
"neg(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)" |
|
5508 | 31 |
|
32 |
(*For simplifying equalities*) |
|
33 |
iszero :: int => bool |
|
5582
a356fb49e69e
many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents:
5562
diff
changeset
|
34 |
"iszero z == z = int 0" |
5508 | 35 |
|
9392 | 36 |
defs (*of overloaded constants*) |
37 |
||
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11451
diff
changeset
|
38 |
Zero_int_def "0 == int 0" |
8937 | 39 |
|
5508 | 40 |
zadd_def |
5594 | 41 |
"z + w == |
7375
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
paulson
parents:
7127
diff
changeset
|
42 |
Abs_Integ(UN (x1,y1):Rep_Integ(z). UN (x2,y2):Rep_Integ(w). |
10834 | 43 |
intrel``{(x1+x2, y1+y2)})" |
5508 | 44 |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
5594
diff
changeset
|
45 |
zdiff_def "z - (w::int) == z + (-w)" |
5508 | 46 |
|
5594 | 47 |
zless_def "z<w == neg(z - w)" |
5508 | 48 |
|
5594 | 49 |
zle_def "z <= (w::int) == ~(w < z)" |
5508 | 50 |
|
51 |
zmult_def |
|
5594 | 52 |
"z * w == |
7375
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
paulson
parents:
7127
diff
changeset
|
53 |
Abs_Integ(UN (x1,y1):Rep_Integ(z). UN (x2,y2):Rep_Integ(w). |
10834 | 54 |
intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)})" |
5508 | 55 |
|
56 |
end |