This commit is contained in:
stefan-hoeck 2021-02-06 07:48:39 +01:00 committed by G. Allais
parent 8102e9e495
commit 1104776430
6 changed files with 18 additions and 1 deletions

View File

@ -141,7 +141,7 @@ boundedUInt : {auto c : Ref ESs ESSt} -> Int -> String -> Core String
boundedUInt bits e =
do
n <- makeIntBound bits
fn <- addConstToPreamble ("truncToUInt"++show bits) ("x=>{const m = x%" ++ n ++ ";return m>0?m:m+" ++ n ++ "}")
fn <- addConstToPreamble ("truncToUInt"++show bits) ("x=>{const m = x%" ++ n ++ ";return m>=0?m:m+" ++ n ++ "}")
pure $ fn ++ "(" ++ e ++ ")"
boundedIntOp : {auto c : Ref ESs ESSt} -> Int -> String -> String -> String -> Core String

View File

@ -194,6 +194,7 @@ nodeTests = MkTestPool [Node]
[ "node001", "node002", "node003", "node004", "node005", "node006"
, "node007", "node008", "node009", "node011", "node012", "node015"
, "node017", "node018", "node019", "node021", "node022", "node023"
, "node025"
-- , "node14", "node020"
, "reg001"
, "syntax001"

View File

@ -0,0 +1,5 @@
main : HasIO io => io ()
main = do printLn (0 * the Bits8 1)
printLn (0 * the Bits16 1)
printLn (0 * the Bits32 1)
printLn (0 * the Bits64 1)

View File

@ -0,0 +1,6 @@
0
0
0
0
1/1: Building Fix1037 (Fix1037.idr)
Main> Main> Bye for now!

2
tests/node/node025/input Normal file
View File

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

3
tests/node/node025/run Normal file
View File

@ -0,0 +1,3 @@
$1 --cg node --no-banner --no-color --console-width 0 Fix1037.idr < input
rm -rf build