mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-21 10:41:59 +03:00
70ef197cf6
* [ base ] Deprecate setByte in favour of setBits8 Deprecate getByte; fix Core.Binary.Prims Along with `setByte`, the `getByte` function should similarly be deprecated since it also assumes the value will have the given size, rather than guaranteeing it in the type. CI highlighted some required changes in `Core.Binary.Prims` thanks to `-Werror`. The fix was to add some `cast` calls where the old `getByte` and `setByte` used to be. The RefC buffer test will need updating once we remove the functions completely. Added a note for future peeps.
57 lines
1.4 KiB
Idris
57 lines
1.4 KiB
Idris
module TestBuffer
|
|
|
|
import Data.Buffer
|
|
import System.File
|
|
|
|
put : Show a => IO a -> IO ()
|
|
put = (>>= putStrLn . show)
|
|
|
|
main : IO ()
|
|
main = do
|
|
Just buf <- newBuffer 31
|
|
| Nothing => pure ()
|
|
|
|
-- TODO: setByte is deprecated, remove once no longer in libs
|
|
setByte buf 0 1
|
|
setBits8 buf 1 2
|
|
setInt buf 2 0x1122334455667788
|
|
setDouble buf 10 (sqrt 2)
|
|
|
|
let helloWorld = "Hello, world"
|
|
|
|
Just helloWorldBuf <- newBuffer (stringByteLength helloWorld)
|
|
| Nothing => pure ()
|
|
|
|
setString helloWorldBuf 0 "Hello, world"
|
|
copyData helloWorldBuf 0 12 buf 18
|
|
|
|
put $ rawSize buf
|
|
|
|
-- TODO: getByte is deprecated, remove once no longer in libs
|
|
put $ getByte buf 0
|
|
put $ getBits8 buf 1
|
|
put $ getInt buf 2
|
|
put $ getDouble buf 10
|
|
put $ getString buf 18 12
|
|
|
|
put $ bufferData buf
|
|
|
|
Just readBuf <- newBuffer 8
|
|
| Nothing => pure ()
|
|
Right f <- openFile "testRead.buf" Read
|
|
| Left err => put $ pure err
|
|
Right 8 <- readBufferData f readBuf 0 8
|
|
| Right size => put $ pure "\{show size} bytes have been read, 8 expected"
|
|
| Left err => put $ pure err
|
|
put $ bufferData readBuf
|
|
|
|
Just writeBuf <- newBuffer 8
|
|
| Nothing => pure ()
|
|
setInt writeBuf 0 0x7766554433221100
|
|
Right f <- openFile "testWrite.buf" WriteTruncate
|
|
| Left err => put $ pure err
|
|
Right () <- writeBufferData f writeBuf 0 8
|
|
| Left err => put $ pure err
|
|
|
|
pure ()
|