mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-24 04:09:10 +03:00
[ new ] Implement Sized
for Seq
s
This commit is contained in:
parent
0c40a76c2c
commit
276d41d86c
@ -232,6 +232,8 @@
|
|||||||
were made to be indeed lazy by the effect, but their requirements were strengthened
|
were made to be indeed lazy by the effect, but their requirements were strengthened
|
||||||
from `Applicative` to `Monad`.
|
from `Applicative` to `Monad`.
|
||||||
|
|
||||||
|
* Implements `Sized` for `Data.Seq.Sized` and `Data.Seq.Unsized`.
|
||||||
|
|
||||||
#### Papers
|
#### Papers
|
||||||
|
|
||||||
* In `Control.DivideAndConquer`: a port of the paper
|
* In `Control.DivideAndConquer`: a port of the paper
|
||||||
|
@ -207,3 +207,7 @@ public export
|
|||||||
implementation {n : Nat} -> Applicative (Seq n) where
|
implementation {n : Nat} -> Applicative (Seq n) where
|
||||||
pure = replicate n
|
pure = replicate n
|
||||||
(<*>) = zipWith ($)
|
(<*>) = zipWith ($)
|
||||||
|
|
||||||
|
public export
|
||||||
|
implementation Sized (Seq n a) where
|
||||||
|
size (MkSeq s) = size s
|
||||||
|
@ -218,3 +218,7 @@ public export
|
|||||||
public export
|
public export
|
||||||
implementation Monad Seq where
|
implementation Monad Seq where
|
||||||
xs >>= f = foldMap f xs
|
xs >>= f = foldMap f xs
|
||||||
|
|
||||||
|
public export
|
||||||
|
implementation Sized (Seq a) where
|
||||||
|
size (MkSeq s) = size s
|
||||||
|
Loading…
Reference in New Issue
Block a user