| 10751 |      1 | (*  Title       : Star.thy
 | 
|  |      2 |     Author      : Jacques D. Fleuriot
 | 
|  |      3 |     Copyright   : 1998  University of Cambridge
 | 
|  |      4 |     Description : defining *-transforms in NSA which extends sets of reals, 
 | 
|  |      5 |                   and real=>real functions
 | 
|  |      6 | *) 
 | 
|  |      7 | 
 | 
|  |      8 | Star = NSA +
 | 
|  |      9 | 
 | 
|  |     10 | constdefs
 | 
|  |     11 |     (* nonstandard extension of sets *)
 | 
|  |     12 |     starset :: real set => hypreal set          ("*s* _" [80] 80)
 | 
|  |     13 |     "*s* A  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}"
 | 
|  |     14 | 
 | 
|  |     15 |     (* internal sets *)
 | 
|  |     16 |     starset_n :: (nat => real set) => hypreal set        ("*sn* _" [80] 80)
 | 
|  |     17 |     "*sn* As  == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}"   
 | 
|  |     18 |     
 | 
|  |     19 |     InternalSets :: "hypreal set set"
 | 
|  |     20 |     "InternalSets == {X. EX As. X = *sn* As}"
 | 
|  |     21 | 
 | 
|  |     22 |     (* nonstandard extension of function *)
 | 
|  |     23 |     is_starext  ::  [hypreal => hypreal, real => real] => bool
 | 
|  |     24 |     "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y).
 | 
|  |     25 |                         ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
 | 
|  |     26 |     
 | 
|  |     27 |     starfun :: (real => real) => hypreal => hypreal        ("*f* _" [80] 80)
 | 
|  |     28 |     "*f* f  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. f (X n)}))" 
 | 
|  |     29 | 
 | 
|  |     30 |     (* internal functions *)
 | 
|  |     31 |     starfun_n :: (nat => (real => real)) => hypreal => hypreal        ("*fn* _" [80] 80)
 | 
|  |     32 |     "*fn* F  == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel^^{%n. (F n)(X n)}))" 
 | 
|  |     33 | 
 | 
|  |     34 |     InternalFuns :: (hypreal => hypreal) set
 | 
|  |     35 |     "InternalFuns == {X. EX F. X = *fn* F}"
 | 
|  |     36 | end
 | 
|  |     37 | 
 | 
|  |     38 | 
 | 
|  |     39 | 
 |