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