Idris2/libs/base/Data/IOArray.idr

81 lines
2.4 KiB
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module Data.IOArray
import Data.IOArray.Prims
2020-05-18 15:59:07 +03:00
import Data.List
export
record IOArray elem where
constructor MkIOArray
2020-05-30 19:40:48 +03:00
maxSize : Int
2020-05-18 15:59:07 +03:00
content : ArrayData (Maybe elem)
2020-05-30 19:40:48 +03:00
export
max : IOArray elem -> Int
max = maxSize
2020-05-18 15:59:07 +03:00
export
newArray : HasIO io => Int -> io (IOArray elem)
2020-05-18 15:59:07 +03:00
newArray size
= pure (MkIOArray size !(primIO (prim__newArray size Nothing)))
export
writeArray : HasIO io => IOArray elem -> Int -> elem -> io ()
2020-05-18 15:59:07 +03:00
writeArray arr pos el
= if pos < 0 || pos >= max arr
then pure ()
else primIO (prim__arraySet (content arr) pos (Just el))
export
readArray : HasIO io => IOArray elem -> Int -> io (Maybe elem)
2020-05-18 15:59:07 +03:00
readArray arr pos
= if pos < 0 || pos >= max arr
then pure Nothing
else primIO (prim__arrayGet (content arr) pos)
-- Make a new array of the given size with the elements copied from the
-- other array
export
newArrayCopy : HasIO io =>
(newsize : Int) -> IOArray elem -> io (IOArray elem)
2020-05-18 15:59:07 +03:00
newArrayCopy newsize arr
= do let newsize' = if newsize < max arr then max arr else newsize
arr' <- newArray newsize'
copyFrom (content arr) (content arr') (max arr - 1)
pure arr'
where
copyFrom : ArrayData (Maybe elem) ->
ArrayData (Maybe elem) ->
Int -> io ()
2020-05-18 15:59:07 +03:00
copyFrom old new pos
= if pos < 0
then pure ()
else do el <- primIO $ prim__arrayGet old pos
primIO $ prim__arraySet new pos el
assert_total (copyFrom old new (pos - 1))
export
toList : HasIO io => IOArray elem -> io (List (Maybe elem))
2020-05-18 15:59:07 +03:00
toList arr = iter 0 (max arr) []
where
iter : Int -> Int -> List (Maybe elem) -> io (List (Maybe elem))
2020-05-18 15:59:07 +03:00
iter pos end acc
= if pos >= end
then pure (reverse acc)
else do el <- readArray arr pos
assert_total (iter (pos + 1) end (el :: acc))
export
fromList : List (Maybe elem) -> IO (IOArray elem)
fromList ns
= do arr <- newArray (cast (length ns))
addToArray 0 ns arr
pure arr
where
addToArray : Int -> List (Maybe elem) -> IOArray elem -> IO ()
addToArray loc [] arr = pure ()
addToArray loc (Nothing :: ns) arr
= assert_total (addToArray (loc + 1) ns arr)
addToArray loc (Just el :: ns) arr
= do primIO $ prim__arraySet (content arr) loc (Just el)
assert_total (addToArray (loc + 1) ns arr)