[ base ] Deprecate setByte in favour of setBits8 (#2764)

* [ 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.
This commit is contained in:
CodingCellist 2022-11-15 10:42:07 +01:00 committed by GitHub
parent f5f7e2eb80
commit 70ef197cf6
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
5 changed files with 21 additions and 6 deletions

View File

@ -13,6 +13,13 @@
### Library changes
#### Base
* Deprecates `setByte` and `getByte` from `Data.Buffer` for removal in a future
release. Use `setBits8` and `getBits8` instead (with `cast` if you need to
convert a `Bits8` to an `Int`), as their values are limited, as opposed to the
assumption in `setByte` that the value is between 0 and 255.
#### System
* Changes `getNProcessors` to return the number of online processors rather than

View File

@ -39,6 +39,7 @@ newBuffer size
-- then pure Nothing
-- else pure $ Just $ MkBuffer buf size 0
-- TODO: remove me when we remove the deprecated `setByte` in a future release
%foreign "scheme:blodwen-buffer-setbyte"
"RefC:setBufferByte"
"node:lambda:(buf,offset,value)=>buf.writeUInt8(value, offset)"
@ -50,6 +51,8 @@ prim__setByte : Buffer -> (offset : Int) -> (val : Int) -> PrimIO ()
prim__setBits8 : Buffer -> (offset : Int) -> (val: Bits8) -> PrimIO ()
-- Assumes val is in the range 0-255
||| Use `setBits8` instead, as its value is correctly limited.
%deprecate
export %inline
setByte : HasIO io => Buffer -> (offset : Int) -> (val : Int) -> io ()
setByte buf offset val
@ -60,6 +63,7 @@ setBits8 : HasIO io => Buffer -> (offset : Int) -> (val : Bits8) -> io ()
setBits8 buf offset val
= primIO (prim__setBits8 buf offset val)
-- TODO: remove me when we remove the deprecated `getByte` in a future release
%foreign "scheme:blodwen-buffer-getbyte"
"RefC:getBufferByte"
"node:lambda:(buf,offset)=>buf.readUInt8(offset)"
@ -70,6 +74,8 @@ prim__getByte : Buffer -> (offset : Int) -> PrimIO Int
"node:lambda:(buf,offset)=>buf.readUInt8(offset)"
prim__getBits8 : Buffer -> (offset : Int) -> PrimIO Bits8
||| Use `getBits8` instead, as its value is correctly limited.
%deprecate
export %inline
getByte : HasIO io => Buffer -> (offset : Int) -> io Int
getByte buf offset

View File

@ -71,10 +71,10 @@ tag {b} val
= do chunk <- get Bin
if avail chunk >= 1
then
do coreLift $ setByte (buf chunk) (loc chunk) val
do coreLift $ setBits8 (buf chunk) (loc chunk) $ cast val
put Bin (appended 1 chunk)
else do chunk' <- extendBinary 1 chunk
coreLift $ setByte (buf chunk') (loc chunk') val
coreLift $ setBits8 (buf chunk') (loc chunk') $ cast val
put Bin (appended 1 chunk')
export
@ -83,10 +83,10 @@ getTag {b}
= do chunk <- get Bin
if toRead chunk >= 1
then
do val <- coreLift $ getByte (buf chunk) (loc chunk)
do val <- coreLift $ getBits8 (buf chunk) (loc chunk)
put Bin (incLoc 1 chunk)
pure val
else throw (TTCError (EndOfBuffer "Byte"))
pure $ cast val
else throw (TTCError (EndOfBuffer "Bits8"))
-- Primitives; these are the only things that have to deal with growing
-- the buffer list

View File

@ -46,7 +46,7 @@ main
ds <- bufferData buf2
printLn ds
setByte buf2 0 1
setBits8 buf2 0 1
Just ccBuf <- concatBuffers [buf, buf2]
| Nothing => putStrLn "Buffer concat failed"
printLn !(bufferData ccBuf)

View File

@ -11,6 +11,7 @@ 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
@ -26,6 +27,7 @@ main = do
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