Idris2/libs/base/Data/List1.idr

208 lines
5.1 KiB
Idris
Raw Normal View History

module Data.List1
2021-01-27 21:23:08 +03:00
import public Data.Zippable
%default total
infixr 7 :::
2021-03-17 17:07:52 +03:00
||| Non-empty lists.
public export
record List1 a where
constructor (:::)
head : a
tail : List a
2021-01-26 13:39:16 +03:00
%name List1 xs, ys, zs
------------------------------------------------------------------------
-- Basic functions
public export
fromList : List a -> Maybe (List1 a)
fromList [] = Nothing
fromList (x :: xs) = Just (x ::: xs)
public export
singleton : (x : a) -> List1 a
singleton a = a ::: []
2021-03-17 17:07:52 +03:00
||| Forget that a list is non-empty.
public export
forget : List1 a -> List a
forget (x ::: xs) = x :: xs
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
2021-03-17 17:07:52 +03:00
init : List1 a -> List a
init (x ::: xs) = loop x xs where
loop : a -> List a -> List a
loop x [] = []
loop x (y :: xs) = x :: loop y xs
2021-03-17 17:07:52 +03:00
export
foldr1By : (func : a -> b -> b) -> (map : a -> b) -> (l : List1 a) -> b
foldr1By f map (x ::: xs) = loop x xs where
loop : a -> List a -> b
2021-03-17 17:07:52 +03:00
loop x [] = map x
loop x (y :: xs) = f x (loop y xs)
export
foldl1By : (func : b -> a -> b) -> (map : a -> b) -> (l : List1 a) -> b
foldl1By f map (x ::: xs) = foldl f (map x) xs
export
2021-03-17 17:07:52 +03:00
foldr1 : (func : a -> a -> a) -> (l : List1 a) -> a
foldr1 f = foldr1By f id
export
2021-03-17 17:07:52 +03:00
foldl1 : (func : a -> a -> a) -> (l : List1 a) -> a
foldl1 f = foldl1By f id
------------------------------------------------------------------------
-- Append
export
appendl : (xs : List1 a) -> (ys : List a) -> List1 a
appendl (x ::: xs) ys = x ::: xs ++ ys
export
2021-03-17 17:07:52 +03:00
(++) : (xs, ys : List1 a) -> List1 a
(++) xs ys = appendl xs (forget ys)
export
lappend : (xs : List a) -> (ys : List1 a) -> List1 a
lappend [] ys = ys
2021-03-17 17:07:52 +03:00
lappend (x :: xs) ys = (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
2021-03-17 17:07:52 +03:00
snoc xs x = xs ++ (singleton x)
public export
unsnoc : (xs : List1 a) -> (List a, a)
unsnoc (h ::: Nil) = (Nil, h)
unsnoc (h ::: (x :: xs)) =
let (ini,lst) = unsnoc (x ::: xs)
in (h :: ini, lst)
------------------------------------------------------------------------
-- 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
2021-03-17 17:07:52 +03:00
(<+>) = (++)
public export
Functor List1 where
map f (x ::: xs) = f x ::: map f xs
public 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)
2021-06-01 17:05:04 +03:00
foldl f z (x ::: xs) = foldl f (f z x) xs
null _ = False
toList = forget
2021-06-01 17:05:04 +03:00
foldMap f (x ::: xs) = f x <+> foldMap f xs
2020-12-21 06:46:14 +03:00
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
2021-06-28 14:49:33 +03:00
consInjective : (x ::: xs) === (y ::: ys) -> (x === y, xs === ys)
consInjective Refl = (Refl, Refl)
2021-01-26 13:39:16 +03:00
------------------------------------------------------------------------
2021-01-27 21:23:08 +03:00
-- Zippable
2021-01-26 13:39:16 +03:00
public export
2021-01-27 21:23:08 +03:00
Zippable List1 where
zipWith f (x ::: xs) (y ::: ys) = f x y ::: zipWith' xs ys
2021-01-26 13:39:16 +03:00
where
2021-01-27 21:23:08 +03:00
zipWith' : List a -> List b -> List c
zipWith' [] _ = []
zipWith' _ [] = []
zipWith' (x :: xs) (y :: ys) = f x y :: zipWith' xs ys
zipWith3 f (x ::: xs) (y ::: ys) (z ::: zs) = f x y z ::: zipWith3' xs ys zs
where
zipWith3' : List a -> List b -> List c -> List d
zipWith3' [] _ _ = []
zipWith3' _ [] _ = []
zipWith3' _ _ [] = []
zipWith3' (x :: xs) (y :: ys) (z :: zs) = f x y z :: zipWith3' xs ys zs
unzipWith f (x ::: xs) = let (b, c) = f x
2021-01-26 13:39:16 +03:00
(bs, cs) = unzipWith' xs in
2021-01-27 21:23:08 +03:00
(b ::: bs, c ::: cs)
where
unzipWith' : List a -> (List b, List c)
unzipWith' [] = ([], [])
unzipWith' (x :: xs) = let (b, c) = f x
(bs, cs) = unzipWith' xs in
(b :: bs, c :: cs)
unzipWith3 f (x ::: xs) = let (b, c, d) = f x
2021-01-26 13:39:16 +03:00
(bs, cs, ds) = unzipWith3' xs in
2021-01-27 21:23:08 +03:00
(b ::: bs, c ::: cs, d ::: ds)
where
unzipWith3' : List a -> (List b, List c, List d)
unzipWith3' [] = ([], [], [])
unzipWith3' (x :: xs) = let (b, c, d) = f x
(bs, cs, ds) = unzipWith3' xs in
(b :: bs, c :: cs, d :: ds)
export
Uninhabited a => Uninhabited (List1 a) where
uninhabited (hd ::: _) = uninhabited hd