mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 17:21:59 +03:00
[ base ] Add bindings for ieee Double
number consts (#3116)
This commit is contained in:
parent
a945b5d2df
commit
d80bc1537d
@ -239,6 +239,9 @@
|
|||||||
|
|
||||||
* Adds `Semigroup`, `Applicative`, `Traversable` and `Zippable` for `Data.These`.
|
* Adds `Semigroup`, `Applicative`, `Traversable` and `Zippable` for `Data.These`.
|
||||||
|
|
||||||
|
* Adds bindings for IEEE floating point constants NaN and (+/-) Inf, as well as
|
||||||
|
machine epsilon and unit roundoff. Speeds vary depending on backend.
|
||||||
|
|
||||||
#### System
|
#### System
|
||||||
|
|
||||||
* Changes `getNProcessors` to return the number of online processors rather than
|
* Changes `getNProcessors` to return the number of online processors rather than
|
||||||
|
34
libs/base/Data/Double.idr
Normal file
34
libs/base/Data/Double.idr
Normal file
@ -0,0 +1,34 @@
|
|||||||
|
||| Various IEEE floating-point number constants
|
||||||
|
module Data.Double
|
||||||
|
|
||||||
|
||| Largest number that can be added to a floating-point number without changing
|
||||||
|
||| its value, i.e. `1.0 + unitRoundoff == 1.0`.
|
||||||
|
%foreign "scheme:blodwen-calcFlonumUnitRoundoff"
|
||||||
|
"node:lambda:()=>Number.EPSILON / 2"
|
||||||
|
export
|
||||||
|
unitRoundoff : Double
|
||||||
|
|
||||||
|
||| Machine epsilon is the smallest floating-point number that distinguishes two
|
||||||
|
||| floating-point numbers; the step size on the floating-point number line.
|
||||||
|
-- /!\ See `support/racket/support.rkt:42-45`
|
||||||
|
-- %foreign "scheme,chez:blodwen-calcFlonumEpsilon"
|
||||||
|
-- "scheme,racket:blodwen-flonumEpsilon"
|
||||||
|
%foreign "scheme:blodwen-calcFlonumEpsilon"
|
||||||
|
"node:lambda:()=>Number.EPSILON"
|
||||||
|
export
|
||||||
|
epsilon : Double
|
||||||
|
|
||||||
|
|
||||||
|
||| Not a number, e.g. `0.0 / 0.0`. Never equal to anything, including itself.
|
||||||
|
%foreign "scheme:blodwen-flonumNaN"
|
||||||
|
"node:lambda:()=>Number.NaN"
|
||||||
|
export
|
||||||
|
nan : Double
|
||||||
|
|
||||||
|
|
||||||
|
||| Positive Infinity. Can be negated to obtain Negative Infinity.
|
||||||
|
%foreign "scheme:blodwen-flonumInf"
|
||||||
|
"node:lambda:()=>Infinity"
|
||||||
|
export
|
||||||
|
inf : Double
|
||||||
|
|
@ -42,6 +42,7 @@ modules = Control.App,
|
|||||||
Data.Colist,
|
Data.Colist,
|
||||||
Data.Colist1,
|
Data.Colist1,
|
||||||
Data.Contravariant,
|
Data.Contravariant,
|
||||||
|
Data.Double,
|
||||||
Data.DPair,
|
Data.DPair,
|
||||||
Data.Either,
|
Data.Either,
|
||||||
Data.Fin,
|
Data.Fin,
|
||||||
|
@ -56,6 +56,7 @@ schHeader prof libs = fromString """
|
|||||||
(require ffi/unsafe ffi/unsafe/define) ; for calling C
|
(require ffi/unsafe ffi/unsafe/define) ; for calling C
|
||||||
\{ ifThenElse prof "(require profile)" "" }
|
\{ ifThenElse prof "(require profile)" "" }
|
||||||
(require racket/flonum) ; for float-typed transcendental functions
|
(require racket/flonum) ; for float-typed transcendental functions
|
||||||
|
(require math/flonum) ; for flonum constants
|
||||||
|
|
||||||
""" ++ libs ++ """
|
""" ++ libs ++ """
|
||||||
|
|
||||||
|
@ -40,6 +40,25 @@
|
|||||||
(if (> b 0) (+ r b) (- r b))
|
(if (> b 0) (+ r b) (- r b))
|
||||||
r)))
|
r)))
|
||||||
|
|
||||||
|
; flonum constants
|
||||||
|
|
||||||
|
(define (blodwen-calcFlonumUnitRoundoff)
|
||||||
|
(let loop [(uro 1.0)]
|
||||||
|
(if (fl= 1.0 (fl+ 1.0 uro))
|
||||||
|
uro
|
||||||
|
(loop (fl/ uro 2.0)))))
|
||||||
|
|
||||||
|
(define (blodwen-calcFlonumEpsilon)
|
||||||
|
(fl* (blodwen-calcFlonumUnitRoundoff) 2.0))
|
||||||
|
|
||||||
|
(define (blodwen-flonumNaN)
|
||||||
|
+nan.0)
|
||||||
|
|
||||||
|
(define (blodwen-flonumInf)
|
||||||
|
+inf.0)
|
||||||
|
|
||||||
|
; Bits
|
||||||
|
|
||||||
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
|
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
|
||||||
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
|
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
|
||||||
(define bu* (lambda (x y bits) (blodwen-toUnsignedInt (* x y) bits)))
|
(define bu* (lambda (x y bits) (blodwen-toUnsignedInt (* x y) bits)))
|
||||||
|
@ -37,6 +37,31 @@
|
|||||||
(if (> b 0) (+ r b) (- r b))
|
(if (> b 0) (+ r b) (- r b))
|
||||||
r)))
|
r)))
|
||||||
|
|
||||||
|
; flonum constants
|
||||||
|
|
||||||
|
;; /!\ this code is cursed for some reason?...
|
||||||
|
;;
|
||||||
|
;; (define (blodwen-flonumEpsilon)
|
||||||
|
;; epsilon.0)
|
||||||
|
|
||||||
|
(define (blodwen-calcFlonumUnitRoundoff)
|
||||||
|
;; (fl/ (blodwen-flonumEpsilon) 2.0))
|
||||||
|
(let loop [(uro 1.0)]
|
||||||
|
(if (fl= 1.0 (fl+ 1.0 uro))
|
||||||
|
uro
|
||||||
|
(loop (fl/ uro 2.0)))))
|
||||||
|
|
||||||
|
(define (blodwen-calcFlonumEpsilon)
|
||||||
|
(fl* (blodwen-calcFlonumUnitRoundoff) 2.0))
|
||||||
|
|
||||||
|
(define (blodwen-flonumNaN)
|
||||||
|
+nan.0)
|
||||||
|
|
||||||
|
(define (blodwen-flonumInf)
|
||||||
|
+inf.0)
|
||||||
|
|
||||||
|
; Bits
|
||||||
|
|
||||||
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
|
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
|
||||||
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
|
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
|
||||||
(define bu* (lambda (x y bits) (blodwen-toUnsignedInt (* x y) bits)))
|
(define bu* (lambda (x y bits) (blodwen-toUnsignedInt (* x y) bits)))
|
||||||
|
19
tests/allschemes/double001/URandEpsilon.idr
Normal file
19
tests/allschemes/double001/URandEpsilon.idr
Normal file
@ -0,0 +1,19 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- adding the unit roundoff to a Double should not modify it
|
||||||
|
|
||||||
|
testOnePlusUR : Bool
|
||||||
|
testOnePlusUR = 1.0 + unitRoundoff == 1.0
|
||||||
|
|
||||||
|
testURPlusOne : Bool
|
||||||
|
testURPlusOne = unitRoundoff + 1.0 == 1.0
|
||||||
|
|
||||||
|
testURComm : Bool
|
||||||
|
testURComm = testOnePlusUR == testURPlusOne
|
||||||
|
|
||||||
|
|
||||||
|
-- the machine epsilon should be double the unit roundoff
|
||||||
|
|
||||||
|
testEps2UR : Bool
|
||||||
|
testEps2UR = epsilon == unitRoundoff * 2.0
|
||||||
|
|
7
tests/allschemes/double001/expected
Normal file
7
tests/allschemes/double001/expected
Normal file
@ -0,0 +1,7 @@
|
|||||||
|
1/1: Building URandEpsilon (URandEpsilon.idr)
|
||||||
|
Main> Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
5
tests/allschemes/double001/input
Normal file
5
tests/allschemes/double001/input
Normal file
@ -0,0 +1,5 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn testOnePlusUR
|
||||||
|
:exec printLn testURPlusOne
|
||||||
|
:exec printLn testURComm
|
||||||
|
:exec printLn testEps2UR
|
3
tests/allschemes/double001/run
Executable file
3
tests/allschemes/double001/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 URandEpsilon.idr < input
|
28
tests/allschemes/double002/NaN.idr
Normal file
28
tests/allschemes/double002/NaN.idr
Normal file
@ -0,0 +1,28 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- undefined things are NaN
|
||||||
|
divZeroZero : Double
|
||||||
|
divZeroZero = 0.0 / 0.0
|
||||||
|
|
||||||
|
-- NaN with anything is NaN
|
||||||
|
|
||||||
|
nanPlus : Double
|
||||||
|
nanPlus = 1.0 + nan
|
||||||
|
|
||||||
|
nanSub : Double
|
||||||
|
nanSub = 1.0 - nan
|
||||||
|
|
||||||
|
nanMult : Double
|
||||||
|
nanMult = 2.0 * nan
|
||||||
|
|
||||||
|
nanDiv : Double
|
||||||
|
nanDiv = nan / 2.0
|
||||||
|
|
||||||
|
-- neg NaN is still NaN
|
||||||
|
negNaN : Double
|
||||||
|
negNaN = negate nan
|
||||||
|
|
||||||
|
-- NaNs are never equal
|
||||||
|
nansNotEq : Bool
|
||||||
|
nansNotEq = nan == nan
|
||||||
|
|
10
tests/allschemes/double002/expected
Normal file
10
tests/allschemes/double002/expected
Normal file
@ -0,0 +1,10 @@
|
|||||||
|
1/1: Building NaN (NaN.idr)
|
||||||
|
Main> Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> False
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
8
tests/allschemes/double002/input
Normal file
8
tests/allschemes/double002/input
Normal file
@ -0,0 +1,8 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn divZeroZero
|
||||||
|
:exec printLn nanPlus
|
||||||
|
:exec printLn nanSub
|
||||||
|
:exec printLn nanMult
|
||||||
|
:exec printLn nanDiv
|
||||||
|
:exec printLn negNaN
|
||||||
|
:exec printLn nansNotEq
|
3
tests/allschemes/double002/run
Executable file
3
tests/allschemes/double002/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 NaN.idr < input
|
52
tests/allschemes/double003/Inf.idr
Normal file
52
tests/allschemes/double003/Inf.idr
Normal file
@ -0,0 +1,52 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- negating positive inf (should give negative inf)
|
||||||
|
negInf : Double
|
||||||
|
negInf = negate inf
|
||||||
|
|
||||||
|
-- dividing non-zero by zero tends to infinity
|
||||||
|
divNZbyZero : Bool
|
||||||
|
divNZbyZero = 1.0 / 0.0 == inf
|
||||||
|
|
||||||
|
-- adding pos-inf to pos-inf gives pos-inf
|
||||||
|
addPosInfs : Bool
|
||||||
|
addPosInfs = inf + inf == inf
|
||||||
|
|
||||||
|
-- subbing inf from neg-inf gives neg-inf
|
||||||
|
subNegInfs : Bool
|
||||||
|
subNegInfs = -inf - inf == -inf
|
||||||
|
|
||||||
|
-- subtracting pos-inf from pos-inf is nan
|
||||||
|
-- (can't use `==` because no nans are equal)
|
||||||
|
subPosInf : Double
|
||||||
|
subPosInf = inf - inf
|
||||||
|
|
||||||
|
-- adding pos-inf to neg-inf is nan
|
||||||
|
-- (can't use `==` because no nans are equal)
|
||||||
|
addNegInf : Double
|
||||||
|
addNegInf = -inf + inf
|
||||||
|
|
||||||
|
-- neg-inf by neg-inf gives pos-inf
|
||||||
|
squareNegInfIsPosInf : Bool
|
||||||
|
squareNegInfIsPosInf = pow (-inf) 2 == inf
|
||||||
|
|
||||||
|
-- pos-inf by pos-inf gives pos-inf
|
||||||
|
squarePosInfIsPosInf : Bool
|
||||||
|
squarePosInfIsPosInf = pow inf 2 == inf
|
||||||
|
|
||||||
|
-- pos-inf by neg-inf is neg-inf
|
||||||
|
multInfSignFlip : Bool
|
||||||
|
multInfSignFlip = inf * (-inf) == -inf
|
||||||
|
|
||||||
|
-- pos-inf is equal to some other pos-inf
|
||||||
|
posInfIsPosInf : Bool
|
||||||
|
posInfIsPosInf = inf == pow inf 2
|
||||||
|
|
||||||
|
-- neg-inf is equal to some other neg-inf
|
||||||
|
negInfIsNegInf : Bool
|
||||||
|
negInfIsNegInf = -inf == pow (-inf) 3
|
||||||
|
|
||||||
|
-- pos-inf is not neg-inf
|
||||||
|
posInfEqNegInf : Bool
|
||||||
|
posInfEqNegInf = inf == -inf
|
||||||
|
|
15
tests/allschemes/double003/expected
Normal file
15
tests/allschemes/double003/expected
Normal file
@ -0,0 +1,15 @@
|
|||||||
|
1/1: Building Inf (Inf.idr)
|
||||||
|
Main> Main> -inf.0
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> +nan.0
|
||||||
|
Main> +nan.0
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> False
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
13
tests/allschemes/double003/input
Normal file
13
tests/allschemes/double003/input
Normal file
@ -0,0 +1,13 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn negInf
|
||||||
|
:exec printLn divNZbyZero
|
||||||
|
:exec printLn addPosInfs
|
||||||
|
:exec printLn subNegInfs
|
||||||
|
:exec printLn subPosInf
|
||||||
|
:exec printLn addNegInf
|
||||||
|
:exec printLn squareNegInfIsPosInf
|
||||||
|
:exec printLn squarePosInfIsPosInf
|
||||||
|
:exec printLn multInfSignFlip
|
||||||
|
:exec printLn posInfIsPosInf
|
||||||
|
:exec printLn negInfIsNegInf
|
||||||
|
:exec printLn posInfEqNegInf
|
3
tests/allschemes/double003/run
Executable file
3
tests/allschemes/double003/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 Inf.idr < input
|
19
tests/node/double001/URandEpsilon.idr
Normal file
19
tests/node/double001/URandEpsilon.idr
Normal file
@ -0,0 +1,19 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- adding the rounding unit to a Double should not modify it
|
||||||
|
|
||||||
|
testOnePlusUR : Bool
|
||||||
|
testOnePlusUR = 1.0 + unitRoundoff == 1.0
|
||||||
|
|
||||||
|
testURPlusOne : Bool
|
||||||
|
testURPlusOne = unitRoundoff + 1.0 == 1.0
|
||||||
|
|
||||||
|
testURComm : Bool
|
||||||
|
testURComm = testOnePlusUR == testURPlusOne
|
||||||
|
|
||||||
|
|
||||||
|
-- the machine epsilon should be double the rounding unit
|
||||||
|
|
||||||
|
testEps2UR : Bool
|
||||||
|
testEps2UR = epsilon == unitRoundoff * 2.0
|
||||||
|
|
7
tests/node/double001/expected
Normal file
7
tests/node/double001/expected
Normal file
@ -0,0 +1,7 @@
|
|||||||
|
1/1: Building URandEpsilon (URandEpsilon.idr)
|
||||||
|
Main> Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
5
tests/node/double001/input
Normal file
5
tests/node/double001/input
Normal file
@ -0,0 +1,5 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn testOnePlusUR
|
||||||
|
:exec printLn testURPlusOne
|
||||||
|
:exec printLn testURComm
|
||||||
|
:exec printLn testEps2UR
|
3
tests/node/double001/run
Executable file
3
tests/node/double001/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 URandEpsilon.idr < input
|
28
tests/node/double002/NaN.idr
Normal file
28
tests/node/double002/NaN.idr
Normal file
@ -0,0 +1,28 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- undefined things are NaN
|
||||||
|
divZeroZero : Double
|
||||||
|
divZeroZero = 0.0 / 0.0
|
||||||
|
|
||||||
|
-- NaN with anything is NaN
|
||||||
|
|
||||||
|
nanPlus : Double
|
||||||
|
nanPlus = 1.0 + nan
|
||||||
|
|
||||||
|
nanSub : Double
|
||||||
|
nanSub = 1.0 - nan
|
||||||
|
|
||||||
|
nanMult : Double
|
||||||
|
nanMult = 2.0 * nan
|
||||||
|
|
||||||
|
nanDiv : Double
|
||||||
|
nanDiv = nan / 2.0
|
||||||
|
|
||||||
|
-- neg NaN is still NaN
|
||||||
|
negNaN : Double
|
||||||
|
negNaN = negate nan
|
||||||
|
|
||||||
|
-- NaNs are never equal
|
||||||
|
nansNotEq : Bool
|
||||||
|
nansNotEq = nan == nan
|
||||||
|
|
10
tests/node/double002/expected
Normal file
10
tests/node/double002/expected
Normal file
@ -0,0 +1,10 @@
|
|||||||
|
1/1: Building NaN (NaN.idr)
|
||||||
|
Main> Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> False
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
8
tests/node/double002/input
Normal file
8
tests/node/double002/input
Normal file
@ -0,0 +1,8 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn divZeroZero
|
||||||
|
:exec printLn nanPlus
|
||||||
|
:exec printLn nanSub
|
||||||
|
:exec printLn nanMult
|
||||||
|
:exec printLn nanDiv
|
||||||
|
:exec printLn negNaN
|
||||||
|
:exec printLn nansNotEq
|
3
tests/node/double002/run
Executable file
3
tests/node/double002/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 NaN.idr < input
|
52
tests/node/double003/Inf.idr
Normal file
52
tests/node/double003/Inf.idr
Normal file
@ -0,0 +1,52 @@
|
|||||||
|
import Data.Double
|
||||||
|
|
||||||
|
-- negating positive inf (should give negative inf)
|
||||||
|
negInf : Double
|
||||||
|
negInf = negate inf
|
||||||
|
|
||||||
|
-- dividing non-zero by zero tends to infinity
|
||||||
|
divNZbyZero : Bool
|
||||||
|
divNZbyZero = 1.0 / 0.0 == inf
|
||||||
|
|
||||||
|
-- adding pos-inf to pos-inf gives pos-inf
|
||||||
|
addPosInfs : Bool
|
||||||
|
addPosInfs = inf + inf == inf
|
||||||
|
|
||||||
|
-- subbing inf from neg-inf gives neg-inf
|
||||||
|
subNegInfs : Bool
|
||||||
|
subNegInfs = -inf - inf == -inf
|
||||||
|
|
||||||
|
-- subtracting pos-inf from pos-inf is nan
|
||||||
|
-- (can't use `==` because no nans are equal)
|
||||||
|
subPosInf : Double
|
||||||
|
subPosInf = inf - inf
|
||||||
|
|
||||||
|
-- adding pos-inf to neg-inf is nan
|
||||||
|
-- (can't use `==` because no nans are equal)
|
||||||
|
addNegInf : Double
|
||||||
|
addNegInf = -inf + inf
|
||||||
|
|
||||||
|
-- neg-inf by neg-inf gives pos-inf
|
||||||
|
squareNegInfIsPosInf : Bool
|
||||||
|
squareNegInfIsPosInf = pow (-inf) 2 == inf
|
||||||
|
|
||||||
|
-- pos-inf by pos-inf gives pos-inf
|
||||||
|
squarePosInfIsPosInf : Bool
|
||||||
|
squarePosInfIsPosInf = pow inf 2 == inf
|
||||||
|
|
||||||
|
-- pos-inf by neg-inf is neg-inf
|
||||||
|
multInfSignFlip : Bool
|
||||||
|
multInfSignFlip = inf * (-inf) == -inf
|
||||||
|
|
||||||
|
-- pos-inf is equal to some other pos-inf
|
||||||
|
posInfIsPosInf : Bool
|
||||||
|
posInfIsPosInf = inf == pow inf 2
|
||||||
|
|
||||||
|
-- neg-inf is equal to some other neg-inf
|
||||||
|
negInfIsNegInf : Bool
|
||||||
|
negInfIsNegInf = -inf == pow (-inf) 3
|
||||||
|
|
||||||
|
-- pos-inf is not neg-inf
|
||||||
|
posInfEqNegInf : Bool
|
||||||
|
posInfEqNegInf = inf == -inf
|
||||||
|
|
15
tests/node/double003/expected
Normal file
15
tests/node/double003/expected
Normal file
@ -0,0 +1,15 @@
|
|||||||
|
1/1: Building Inf (Inf.idr)
|
||||||
|
Main> Main> -Infinity
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> NaN
|
||||||
|
Main> NaN
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> True
|
||||||
|
Main> False
|
||||||
|
Main>
|
||||||
|
Bye for now!
|
13
tests/node/double003/input
Normal file
13
tests/node/double003/input
Normal file
@ -0,0 +1,13 @@
|
|||||||
|
-- we need to exec to get the prim__* functions to disappear
|
||||||
|
:exec printLn negInf
|
||||||
|
:exec printLn divNZbyZero
|
||||||
|
:exec printLn addPosInfs
|
||||||
|
:exec printLn subNegInfs
|
||||||
|
:exec printLn subPosInf
|
||||||
|
:exec printLn addNegInf
|
||||||
|
:exec printLn squareNegInfIsPosInf
|
||||||
|
:exec printLn squarePosInfIsPosInf
|
||||||
|
:exec printLn multInfSignFlip
|
||||||
|
:exec printLn posInfIsPosInf
|
||||||
|
:exec printLn negInfIsNegInf
|
||||||
|
:exec printLn posInfEqNegInf
|
3
tests/node/double003/run
Executable file
3
tests/node/double003/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
. ../../testutils.sh
|
||||||
|
|
||||||
|
idris2 Inf.idr < input
|
Loading…
Reference in New Issue
Block a user