src/HOL/Real/Hyperreal/NSA.thy
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 10045 c76b73e16711
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10045
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     1
(*  Title       : NSA.thy
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     4
    Description : Infinite numbers, Infinitesimals,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     5
                  infinitely close relation etc.
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     6
*) 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     7
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     8
NSA = HRealAbs + RComplete +
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     9
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    10
constdefs
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    11
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    12
   (* standard real numbers reagarded as *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    13
   (* an embedded subset of hyperreals   *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    14
   SReal  :: "hypreal set"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    15
   "SReal == {x. EX r. x = hypreal_of_real r}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    16
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    17
   Infinitesimal  :: "hypreal set"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    18
   "Infinitesimal == {x. ALL r: SReal. 0 < r --> abs x < r}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    19
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    20
   HFinite :: "hypreal set"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    21
   "HFinite == {x. EX r: SReal. abs x < r}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    22
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    23
   HInfinite :: "hypreal set"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    24
   "HInfinite == {x. ALL r: SReal. r < abs x}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    25
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    26
consts   
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    27
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    28
    (* standard part map *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    29
    st       :: hypreal => hypreal
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    30
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    31
    (* infinitely close *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    32
    "@="     :: [hypreal,hypreal] => bool  (infixl 50)  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    33
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    34
    monad    :: hypreal => hypreal set
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    35
    galaxy   :: hypreal => hypreal set
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    36
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    37
defs  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    38
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    39
   inf_close_def  "x @= y       == (x + -y) : Infinitesimal"     
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    40
   st_def         "st           == (%x. @r. x : HFinite & r:SReal & r @= x)"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    41
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    42
   monad_def      "monad x      == {y. x @= y}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    43
   galaxy_def     "galaxy x     == {y. (x + -y) : HFinite}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    44
 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    45
end
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    46
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    47
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    48
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    49