| author | berghofe |
| Fri, 09 Nov 2001 10:25:10 +0100 | |
| changeset 12125 | 316d11f760f7 |
| parent 10919 | 144ede948e58 |
| child 14269 | 502a7c95de73 |
| 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 |
||
| 8856 | 8 |
RealInt = RealOrd + |
| 7292 | 9 |
|
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
10 |
defs |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
11 |
(*overloaded*) |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
12 |
real_of_int_def |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
13 |
"real z == Abs_REAL(UN (i,j): Rep_Integ z. realrel `` |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
14 |
{(preal_of_prat(prat_of_pnat(pnat_of_nat i)),
|
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
15 |
preal_of_prat(prat_of_pnat(pnat_of_nat j)))})" |
| 7292 | 16 |
|
17 |
end |