[ fix ] Ensure casting from String to Double results in floating numbers on scheme backends

This commit is contained in:
Jesse Nava 2022-05-07 19:07:28 -04:00 committed by G. Allais
parent 7a5f63eab0
commit db13a35bf5
13 changed files with 242 additions and 5 deletions

View File

@ -241,7 +241,7 @@
(define ct-cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(exact->inexact (cast-num (string->number (destroy-prefix x))))))
(define ct-cast-char-boundedInt
(lambda (x y)

View File

@ -128,7 +128,8 @@
(define cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(exact->inexact (cast-num (string->number (destroy-prefix x))))))
(define (string-concat xs) (apply string-append xs))
(define (string-unpack s) (string->list s))

View File

@ -126,7 +126,8 @@
(if (number? x) x 0))
(define (destroy-prefix x)
(if (or (string=? x "") (char=? (string-ref x 0) #\#)) "" x))
(cast-num (string->number (destroy-prefix x))))
(exact->inexact (cast-num (string->number (destroy-prefix x)))))
(define-macro (cast-string-int x)
`(exact-truncate (cast-string-double ,x)))

View File

@ -241,7 +241,7 @@
(define ct-cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(exact->inexact (cast-num (string->number (destroy-prefix x))))))
(define ct-cast-char-boundedInt
(lambda (x y)

View File

@ -129,7 +129,7 @@
(define cast-string-double
(lambda (x)
(cast-num (string->number (destroy-prefix x)))))
(exact->inexact (cast-num (string->number (destroy-prefix x))))))
(define (string-concat xs) (apply string-append xs))
(define (string-unpack s) (string->list s))
(define (string-pack xs) (list->string xs))

View File

@ -174,6 +174,13 @@ idrisTestsEvaluator = MkTestPool "Evaluation" [] Nothing
"interpreter001", "interpreter002", "interpreter003", "interpreter004",
"interpreter005", "interpreter006", "interpreter007", "interpreter008"]
idrisTestsAllSchemes : Requirement -> TestPool
idrisTestsAllSchemes cg = MkTestPool
("Test across all scheme backends: " ++ show cg ++ " instance")
[] (Just cg)
[ "scheme001"
]
idrisTestsAllBackends : Requirement -> TestPool
idrisTestsAllBackends cg = MkTestPool
("Test across all backends: " ++ show cg ++ " instance")
@ -183,6 +190,7 @@ idrisTestsAllBackends cg = MkTestPool
-- Unfortunately the behaviour of Double is platform dependent so the
-- following test is turned off.
-- "evaluator005",
"issue2362",
"basic048",
"perf006"]
@ -381,6 +389,7 @@ main = runner $
, !templateTests
, !codegenTests
]
++ map (testPaths "allschemes" . idrisTestsAllSchemes) [Chez, Racket]
++ map (testPaths "allbackends" . idrisTestsAllBackends) [Chez, Node, Racket]

View File

@ -0,0 +1,19 @@
module Issue2362
import Data.Vect
%default total
items : Vect ? String
items = [ "0.0", "0", "0.", ".0", "+0", "-0", "string", "" ]
zeroes_double : Vect ? Double
zeroes_double = cast <$> items
zeroes_show : Vect ? String
zeroes_show = show <$> zeroes_double
main : IO ()
main = do
printLn $ all (== head zeroes_double) zeroes_double
printLn $ all (== head zeroes_show) zeroes_show

View File

@ -0,0 +1,2 @@
True
True

View File

@ -0,0 +1,4 @@
rm -rf build
$1 --exec main Issue2362.idr

View File

@ -0,0 +1,9 @@
module CastStringDouble
%default total
main : IO ()
main = traverse_ printLn $ do
f <- [(1/), log, sin, cos, tan, asin, acos, atan, sqrt, floor, ceiling]
v <- [ "", "-1", "-1.0", "-1.", "0", "0.", ".0", ".1", "1", "1.0", "+0", "+1", "+.1", "+1.0", "string", "+nan.0", "+inf.0" ]
pure (floor $ f (cast v))

View File

@ -0,0 +1,187 @@
+inf.0
-1.0
-1.0
-1.0
+inf.0
+inf.0
+inf.0
10.0
1.0
1.0
+inf.0
1.0
10.0
1.0
+inf.0
+nan.0
0.0
-inf.0
+nan.0
+nan.0
+nan.0
-inf.0
-inf.0
-inf.0
-3.0
0.0
0.0
-inf.0
0.0
-3.0
0.0
-inf.0
+nan.0
+inf.0
0.0
-1.0
-1.0
-1.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
+nan.0
+nan.0
1.0
0.0
0.0
0.0
1.0
1.0
1.0
0.0
0.0
0.0
1.0
0.0
0.0
0.0
1.0
+nan.0
+nan.0
0.0
-2.0
-2.0
-2.0
0.0
0.0
0.0
0.0
1.0
1.0
0.0
1.0
0.0
1.0
0.0
+nan.0
+nan.0
0.0
-2.0
-2.0
-2.0
0.0
0.0
0.0
0.0
1.0
1.0
0.0
1.0
0.0
1.0
0.0
+nan.0
+nan.0
1.0
3.0
3.0
3.0
1.0
1.0
1.0
1.0
0.0
0.0
1.0
0.0
1.0
0.0
1.0
+nan.0
+nan.0
0.0
-1.0
-1.0
-1.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
0.0
+nan.0
1.0
0.0
+nan.0
+nan.0
+nan.0
0.0
0.0
0.0
0.0
1.0
1.0
0.0
1.0
0.0
1.0
0.0
+nan.0
+inf.0
0.0
-1.0
-1.0
-1.0
0.0
0.0
0.0
0.0
1.0
1.0
0.0
1.0
0.0
1.0
0.0
+nan.0
+inf.0
0.0
-1.0
-1.0
-1.0
0.0
0.0
0.0
1.0
1.0
1.0
0.0
1.0
1.0
1.0
0.0
+nan.0
+inf.0

View File

@ -0,0 +1,2 @@
:exec main
:q

View File

@ -0,0 +1,3 @@
rm -rf build
$1 --exec main CastStringDouble.idr