mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 02:23:44 +03:00
fix 1037
This commit is contained in:
parent
8102e9e495
commit
1104776430
@ -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
|
||||
|
@ -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"
|
||||
|
5
tests/node/node025/Fix1037.idr
Normal file
5
tests/node/node025/Fix1037.idr
Normal 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)
|
6
tests/node/node025/expected
Normal file
6
tests/node/node025/expected
Normal 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
2
tests/node/node025/input
Normal file
@ -0,0 +1,2 @@
|
||||
:exec main
|
||||
:q
|
3
tests/node/node025/run
Normal file
3
tests/node/node025/run
Normal file
@ -0,0 +1,3 @@
|
||||
$1 --cg node --no-banner --no-color --console-width 0 Fix1037.idr < input
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user