2020-06-13 22:20:53 +03:00
|
|
|
module Debug.Buffer
|
|
|
|
|
|
|
|
import Data.Buffer
|
|
|
|
import Data.List
|
|
|
|
|
2020-06-13 22:53:06 +03:00
|
|
|
toHex : Int -> Int -> String
|
|
|
|
toHex d n = pack $ reverse $ foldl toHexDigit [] (slice d n [])
|
2020-06-13 22:20:53 +03:00
|
|
|
where
|
|
|
|
toHexDigit : List Char ->Int -> List Char
|
2020-06-13 22:53:06 +03:00
|
|
|
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 0 _ acc = acc
|
|
|
|
slice d n acc = slice (d-1) (n `div` 16) ((n `mod` 16)::acc)
|
2020-06-13 22:20:53 +03:00
|
|
|
|
|
|
|
showSep : String -> List String -> String
|
|
|
|
showSep sep [] = ""
|
|
|
|
showSep sep [x] = x
|
|
|
|
showSep sep (x :: xs) = x ++ sep ++ showSep sep xs
|
|
|
|
|
|
|
|
renderRow : List Int -> String
|
2020-06-13 22:53:06 +03:00
|
|
|
renderRow dat = showSep " " (map (toHex 2) dat) ++ " " ++
|
|
|
|
pack (map (\i => if i > 0x1f && i < 0x80 then chr i else '.') dat)
|
2020-06-13 22:20:53 +03:00
|
|
|
|
|
|
|
group : Nat -> List a -> List (List a)
|
|
|
|
group n xs = worker [] xs
|
|
|
|
where
|
|
|
|
worker : List (List a) -> List a -> List (List a)
|
|
|
|
worker acc [] = reverse acc
|
|
|
|
worker acc ys = worker ((take n ys)::acc) (drop n ys)
|
|
|
|
|
|
|
|
export
|
|
|
|
dumpBuffer : Buffer -> IO String
|
|
|
|
dumpBuffer buf = do
|
|
|
|
size <- rawSize buf
|
|
|
|
dat <- bufferData buf
|
|
|
|
let rows = group 16 dat
|
|
|
|
let hex = showSep "\n" (map renderRow rows)
|
|
|
|
pure $ hex ++ "\n\ntotal size = " ++ show size
|
|
|
|
|
|
|
|
export
|
|
|
|
printBuffer : Buffer -> IO ()
|
|
|
|
printBuffer buf = putStrLn $ !(dumpBuffer buf)
|