src/HOL/Complex/NSCA.thy
author nipkow
Sun, 25 Jan 2004 00:42:22 +0100
changeset 14360 e654599b114e
parent 14320 fb7a114826be
child 14387 e96d5c42c4b0
permissions -rw-r--r--
Added an exception handler and error msg.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     1
(*  Title       : NSCA.thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Copyright   : 2001,2002 University of Edinburgh
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Description : Infinite, infinitesimal complex number etc! 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 13957
diff changeset
     7
NSCA = NSComplexArith + 
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     8
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     9
consts   
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
    (* infinitely close *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
    "@c="     :: [hcomplex,hcomplex] => bool  (infixl 50)  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
constdefs
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
   (* standard complex numbers reagarded as an embedded subset of NS complex *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
   SComplex  :: "hcomplex set"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
   "SComplex == {x. EX r. x = hcomplex_of_complex r}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
   CInfinitesimal  :: "hcomplex set"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
   "CInfinitesimal == {x. ALL r: Reals. 0 < r --> hcmod x < r}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
   CFinite :: "hcomplex set"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
   "CFinite == {x. EX r: Reals. hcmod x < r}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
   CInfinite :: "hcomplex set"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
   "CInfinite == {x. ALL r: Reals. r < hcmod x}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
   (* standard part map *)  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
   stc :: hcomplex => hcomplex
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
   "stc x == (@r. x : CFinite & r:SComplex & r @c= x)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
   cmonad    :: hcomplex => hcomplex set
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
   "cmonad x  == {y. x @c= y}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
   cgalaxy   :: hcomplex => hcomplex set
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
   "cgalaxy x == {y. (x - y) : CFinite}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
defs  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
   capprox_def  "x @c= y == (x - y) : CInfinitesimal"     
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
end