src/HOL/Real/RealDef.thy
author paulson
Sat, 30 Dec 2000 22:13:18 +0100
changeset 10752 c4f1bf2acf4c
parent 10648 a8c647cfa31f
child 10778 2c6605049646
permissions -rw-r--r--
tidying, and separation of HOL-Hyperreal from HOL-Real
paulson@5588
     1
(*  Title       : Real/RealDef.thy
paulson@7219
     2
    ID          : $Id$
paulson@5588
     3
    Author      : Jacques D. Fleuriot
paulson@5588
     4
    Copyright   : 1998  University of Cambridge
paulson@5588
     5
    Description : The reals
paulson@5588
     6
*) 
paulson@5588
     7
paulson@5588
     8
RealDef = PReal +
paulson@5588
     9
paulson@10752
    10
instance preal :: order (preal_le_refl,preal_le_trans,preal_le_anti_sym,
paulson@10752
    11
                         preal_less_le)
paulson@10752
    12
paulson@5588
    13
constdefs
paulson@5588
    14
  realrel   ::  "((preal * preal) * (preal * preal)) set"
paulson@5588
    15
  "realrel == {p. ? x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" 
paulson@5588
    16
paulson@9391
    17
typedef real = "UNIV//realrel"  (Equiv.quotient_def)
paulson@5588
    18
paulson@5588
    19
paulson@5588
    20
instance
bauerg@10606
    21
   real  :: {ord, zero, plus, times, minus, inverse}
paulson@5588
    22
paulson@5588
    23
consts 
paulson@10752
    24
  "1r"   :: real               ("1r")  
paulson@5588
    25
paulson@10752
    26
   (*Overloaded constant: denotes the Real subset of enclosing types such as
paulson@10752
    27
     hypreal and complex*)
paulson@10752
    28
   SReal :: "'a set"
paulson@5588
    29
paulson@5588
    30
defs
paulson@5588
    31
paulson@7077
    32
  real_zero_def  
bauerg@10606
    33
  "0 == Abs_real(realrel^^{(preal_of_prat(prat_of_pnat 1p),
paulson@7077
    34
                                preal_of_prat(prat_of_pnat 1p))})"
paulson@7077
    35
  real_one_def   
bauerg@10606
    36
  "1r == Abs_real(realrel^^{(preal_of_prat(prat_of_pnat 1p) + 
paulson@7077
    37
            preal_of_prat(prat_of_pnat 1p),preal_of_prat(prat_of_pnat 1p))})"
paulson@5588
    38
paulson@5588
    39
  real_minus_def
bauerg@10606
    40
  "- R ==  Abs_real(UN (x,y):Rep_real(R). realrel^^{(y,x)})"
paulson@5588
    41
bauerg@10606
    42
  real_diff_def
bauerg@10606
    43
  "R - (S::real) == R + - S"
paulson@5588
    44
bauerg@10606
    45
  real_inverse_def
paulson@10648
    46
  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1r)"
bauerg@10606
    47
bauerg@10606
    48
  real_divide_def
bauerg@10606
    49
  "R / (S::real) == R * inverse S"
bauerg@10606
    50
  
paulson@5588
    51
constdefs
paulson@5588
    52
paulson@7077
    53
  real_of_preal :: preal => real            
paulson@7077
    54
  "real_of_preal m     ==
paulson@7077
    55
           Abs_real(realrel^^{(m+preal_of_prat(prat_of_pnat 1p),
paulson@7077
    56
                               preal_of_prat(prat_of_pnat 1p))})"
paulson@5588
    57
paulson@7077
    58
  real_of_posnat :: nat => real             
paulson@7077
    59
  "real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))"
paulson@7077
    60
paulson@7077
    61
  real_of_nat :: nat => real          
paulson@7127
    62
  "real_of_nat n    == real_of_posnat n + (-1r)"
paulson@5588
    63
paulson@5588
    64
defs
paulson@5588
    65
paulson@5588
    66
  real_add_def  
paulson@10752
    67
  "P+Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q).
paulson@10752
    68
                   (%(x1,y1). (%(x2,y2). realrel^^{(x1+x2, y1+y2)}) p2) p1)"
paulson@5588
    69
  
paulson@5588
    70
  real_mult_def  
paulson@10752
    71
  "P*Q == Abs_real(UN p1:Rep_real(P). UN p2:Rep_real(Q).
paulson@10752
    72
                   (%(x1,y1). (%(x2,y2). realrel^^{(x1*x2+y1*y2,x1*y2+x2*y1)})
paulson@10752
    73
		   p2) p1)"
paulson@5588
    74
paulson@5588
    75
  real_less_def
paulson@10752
    76
  "P<Q == EX x1 y1 x2 y2. x1 + y2 < x2 + y1 & 
paulson@10752
    77
                            (x1,y1):Rep_real(P) & (x2,y2):Rep_real(Q)" 
paulson@5588
    78
  real_le_def
paulson@5588
    79
  "P <= (Q::real) == ~(Q < P)"
paulson@5588
    80
paulson@10752
    81
syntax (symbols)
paulson@10752
    82
  SReal     :: "'a set"                   ("\\<real>")
paulson@10752
    83
paulson@5588
    84
end