[ base ] Implement Zippable for several standard types + small cleanup

This commit is contained in:
Denis Buzdalov 2023-10-16 20:21:00 +03:00 committed by G. Allais
parent 7c8076c149
commit 2358a74a29
8 changed files with 97 additions and 33 deletions

View File

@ -220,6 +220,10 @@
* Implements `MonadState` for `Data.Ref` with a named implementation requiring
a particular reference.
* Adds implementations of `Zippable` to `Either`, `Pair`, `Maybe`, `SortedMap`.
* Adds a `Compose` and `FromApplicative` named implementations for `Zippable`.
#### System
* Changes `getNProcessors` to return the number of online processors rather than

View File

@ -249,19 +249,6 @@ Applicative Colist where
_ <*> [] = []
f :: fs <*> a :: as = f a :: (fs <*> as)
public export
Zippable Colist where
zipWith f as bs = [| f as bs |]
zipWith3 f as bs cs = [| f as bs cs |]
unzip xs = (map fst xs, map snd xs)
unzip3 xs = ( map (\(a,_,_) => a) xs
, map (\(_,b,_) => b) xs
, map (\(_,_,c) => c) xs
)
unzipWith f = unzip . map f
unzipWith3 f = unzip3 . map f
public export %hint %inline
ZippableColist : Zippable Colist
ZippableColist = FromApplicative

View File

@ -196,23 +196,9 @@ Applicative Colist1 where
(f ::: fs) <*> (a ::: as) = f a ::: (fs <*> as)
public export
Zippable Colist1 where
zipWith f (x ::: xs) (y ::: ys) = f x y ::: zipWith f xs ys
zipWith3 f (x ::: xs) (y ::: ys) (z ::: zs) =
f x y z ::: zipWith3 f xs ys zs
unzip xs = (map fst xs, map snd xs)
unzip3 xs = ( map (\(a,_,_) => a) xs
, map (\(_,b,_) => b) xs
, map (\(_,_,c) => c) xs
)
unzipWith f = unzip . map f
unzipWith3 f = unzip3 . map f
public export %hint %inline
ZippableColist1 : Zippable Colist1
ZippableColist1 = FromApplicative
--------------------------------------------------------------------------------
-- Interleavings

View File

@ -106,6 +106,17 @@ mirror : Either a b -> Either b a
mirror (Left x) = Right x
mirror (Right x) = Left x
public export
Zippable (Either a) where
zipWith f x y = [| f x y |]
zipWith3 f x y z = [| f x y z |]
unzipWith f (Left e) = (Left e, Left e)
unzipWith f (Right xy) = let (x, y) = f xy in (Right x, Right y)
unzipWith3 f (Left e) = (Left e, Left e, Left e)
unzipWith3 f (Right xyz) = let (x, y, z) = f xyz in (Right x, Right y, Right z)
--------------------------------------------------------------------------------
-- Bifunctor

View File

@ -2,6 +2,8 @@ module Data.Maybe
import Control.Function
import Data.Zippable
%default total
public export
@ -84,3 +86,14 @@ namespace Monoid
public export
[Deep] Semigroup a => Monoid (Maybe a) using Semigroup.Deep where
neutral = Nothing
public export
Zippable Maybe where
zipWith f x y = [| f x y |]
zipWith3 f x y z = [| f x y z |]
unzipWith f Nothing = (Nothing, Nothing)
unzipWith f (Just x) = let (a, b) = f x in (Just a, Just b)
unzipWith3 f Nothing = (Nothing, Nothing, Nothing)
unzipWith3 f (Just x) = let (a, b, c) = f x in (Just a, Just b, Just c)

View File

@ -1,6 +1,7 @@
module Data.SortedMap
import Data.SortedMap.Dependent
import Data.Zippable
%hide Prelude.toList
@ -94,6 +95,13 @@ export
implementation Traversable (SortedMap k) where
traverse f = map M . traverse f . unM
export
implementation Ord k => Zippable (SortedMap k) where
zipWith f mx my = fromList $ mapMaybe (\(k, x) => (k,) . f x <$> lookup k my) $ toList mx
zipWith3 f mx my mz = fromList $ mapMaybe (\(k, x) => (k,) .: f x <$> lookup k my <*> lookup k mz) $ toList mx
unzipWith f m = let m' = map f m in (map fst m', map snd m')
unzipWith3 f m = let m' = map f m in (map fst m', map (fst . snd) m', map (snd . snd) m')
||| Merge two maps. When encountering duplicate keys, using a function to combine the values.
||| Uses the ordering of the first map given.
export

View File

@ -55,3 +55,42 @@ interface Zippable z where
||| @ z the parameterised type
unzip3 : z (a, b, c) -> (z a, z b, z c)
unzip3 = unzipWith3 id
public export
[Compose] Zippable f => Zippable g => Zippable (f . g) where
zipWith f = zipWith $ zipWith f
zipWith3 f = zipWith3 $ zipWith3 f
unzipWith f = unzipWith $ unzipWith f
unzipWith3 f = unzipWith3 $ unzipWith3 f
zip = zipWith zip
zip3 = zipWith3 zip3
unzip = unzipWith unzip
unzip3 = unzipWith3 unzip3
||| Variant of composing and decomposing using existing `Applicative` implementation.
|||
||| This implementation is lazy during unzipping.
||| Caution! It may repeat the whole original work, each time the unzipped elements are used.
|||
||| Please be aware that for every `Applicative` has the same semantics as `Zippable`.
||| Consider `List` as a simple example.
||| However, this implementation is applicable for lazy data types or deferred computations.
public export
[FromApplicative] Applicative f => Zippable f where
zipWith f x y = [| f x y |]
zipWith3 f x y z = [| f x y z |]
unzip u = (map fst u, map snd u)
unzip3 u = (map fst u, map (fst . snd) u, map (snd . snd) u)
unzipWith f = unzip . map f
unzipWith3 f = unzip3 . map f
-- To be moved appropriately as soon as we have some module like `Data.Pair`, like other prelude types have.
public export
Semigroup a => Zippable (Pair a) where
zipWith f (l, x) (r, y) = (l <+> r, f x y)
zipWith3 f (l, x) (m, y) (r, z) = (l <+> m <+> r, f x y z)
unzipWith f (l, x) = bimap (l,) (l,) (f x)
unzipWith3 f (l, x) = bimap (l,) (bimap (l,) (l,)) (f x)

View File

@ -1,6 +1,8 @@
||| Version of Maybe indexed by an `isJust' boolean
module Data.IMaybe
import Data.Zippable
%default total
public export
@ -21,3 +23,17 @@ public export
Applicative (IMaybe True) where
pure = Just
Just f <*> Just x = Just (f x)
public export
Zippable (IMaybe b) where
zipWith f Nothing Nothing = Nothing
zipWith f (Just x) (Just y) = Just $ f x y
zipWith3 f Nothing Nothing Nothing = Nothing
zipWith3 f (Just x) (Just y) (Just z) = Just $ f x y z
unzipWith f Nothing = (Nothing, Nothing)
unzipWith f (Just x) = let (a, b) = f x in (Just a, Just b)
unzipWith3 f Nothing = (Nothing, Nothing, Nothing)
unzipWith3 f (Just x) = let (a, b, c) = f x in (Just a, Just b, Just c)