mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
[ prelude ] Add lacking implementation of Traversable
for Pair
This commit is contained in:
parent
4005b40a95
commit
936b7270ae
@ -57,6 +57,8 @@
|
||||
|
||||
* Improved performance of functions `isNL`, `isSpace`, and `isHexDigit`.
|
||||
|
||||
* Implements `Foldable` and `Traversable` for pairs, right-biased as `Functor`.
|
||||
|
||||
#### Base
|
||||
|
||||
* Deprecates `setByte` and `getByte` from `Data.Buffer` for removal in a future
|
||||
|
@ -133,6 +133,18 @@ public export
|
||||
Functor (Pair a) where
|
||||
map = mapSnd
|
||||
|
||||
%inline
|
||||
public export
|
||||
Foldable (Pair a) where
|
||||
foldr op init (_, x) = x `op` init
|
||||
foldl op init (_, x) = init `op` x
|
||||
null _ = False
|
||||
|
||||
%inline
|
||||
public export
|
||||
Traversable (Pair a) where
|
||||
traverse f (l, r) = (l,) <$> f r
|
||||
|
||||
%inline
|
||||
public export
|
||||
Monoid a => Applicative (Pair a) where
|
||||
|
@ -17,7 +17,6 @@ maybe = %runElab derive
|
||||
either : Foldable (Either err)
|
||||
either = %runElab derive
|
||||
|
||||
%hint
|
||||
pair : Foldable (Pair a)
|
||||
pair = %runElab derive
|
||||
|
||||
|
@ -19,24 +19,24 @@ maybeT = %runElab derive
|
||||
eitherT : Traversable (Either err)
|
||||
eitherT = %runElab derive
|
||||
|
||||
-- Here we don't have a `Foldable (Pair a)` instance and the tactics
|
||||
-- If we didn't have a `Foldable (Pair a)` instance and the tactics
|
||||
-- unfortunately builds
|
||||
-- pairT = let traversePair = (...) in
|
||||
-- MkTraversable @{pairT .foldable} traversePair
|
||||
-- because the `pairT` name is a %hint.
|
||||
--
|
||||
-- TODO: find a way to program defensively against this!
|
||||
failing "pairT is not total"
|
||||
--failing "pairT is not total"
|
||||
--
|
||||
-- %hint total
|
||||
-- pairT : Traversable (Pair a)
|
||||
-- pairT = %runElab derive
|
||||
|
||||
%hint total
|
||||
pairT : Traversable (Pair a)
|
||||
pairT = %runElab derive
|
||||
|
||||
%hint total
|
||||
total
|
||||
pairM : Foldable (Pair a)
|
||||
pairM = %runElab derive
|
||||
|
||||
%hint total
|
||||
total
|
||||
pairT : Traversable (Pair a)
|
||||
pairT = %runElab derive
|
||||
|
||||
|
@ -14,9 +14,6 @@ LOG derive.traversable.clauses:1:
|
||||
LOG derive.traversable.clauses:1:
|
||||
traversePair : {0 a : _} -> {0 a0, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a0 -> f b) -> Pair a a0 -> f (Pair a b)
|
||||
traversePair f (MkPair x2 x3) = (MkPair x2) <$> (f x3)
|
||||
LOG derive.traversable.clauses:1:
|
||||
traversePair : {0 a : _} -> {0 a0, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a0 -> f b) -> Pair a a0 -> f (Pair a b)
|
||||
traversePair f (MkPair x2 x3) = (MkPair x2) <$> (f x3)
|
||||
LOG derive.traversable.clauses:1:
|
||||
traverseConstant : {0 a : _} -> {0 a0, b : Type} -> {0 f : Type -> Type} -> Applicative f => (a0 -> f b) -> Constant a a0 -> f (Constant a b)
|
||||
traverseConstant f (MkConstant x2) = pure (MkConstant x2)
|
||||
|
@ -12,6 +12,6 @@ LOG declare.record.parameters:50: Decided to bind the following extra parameters
|
||||
|
||||
LOG declare.record.parameters:60: We elaborated Main.EtaProof in a non-empty local context.
|
||||
Dropped: [b, a]
|
||||
Remaining type: (p : ($resolved2672 a[1] b[0])) -> Type
|
||||
Remaining type: (p : ($resolved2700 a[1] b[0])) -> Type
|
||||
|
||||
LOG declare.record.parameters:30: Unelaborated type: (%pi RigW Explicit Nothing Main.Product %type)
|
||||
|
Loading…
Reference in New Issue
Block a user