| author | berghofe | 
| Thu, 10 Oct 2002 14:19:17 +0200 | |
| changeset 13636 | fdf7e9388be7 | 
| parent 13487 | 1291c6375c29 | 
| permissions | -rw-r--r-- | 
| 5078 | 1  | 
(* Title : PRat.thy  | 
| 7219 | 2  | 
ID : $Id$  | 
| 5078 | 3  | 
Author : Jacques D. Fleuriot  | 
4  | 
Copyright : 1998 University of Cambridge  | 
|
5  | 
Description : The positive rationals  | 
|
6  | 
*)  | 
|
7  | 
||
| 8856 | 8  | 
PRat = PNat +  | 
| 5078 | 9  | 
|
10  | 
constdefs  | 
|
11  | 
ratrel :: "((pnat * pnat) * (pnat * pnat)) set"  | 
|
12  | 
    "ratrel  ==  {p. ? x1 y1 x2 y2. p=((x1::pnat,y1),(x2,y2)) & x1*y2 = x2*y1}" 
 | 
|
13  | 
||
| 13487 | 14  | 
typedef prat = "UNIV//ratrel" (quotient_def)  | 
| 5078 | 15  | 
|
16  | 
instance  | 
|
17  | 
   prat  :: {ord,plus,times}
 | 
|
18  | 
||
19  | 
||
20  | 
constdefs  | 
|
21  | 
||
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents: 
5078 
diff
changeset
 | 
22  | 
prat_of_pnat :: pnat => prat  | 
| 10834 | 23  | 
  "prat_of_pnat m == Abs_prat(ratrel``{(m,Abs_pnat 1)})"
 | 
| 5078 | 24  | 
|
25  | 
qinv :: prat => prat  | 
|
| 10834 | 26  | 
  "qinv(Q)  == Abs_prat(UN (x,y):Rep_prat(Q). ratrel``{(y,x)})" 
 | 
| 5078 | 27  | 
|
28  | 
defs  | 
|
29  | 
||
30  | 
prat_add_def  | 
|
| 
7376
 
46f92a120af9
tidied, allowing pattern-matching in defs of prat_add and prat_mult
 
paulson 
parents: 
7219 
diff
changeset
 | 
31  | 
"P + Q == Abs_prat(UN (x1,y1):Rep_prat(P). UN (x2,y2):Rep_prat(Q).  | 
| 10834 | 32  | 
		     ratrel``{(x1*y2 + x2*y1, y1*y2)})"
 | 
| 5078 | 33  | 
|
34  | 
prat_mult_def  | 
|
| 
7376
 
46f92a120af9
tidied, allowing pattern-matching in defs of prat_add and prat_mult
 
paulson 
parents: 
7219 
diff
changeset
 | 
35  | 
"P * Q == Abs_prat(UN (x1,y1):Rep_prat(P). UN (x2,y2):Rep_prat(Q).  | 
| 10834 | 36  | 
		     ratrel``{(x1*x2, y1*y2)})"
 | 
| 5078 | 37  | 
|
38  | 
(*** Gleason p. 119 ***)  | 
|
39  | 
prat_less_def  | 
|
| 9391 | 40  | 
"P < (Q::prat) == EX T. P + T = Q"  | 
| 5078 | 41  | 
|
42  | 
prat_le_def  | 
|
43  | 
"P <= (Q::prat) == ~(Q < P)"  | 
|
44  | 
||
45  | 
end  | 
|
46  | 
||
47  | 
||
48  | 
||
49  | 
||
50  |