[ new ] Implement Sized for Seqs

This commit is contained in:
Denis Buzdalov 2023-09-27 21:25:11 +03:00 committed by G. Allais
parent 0c40a76c2c
commit 276d41d86c
3 changed files with 10 additions and 0 deletions

View File

@ -232,6 +232,8 @@
were made to be indeed lazy by the effect, but their requirements were strengthened
from `Applicative` to `Monad`.
* Implements `Sized` for `Data.Seq.Sized` and `Data.Seq.Unsized`.
#### Papers
* In `Control.DivideAndConquer`: a port of the paper

View File

@ -207,3 +207,7 @@ public export
implementation {n : Nat} -> Applicative (Seq n) where
pure = replicate n
(<*>) = zipWith ($)
public export
implementation Sized (Seq n a) where
size (MkSeq s) = size s

View File

@ -218,3 +218,7 @@ public export
public export
implementation Monad Seq where
xs >>= f = foldMap f xs
public export
implementation Sized (Seq a) where
size (MkSeq s) = size s