mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-17 08:11:45 +03:00
ad632d825d
It's disappointing to have to do this, but I think necessary because various issue reports have shown it to be unsound (at least as far as inference goes) and, at the very least, confusing. This patch brings us back to the basic rules of QTT. On the one hand, this makes the 1 multiplicity less useful, because it means we can't flag arguments as being used exactly once which would be useful for optimisation purposes as well as precision in the type. On the other hand, it removes some complexity (and a hack) from unification, and has the advantage of being correct! Also, I still consider the 1 multiplicity an experiment. We can still do interesting things like protocol state tracking, which is my primary motivation at least. Ideally, if the 1 multiplicity is going to be more generall useful, we'll need some kind of way of doing multiplicity polymorphism in the future. I don't think subtyping is the way (I've pretty much always come to regret adding some form of subtyping). Fixes #73 (and maybe some others).
76 lines
2.3 KiB
Idris
76 lines
2.3 KiB
Idris
module Data.Linear.Array
|
|
|
|
import Data.IOArray
|
|
|
|
-- Linear arrays. General idea: mutable arrays are constructed linearly,
|
|
-- using newArray. Once everything is set up, they can be converted to
|
|
-- read only arrays with constant time, pure, access, using toIArray.
|
|
|
|
-- Immutable arrays which can be read in constant time, but not updated
|
|
public export
|
|
interface Array arr where
|
|
read : (1 _ : arr t) -> Int -> Maybe t
|
|
size : (1 _ : arr t) -> Int
|
|
|
|
-- Mutable arrays which can be used linearly
|
|
public export
|
|
interface Array arr => MArray arr where
|
|
newArray : (size : Int) -> (1 _ : (1 _ : arr t) -> a) -> a
|
|
-- Array is unchanged if the index is out of bounds
|
|
write : (1 _ : arr t) -> Int -> t -> arr t
|
|
|
|
mread : (1 _ : arr t) -> Int -> Res (Maybe t) (const (arr t))
|
|
msize : (1 _ : arr t) -> Res Int (const (arr t))
|
|
|
|
export
|
|
data IArray : Type -> Type where
|
|
MkIArray : IOArray t -> IArray t
|
|
|
|
export
|
|
data LinArray : Type -> Type where
|
|
MkLinArray : IOArray t -> LinArray t
|
|
|
|
-- Convert a mutable array to an immutable array
|
|
export
|
|
toIArray : (1 _ : LinArray t) -> (IArray t -> a) -> a
|
|
toIArray (MkLinArray arr) k = k (MkIArray arr)
|
|
|
|
export
|
|
Array LinArray where
|
|
read (MkLinArray a) i = unsafePerformIO (readArray a i)
|
|
size (MkLinArray a) = max a
|
|
|
|
export
|
|
MArray LinArray where
|
|
newArray size k = k (MkLinArray (unsafePerformIO (newArray size)))
|
|
|
|
write (MkLinArray a) i el
|
|
= unsafePerformIO (do writeArray a i el
|
|
pure (MkLinArray a))
|
|
msize (MkLinArray a) = max a # MkLinArray a
|
|
mread (MkLinArray a) i
|
|
= unsafePerformIO (readArray a i) # MkLinArray a
|
|
|
|
export
|
|
Array IArray where
|
|
read (MkIArray a) i = unsafePerformIO (readArray a i)
|
|
size (MkIArray a) = max a
|
|
|
|
export
|
|
copyArray : MArray arr => (newsize : Int) -> (1 _ : arr t) ->
|
|
LPair (arr t) (arr t)
|
|
copyArray newsize a
|
|
= let size # a = msize a in
|
|
newArray newsize $
|
|
(\a' => copyContent (min (newsize - 1) (size - 1)) a a')
|
|
where
|
|
copyContent : Int -> (1 _ : arr t) -> (1 _ : arr t) -> LPair (arr t) (arr t)
|
|
copyContent pos a a'
|
|
= if pos < 0
|
|
then a # a'
|
|
else let val # a = mread a pos
|
|
1 a' = case val of
|
|
Nothing => a'
|
|
Just v => write a' pos v in
|
|
copyContent (pos - 1) a a'
|