| author | paulson |
| Wed, 13 Dec 2000 12:46:47 +0100 | |
| changeset 10662 | cf6be1804912 |
| parent 10607 | 352f6f209775 |
| permissions | -rw-r--r-- |
| 7218 | 1 |
(* Title : HOL/Real/Hyperreal/HyperDef.thy |
2 |
ID : $Id$ |
|
3 |
Author : Jacques D. Fleuriot |
|
4 |
Copyright : 1998 University of Cambridge |
|
5 |
Description : Construction of hyperreals using ultrafilters |
|
6 |
*) |
|
7 |
||
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
7292
diff
changeset
|
8 |
HyperDef = Filter + Real + |
| 7218 | 9 |
|
10 |
consts |
|
11 |
||
| 10568 | 12 |
FreeUltrafilterNat :: nat set set ("\\<U>")
|
| 7218 | 13 |
|
14 |
defs |
|
15 |
||
16 |
FreeUltrafilterNat_def |
|
17 |
"FreeUltrafilterNat == (@U. U : FreeUltrafilter (UNIV:: nat set))" |
|
18 |
||
19 |
||
20 |
constdefs |
|
21 |
hyprel :: "((nat=>real)*(nat=>real)) set" |
|
22 |
"hyprel == {p. ? X Y. p = ((X::nat=>real),Y) &
|
|
23 |
{n::nat. X(n) = Y(n)}: FreeUltrafilterNat}"
|
|
24 |
||
| 9391 | 25 |
typedef hypreal = "UNIV//hyprel" (Equiv.quotient_def) |
| 7218 | 26 |
|
27 |
instance |
|
| 10607 | 28 |
hypreal :: {ord, zero, plus, times, minus, inverse}
|
| 7218 | 29 |
|
30 |
consts |
|
31 |
||
32 |
"1hr" :: hypreal ("1hr")
|
|
33 |
"whr" :: hypreal ("whr")
|
|
34 |
"ehr" :: hypreal ("ehr")
|
|
35 |
||
36 |
||
37 |
defs |
|
38 |
||
| 10607 | 39 |
hypreal_zero_def |
40 |
"0 == Abs_hypreal(hyprel^^{%n::nat. (#0::real)})"
|
|
41 |
||
42 |
hypreal_one_def |
|
43 |
"1hr == Abs_hypreal(hyprel^^{%n::nat. (#1::real)})"
|
|
| 7218 | 44 |
|
45 |
(* an infinite number = [<1,2,3,...>] *) |
|
| 10607 | 46 |
omega_def |
47 |
"whr == Abs_hypreal(hyprel^^{%n::nat. real_of_posnat n})"
|
|
| 7218 | 48 |
|
49 |
(* an infinitesimal number = [<1,1/2,1/3,...>] *) |
|
| 10607 | 50 |
epsilon_def |
51 |
"ehr == Abs_hypreal(hyprel^^{%n::nat. inverse (real_of_posnat n)})"
|
|
| 7218 | 52 |
|
53 |
hypreal_minus_def |
|
| 9055 | 54 |
"- P == Abs_hypreal(UN X: Rep_hypreal(P). hyprel^^{%n::nat. - (X n)})"
|
| 7218 | 55 |
|
56 |
hypreal_diff_def |
|
57 |
"x - y == x + -(y::hypreal)" |
|
58 |
||
| 10607 | 59 |
hypreal_inverse_def |
60 |
"inverse P == Abs_hypreal(UN X: Rep_hypreal(P). |
|
61 |
hyprel^^{%n. if X n = #0 then #0 else inverse (X n)})"
|
|
62 |
||
63 |
hypreal_divide_def |
|
64 |
"P / Q::hypreal == P * inverse Q" |
|
65 |
||
| 7218 | 66 |
constdefs |
67 |
||
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
7292
diff
changeset
|
68 |
hypreal_of_real :: real => hypreal |
| 7218 | 69 |
"hypreal_of_real r == Abs_hypreal(hyprel^^{%n::nat. r})"
|
70 |
||
71 |
(* n::nat --> (n+1)::hypreal *) |
|
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
7292
diff
changeset
|
72 |
hypreal_of_posnat :: nat => hypreal |
| 7218 | 73 |
"hypreal_of_posnat n == (hypreal_of_real(real_of_preal |
74 |
(preal_of_prat(prat_of_pnat(pnat_of_nat n)))))" |
|
75 |
||
|
9013
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents:
7292
diff
changeset
|
76 |
hypreal_of_nat :: nat => hypreal |
| 7218 | 77 |
"hypreal_of_nat n == hypreal_of_posnat n + -1hr" |
78 |
||
79 |
defs |
|
80 |
||
81 |
hypreal_add_def |
|
82 |
"P + Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q). |
|
83 |
hyprel^^{%n::nat. X n + Y n})"
|
|
84 |
||
85 |
hypreal_mult_def |
|
86 |
"P * Q == Abs_hypreal(UN X:Rep_hypreal(P). UN Y:Rep_hypreal(Q). |
|
87 |
hyprel^^{%n::nat. X n * Y n})"
|
|
88 |
||
89 |
hypreal_less_def |
|
90 |
"P < (Q::hypreal) == EX X Y. X: Rep_hypreal(P) & |
|
91 |
Y: Rep_hypreal(Q) & |
|
92 |
{n::nat. X n < Y n} : FreeUltrafilterNat"
|
|
93 |
hypreal_le_def |
|
94 |
"P <= (Q::hypreal) == ~(Q < P)" |
|
95 |
||
96 |
end |
|
97 |
||
98 |