| author | wenzelm | 
| Sat, 18 Mar 2000 19:10:02 +0100 | |
| changeset 8516 | f5f6a97ee43f | 
| parent 7334 | a90fc1e5fb19 | 
| child 8856 | 435187ffc64e | 
| permissions | -rw-r--r-- | 
| 7292 | 1 | (* Title: RealInt.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Jacques D. Fleuriot | |
| 4 | Copyright: 1999 University of Edinburgh | |
| 5 | Description: Embedding the integers in the reals | |
| 6 | *) | |
| 7 | ||
| 7334 | 8 | RealInt = RealOrd + Int + | 
| 7292 | 9 | |
| 10 | constdefs | |
| 11 | real_of_int :: int => real | |
| 12 | "real_of_int z == Abs_real(UN (i,j): Rep_Integ z. realrel ^^ | |
| 13 |                      {(preal_of_prat(prat_of_pnat(pnat_of_nat i)),
 | |
| 14 | preal_of_prat(prat_of_pnat(pnat_of_nat j)))})" | |
| 15 | ||
| 16 | ||
| 17 | end |