mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-23 22:22:07 +03:00
Remove use of deprecated getByte
function (#3190)
* Add CHANGELOG_NEXT entry
This commit is contained in:
parent
6c70048380
commit
073fbef3d1
@ -46,6 +46,10 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
|
|||||||
|
|
||||||
* Added append `(++)` for `List` version of `All`.
|
* Added append `(++)` for `List` version of `All`.
|
||||||
|
|
||||||
|
* Deprecate `bufferData` in favor of `bufferData'`. These functions are the same
|
||||||
|
with the exception of the latter dealing in `Bits8` which is more correct than
|
||||||
|
`Int`.
|
||||||
|
|
||||||
#### Contrib
|
#### Contrib
|
||||||
|
|
||||||
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
* `Data.List.Lazy` was moved from `contrib` to `base`.
|
||||||
|
@ -378,6 +378,8 @@ getString : HasIO io => Buffer -> (offset : Int) -> (len : Int) -> io String
|
|||||||
getString buf offset len
|
getString buf offset len
|
||||||
= primIO (prim__getString buf offset len)
|
= primIO (prim__getString buf offset len)
|
||||||
|
|
||||||
|
||| Use `bufferData'` instead, as its value is correctly limited.
|
||||||
|
%deprecate
|
||||||
export
|
export
|
||||||
covering
|
covering
|
||||||
bufferData : HasIO io => Buffer -> io (List Int)
|
bufferData : HasIO io => Buffer -> io (List Int)
|
||||||
@ -389,7 +391,23 @@ bufferData buf
|
|||||||
unpackTo : List Int -> Int -> io (List Int)
|
unpackTo : List Int -> Int -> io (List Int)
|
||||||
unpackTo acc 0 = pure acc
|
unpackTo acc 0 = pure acc
|
||||||
unpackTo acc offset
|
unpackTo acc offset
|
||||||
= do val <- getByte buf (offset - 1)
|
= do val <- getBits8 buf (offset - 1)
|
||||||
|
unpackTo (cast val :: acc) (offset - 1)
|
||||||
|
|
||||||
|
||| This function performs the same task as `bufferData` but it
|
||||||
|
||| returns `List Bits8` which is more correct than `List Int`.
|
||||||
|
export
|
||||||
|
covering
|
||||||
|
bufferData' : HasIO io => Buffer -> io (List Bits8)
|
||||||
|
bufferData' buf
|
||||||
|
= do len <- rawSize buf
|
||||||
|
unpackTo [] len
|
||||||
|
where
|
||||||
|
covering
|
||||||
|
unpackTo : List Bits8 -> Int -> io (List Bits8)
|
||||||
|
unpackTo acc 0 = pure acc
|
||||||
|
unpackTo acc offset
|
||||||
|
= do val <- getBits8 buf (offset - 1)
|
||||||
unpackTo (val :: acc) (offset - 1)
|
unpackTo (val :: acc) (offset - 1)
|
||||||
|
|
||||||
|
|
||||||
|
@ -9,6 +9,7 @@ toHex d n = pack $ reverse $ foldl toHexDigit [] (slice d n [])
|
|||||||
where
|
where
|
||||||
toHexDigit : List Char -> Int -> List Char
|
toHexDigit : List Char -> Int -> List Char
|
||||||
toHexDigit acc i = chr (if i < 10 then i + ord '0' else (i-10) + ord 'A')::acc
|
toHexDigit acc i = chr (if i < 10 then i + ord '0' else (i-10) + ord 'A')::acc
|
||||||
|
|
||||||
slice : Int -> Int -> List Int -> List Int
|
slice : Int -> Int -> List Int -> List Int
|
||||||
slice 0 _ acc = acc
|
slice 0 _ acc = acc
|
||||||
slice d n acc = slice (d-1) (n `div` 16) ((n `mod` 16)::acc)
|
slice d n acc = slice (d-1) (n `div` 16) ((n `mod` 16)::acc)
|
||||||
@ -19,9 +20,9 @@ showSep sep n [x] = x ++ replicate (3*n) ' '
|
|||||||
showSep sep Z (x :: xs) = x ++ sep ++ showSep sep Z xs
|
showSep sep Z (x :: xs) = x ++ sep ++ showSep sep Z xs
|
||||||
showSep sep (S n) (x :: xs) = x ++ sep ++ showSep sep n xs
|
showSep sep (S n) (x :: xs) = x ++ sep ++ showSep sep n xs
|
||||||
|
|
||||||
renderRow : List Int -> String
|
renderRow : List Bits8 -> String
|
||||||
renderRow dat = showSep " " 16 (map (toHex 2) dat) ++
|
renderRow dat = showSep " " 16 (map (toHex 2 . cast) dat) ++
|
||||||
" " ++ pack (map (\i => if i > 0x1f && i < 0x80 then chr i else '.') dat)
|
" " ++ pack (map (\i => if i > 0x1f && i < 0x80 then chr (cast i) else '.') dat)
|
||||||
|
|
||||||
group : Nat -> List a -> List (List a)
|
group : Nat -> List a -> List (List a)
|
||||||
group n xs = worker [] xs
|
group n xs = worker [] xs
|
||||||
@ -34,7 +35,7 @@ export
|
|||||||
dumpBuffer : HasIO io => Buffer -> io String
|
dumpBuffer : HasIO io => Buffer -> io String
|
||||||
dumpBuffer buf = do
|
dumpBuffer buf = do
|
||||||
size <- liftIO $ rawSize buf
|
size <- liftIO $ rawSize buf
|
||||||
dat <- liftIO $ bufferData buf
|
dat <- liftIO $ bufferData' buf
|
||||||
let rows = group 16 dat
|
let rows = group 16 dat
|
||||||
let hex = showSep "\n" 0 (map renderRow rows)
|
let hex = showSep "\n" 0 (map renderRow rows)
|
||||||
pure $ hex ++ "\n\ntotal size = " ++ show size
|
pure $ hex ++ "\n\ntotal size = " ++ show size
|
||||||
|
@ -43,12 +43,27 @@ export
|
|||||||
incLoc : Integer -> Binary -> Binary
|
incLoc : Integer -> Binary -> Binary
|
||||||
incLoc i c = { loc $= (+i) } c
|
incLoc i c = { loc $= (+i) } c
|
||||||
|
|
||||||
|
-- TODO: remove this function once Idris2 v0.8.0 has been released
|
||||||
|
-- and use the version from base instead.
|
||||||
|
covering
|
||||||
|
bufferData' : HasIO io => Buffer -> io (List Bits8)
|
||||||
|
bufferData' buf
|
||||||
|
= do len <- rawSize buf
|
||||||
|
unpackTo [] len
|
||||||
|
where
|
||||||
|
covering
|
||||||
|
unpackTo : List Bits8 -> Int -> io (List Bits8)
|
||||||
|
unpackTo acc 0 = pure acc
|
||||||
|
unpackTo acc offset
|
||||||
|
= do val <- getBits8 buf (offset - 1)
|
||||||
|
unpackTo (val :: acc) (offset - 1)
|
||||||
|
|
||||||
export
|
export
|
||||||
dumpBin : Binary -> IO ()
|
dumpBin : Binary -> IO ()
|
||||||
dumpBin chunk
|
dumpBin chunk
|
||||||
= do -- printLn !(traverse bufferData (map buf done))
|
= do -- printLn !(traverse Binary.bufferData' (map buf done))
|
||||||
printLn !(bufferData (buf chunk))
|
printLn !(Binary.bufferData' (buf chunk))
|
||||||
-- printLn !(traverse bufferData (map buf rest))
|
-- printLn !(traverse Binary.bufferData' (map buf rest))
|
||||||
|
|
||||||
export
|
export
|
||||||
nonEmptyRev : {xs : _} ->
|
nonEmptyRev : {xs : _} ->
|
||||||
|
@ -35,7 +35,7 @@ main
|
|||||||
val <- getBits16 buf 32
|
val <- getBits16 buf 32
|
||||||
printLn val
|
printLn val
|
||||||
|
|
||||||
ds <- bufferData buf
|
ds <- bufferData' buf
|
||||||
printLn ds
|
printLn ds
|
||||||
|
|
||||||
Right _ <- writeBufferToFile "test.buf" buf 100
|
Right _ <- writeBufferToFile "test.buf" buf 100
|
||||||
@ -43,13 +43,13 @@ main
|
|||||||
Right buf2 <- createBufferFromFile "test.buf"
|
Right buf2 <- createBufferFromFile "test.buf"
|
||||||
| Left err => putStrLn "Buffer read fail"
|
| Left err => putStrLn "Buffer read fail"
|
||||||
|
|
||||||
ds <- bufferData buf2
|
ds <- bufferData' buf2
|
||||||
printLn ds
|
printLn ds
|
||||||
|
|
||||||
setBits8 buf2 0 1
|
setBits8 buf2 0 1
|
||||||
Just ccBuf <- concatBuffers [buf, buf2]
|
Just ccBuf <- concatBuffers [buf, buf2]
|
||||||
| Nothing => putStrLn "Buffer concat failed"
|
| Nothing => putStrLn "Buffer concat failed"
|
||||||
printLn !(bufferData ccBuf)
|
printLn !(bufferData' ccBuf)
|
||||||
|
|
||||||
Just (a, b) <- splitBuffer buf 20
|
Just (a, b) <- splitBuffer buf 20
|
||||||
| Nothing => putStrLn "Buffer split failed"
|
| Nothing => putStrLn "Buffer split failed"
|
||||||
|
@ -28,7 +28,7 @@ main
|
|||||||
val <- getBits16 buf 32
|
val <- getBits16 buf 32
|
||||||
printLn val
|
printLn val
|
||||||
|
|
||||||
ds <- bufferData buf
|
ds <- bufferData' buf
|
||||||
printLn ds
|
printLn ds
|
||||||
|
|
||||||
Right _ <- writeBufferToFile "test.buf" buf 100
|
Right _ <- writeBufferToFile "test.buf" buf 100
|
||||||
@ -36,7 +36,7 @@ main
|
|||||||
Right buf2 <- createBufferFromFile "test.buf"
|
Right buf2 <- createBufferFromFile "test.buf"
|
||||||
| Left err => putStrLn "Buffer read fail"
|
| Left err => putStrLn "Buffer read fail"
|
||||||
|
|
||||||
ds <- bufferData buf2
|
ds <- bufferData' buf2
|
||||||
printLn ds
|
printLn ds
|
||||||
|
|
||||||
-- Put back when the File API is moved to C and these can work again
|
-- Put back when the File API is moved to C and these can work again
|
||||||
|
@ -38,7 +38,7 @@ main = do
|
|||||||
put $ getDouble buf 16
|
put $ getDouble buf 16
|
||||||
put $ getString buf 24 12
|
put $ getString buf 24 12
|
||||||
|
|
||||||
put $ bufferData buf
|
put $ bufferData' buf
|
||||||
|
|
||||||
Just readBuf <- newBuffer 8
|
Just readBuf <- newBuffer 8
|
||||||
| Nothing => pure ()
|
| Nothing => pure ()
|
||||||
@ -47,7 +47,7 @@ main = do
|
|||||||
Right 8 <- readBufferData f readBuf 0 8
|
Right 8 <- readBufferData f readBuf 0 8
|
||||||
| Right size => put $ pure "\{show size} bytes have been read, 8 expected"
|
| Right size => put $ pure "\{show size} bytes have been read, 8 expected"
|
||||||
| Left err => put $ pure err
|
| Left err => put $ pure err
|
||||||
put $ bufferData readBuf
|
put $ bufferData' readBuf
|
||||||
|
|
||||||
Just writeBuf <- newBuffer 8
|
Just writeBuf <- newBuffer 8
|
||||||
| Nothing => pure ()
|
| Nothing => pure ()
|
||||||
|
Loading…
Reference in New Issue
Block a user