src/HOL/Real/Hyperreal/Star.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       : Star.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 : defining *-transforms in NSA which extends sets of reals, 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     5
                  and real=>real functions
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
Star = NSA +
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
    (* nonstandard extension of sets *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    12
    starset :: real set => hypreal set          ("*s* _" [80] 80)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    13
    "*s* A  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    14
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    15
    (* internal sets *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    16
    starset_n :: (nat => real set) => hypreal set        ("*sn* _" [80] 80)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    17
    "*sn* As  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    18
    
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    19
    InternalSets :: "hypreal set set"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    20
    "InternalSets == {X. EX As. X = *sn* As}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    21
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    22
    (* nonstandard extension of function *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    23
    is_starext  ::  [hypreal => hypreal, real => real] => bool
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    24
    "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y).
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    25
                        ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    26
    
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    27
    starfun :: (real => real) => hypreal => hypreal        ("*f* _" [80] 80)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    28
    "*f* f  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. f (X n)}))" 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    29
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    30
    (* internal functions *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    31
    starfun_n :: (nat => (real => real)) => hypreal => hypreal        ("*fn* _" [80] 80)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    32
    "*fn* F  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. (F n)(X n)}))" 
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
    InternalFuns :: (hypreal => hypreal) set
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    35
    "InternalFuns == {X. EX F. X = *fn* F}"
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    36
end
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    37
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