Idris2/libs/base/Data/List1.idr
Edwin Brady ad632d825d Remove linearity subtyping
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).
2020-12-27 19:58:35 +00:00

144 lines
3.1 KiB
Idris

module Data.List1
%default total
infixr 7 :::
public export
record List1 a where
constructor (:::)
head : a
tail : List a
------------------------------------------------------------------------
-- Conversion
||| Forget that a list is non-empty
public export
forget : (xs : List1 a) -> List a
forget (x ::: xs) = x :: xs
||| Check whether a list is non-empty
export
fromList : (xs : List a) -> Maybe (List1 a)
fromList [] = Nothing
fromList (x :: xs) = Just (x ::: xs)
------------------------------------------------------------------------
-- Basic functions
public export
singleton : (x : a) -> List1 a
singleton a = a ::: []
export
last : List1 a -> a
last (x ::: xs) = loop x xs where
loop : a -> List a -> a
loop x [] = x
loop _ (x :: xs) = loop x xs
export
foldr1' : (a -> b -> b) -> (a -> b) -> List1 a -> b
foldr1' c n (x ::: xs) = loop x xs where
loop : a -> List a -> b
loop a [] = n a
loop a (x :: xs) = c a (loop x xs)
export
foldr1 : (a -> a -> a) -> List1 a -> a
foldr1 c = foldr1' c id
export
foldl1 : (a -> b) -> (b -> a -> b) -> List1 a -> b
foldl1 n c (x ::: xs) = foldl c (n x) xs
------------------------------------------------------------------------
-- Append
export
appendl : (xs : List1 a) -> (ys : List a) -> List1 a
appendl (x ::: xs) ys = x ::: xs ++ ys
export
append : (xs, ys : List1 a) -> List1 a
append xs ys = appendl xs (forget ys)
export
lappend : (xs : List a) -> (ys : List1 a) -> List1 a
lappend [] ys = ys
lappend (x :: xs) ys = append (x ::: xs) ys
------------------------------------------------------------------------
-- Cons/Snoc
public export
cons : (x : a) -> (xs : List1 a) -> List1 a
cons x xs = x ::: forget xs
export
snoc : (xs : List1 a) -> (x : a) -> List1 a
snoc xs x = append xs (singleton x)
------------------------------------------------------------------------
-- Reverse
public export
reverseOnto : (acc : List1 a) -> (xs : List a) -> List1 a
reverseOnto acc [] = acc
reverseOnto acc (x :: xs) = reverseOnto (x ::: forget acc) xs
public export
reverse : (xs : List1 a) -> List1 a
reverse (x ::: xs) = reverseOnto (singleton x) xs
------------------------------------------------------------------------
-- Instances
export
Semigroup (List1 a) where
(<+>) = append
export
Functor List1 where
map f (x ::: xs) = f x ::: map f xs
export
Applicative List1 where
pure x = singleton x
f ::: fs <*> xs = appendl (map f xs) (fs <*> forget xs)
export
Monad List1 where
(x ::: xs) >>= f = appendl (f x) (xs >>= forget . f)
export
Foldable List1 where
foldr c n (x ::: xs) = c x (foldr c n xs)
null _ = False
export
Traversable List1 where
traverse f (x ::: xs) = [| f x ::: traverse f xs |]
export
Show a => Show (List1 a) where
show = show . forget
export
Eq a => Eq (List1 a) where
(x ::: xs) == (y ::: ys) = x == y && xs == ys
export
Ord a => Ord (List1 a) where
compare xs ys = compare (forget xs) (forget ys)
------------------------------------------------------------------------
-- Properties
export
consInjective : the (List1 a) (x ::: xs) === (y ::: ys) -> (x === y, xs === ys)
consInjective Refl = (Refl, Refl)