| author | wenzelm | 
| Sat, 18 Mar 2000 19:10:02 +0100 | |
| changeset 8516 | f5f6a97ee43f | 
| parent 7219 | 4e3f386c2e37 | 
| child 9043 | ca761fe227d8 | 
| permissions | -rw-r--r-- | 
| 5588 | 1 | (* Title : Real/RealDef.thy | 
| 7219 | 2 | ID : $Id$ | 
| 5588 | 3 | Author : Jacques D. Fleuriot | 
| 4 | Copyright : 1998 University of Cambridge | |
| 5 | Description : The reals | |
| 6 | *) | |
| 7 | ||
| 8 | RealDef = PReal + | |
| 9 | ||
| 10 | constdefs | |
| 11 | realrel :: "((preal * preal) * (preal * preal)) set" | |
| 12 |   "realrel == {p. ? x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" 
 | |
| 13 | ||
| 14 | typedef real = "{x::(preal*preal).True}/realrel"          (Equiv.quotient_def)
 | |
| 15 | ||
| 16 | ||
| 17 | instance | |
| 18 |    real  :: {ord, plus, times, minus}
 | |
| 19 | ||
| 20 | consts | |
| 21 | ||
| 22 |   "0r"       :: real               ("0r")   
 | |
| 23 |   "1r"       :: real               ("1r")  
 | |
| 24 | ||
| 25 | defs | |
| 26 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 27 | real_zero_def | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 28 |      "0r == Abs_real(realrel^^{(preal_of_prat(prat_of_pnat 1p),
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 29 | preal_of_prat(prat_of_pnat 1p))})" | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 30 | real_one_def | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 31 |      "1r == Abs_real(realrel^^{(preal_of_prat(prat_of_pnat 1p) + 
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 32 | preal_of_prat(prat_of_pnat 1p),preal_of_prat(prat_of_pnat 1p))})" | 
| 5588 | 33 | |
| 34 | real_minus_def | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 35 |     "- R ==  Abs_real(UN (x,y):Rep_real(R). realrel^^{(y,x)})"
 | 
| 5588 | 36 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 37 | real_diff_def "x - y == x + (- y :: real)" | 
| 5588 | 38 | |
| 39 | constdefs | |
| 40 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 41 | real_of_preal :: preal => real | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 42 | "real_of_preal m == | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 43 |            Abs_real(realrel^^{(m+preal_of_prat(prat_of_pnat 1p),
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 44 | preal_of_prat(prat_of_pnat 1p))})" | 
| 5588 | 45 | |
| 46 | rinv :: real => real | |
| 47 | "rinv(R) == (@S. R ~= 0r & S*R = 1r)" | |
| 48 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 49 | real_of_posnat :: nat => real | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 50 | "real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))" | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 51 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5787diff
changeset | 52 | real_of_nat :: nat => real | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 53 | "real_of_nat n == real_of_posnat n + (-1r)" | 
| 5588 | 54 | |
| 55 | defs | |
| 56 | ||
| 57 | real_add_def | |
| 58 | "P + Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q). | |
| 59 |                 split(%x1 y1. split(%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)"
 | |
| 60 | ||
| 61 | real_mult_def | |
| 62 | "P * Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q). | |
| 63 |                 split(%x1 y1. split(%x2 y2. realrel^^{(x1*x2+y1*y2,x1*y2+x2*y1)}) p2) p1)"
 | |
| 64 | ||
| 65 | real_less_def | |
| 66 | "P < Q == EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & | |
| 67 | (x1,y1):Rep_real(P) & | |
| 68 | (x2,y2):Rep_real(Q)" | |
| 69 | real_le_def | |
| 70 | "P <= (Q::real) == ~(Q < P)" | |
| 71 | ||
| 72 | end |