| author | wenzelm | 
| Wed, 07 Nov 2001 18:18:29 +0100 | |
| changeset 12094 | db9a3ad6e90e | 
| parent 11868 | 56db9f3a6b3e | 
| child 13487 | 1291c6375c29 | 
| 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  | 
|
| 
11713
 
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
18  | 
  int :: {ord, zero, one, 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  | 
|
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11713 
diff
changeset
 | 
34  | 
"iszero z == z = (0::int)"  | 
| 5508 | 35  | 
|
| 9392 | 36  | 
defs (*of overloaded constants*)  | 
37  | 
||
| 
11713
 
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
38  | 
Zero_int_def "0 == int 0"  | 
| 
 
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
39  | 
One_int_def "1 == int 1"  | 
| 8937 | 40  | 
|
| 5508 | 41  | 
zadd_def  | 
| 5594 | 42  | 
"z + w ==  | 
| 
7375
 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 
paulson 
parents: 
7127 
diff
changeset
 | 
43  | 
Abs_Integ(UN (x1,y1):Rep_Integ(z). UN (x2,y2):Rep_Integ(w).  | 
| 10834 | 44  | 
		 intrel``{(x1+x2, y1+y2)})"
 | 
| 5508 | 45  | 
|
| 
7127
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
5594 
diff
changeset
 | 
46  | 
zdiff_def "z - (w::int) == z + (-w)"  | 
| 5508 | 47  | 
|
| 5594 | 48  | 
zless_def "z<w == neg(z - w)"  | 
| 5508 | 49  | 
|
| 5594 | 50  | 
zle_def "z <= (w::int) == ~(w < z)"  | 
| 5508 | 51  | 
|
52  | 
zmult_def  | 
|
| 5594 | 53  | 
"z * w ==  | 
| 
7375
 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 
paulson 
parents: 
7127 
diff
changeset
 | 
54  | 
Abs_Integ(UN (x1,y1):Rep_Integ(z). UN (x2,y2):Rep_Integ(w).  | 
| 10834 | 55  | 
		 intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)})"
 | 
| 5508 | 56  | 
|
57  | 
end  |