Skip to content

Commit

Permalink
[ base ] Relevant and irrelevant traversals for Data.Vect.Quantifiers…
Browse files Browse the repository at this point in the history
….All
  • Loading branch information
0xd34df00d authored and gallais committed Oct 16, 2023
1 parent 3e5d8a5 commit 7c8076c
Show file tree
Hide file tree
Showing 2 changed files with 24 additions and 1 deletion.
3 changes: 2 additions & 1 deletion CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -188,7 +188,8 @@
* Generalized `imapProperty` in `Data.List.Quantifiers.All.All`
and `Data.Vect.Quantifiers.All.All`.

* Add `zipPropertyWith` and `remember` to `Data.Vect.Quantifiers.All.All`.
* Add `zipPropertyWith`, `traverseProperty`, `traversePropertyRelevant` and `remember`
to `Data.Vect.Quantifiers.All.All`.

* Add `anyToFin` to `Data.Vect.Quantifiers.Any`,
converting the `Any` witness to the index into the corresponding element.
Expand Down
22 changes: 22 additions & 0 deletions libs/base/Data/Vect/Quantifiers.idr
Original file line number Diff line number Diff line change
Expand Up @@ -173,6 +173,28 @@ namespace All
zipPropertyWith f (px :: pxs) (qx :: qxs)
= f px qx :: zipPropertyWith f pxs qxs

||| A `Traversable`'s `traverse` for `All`,
||| for traversals that don't care about the values of the associated `Vect`.
export
traverseProperty : Applicative f =>
{0 xs : Vect n a} ->
(forall x. p x -> f (q x)) ->
All p xs ->
f (All q xs)
traverseProperty f [] = pure []
traverseProperty f (x :: xs) = [| f x :: traverseProperty f xs |]

||| A `Traversable`'s `traverse` for `All`,
||| in case the elements of the `Vect` that the `All` is proving `p` about are also needed.
export
traversePropertyRelevant : Applicative f =>
{xs : Vect n a} ->
((x : a) -> p x -> f (q x)) ->
All p xs ->
f (All q xs)
traversePropertyRelevant f [] = pure []
traversePropertyRelevant f (x :: xs) = [| f _ x :: traversePropertyRelevant f xs |]

export
All (Show . p) xs => Show (All p xs) where
show pxs = "[" ++ show' "" pxs ++ "]"
Expand Down

0 comments on commit 7c8076c

Please sign in to comment.