Traversable type class
Type classes for traversing collections. The concepts and laws are taken from http://hackage.haskell.org/package/base-4.11.1.0/docs/Data-Traversable.html
Traversable collections are a generalization of functors. Whereas
functors (such as list
) allow us to apply a function to every
element, it does not allow functions which external effects encoded in
a monad. Consider for instance a functor invite : email → io response
that takes an email address, sends an email and waits for a
response. If we have a list guests : list email
, using calling
invite
using map
gives us the following: map invite guests : list
(io response)
. It is not what we need. We need something of type io
(list response)
. Instead of using map
, we can use traverse
to
send all the invites: traverse invite guests : io (list response)
.
traverse
applies invite
to every element of guests
and combines
all the resulting effects. In the example, the effect is encoded in the
monad io
but any applicative functor is accepted by traverse
.
For more on how to use traversable, consider the Haskell tutorial: https://en.wikibooks.org/wiki/Haskell/Traversable
Main definitions
traversable
type class - exposes thetraverse
functionsequence
- based ontraverse
, turns a collection of effects into an effect returning a collection- is_lawful_traversable - laws
Tags
traversable iterator functor applicative
References
- "Applicative Programming with Effects", by Conor McBride and Ross Paterson, Journal of Functional Programming 18:1 (2008) 1-13, online at http://www.soi.city.ac.uk/~ross/papers/Applicative.html
- "The Essence of the Iterator Pattern", by Jeremy Gibbons and Bruno Oliveira, in Mathematically-Structured Functional Programming, 2006, online at http://web.comlab.ox.ac.uk/oucl/work/jeremy.gibbons/publications/#iterator
- "An Investigation of the Laws of Traversals", by Mauro Jaskelioff and Ondrej Rypacek, in Mathematically-Structured Functional Programming, 2012, online at http://arxiv.org/pdf/1202.2919
- app : Π (α : Type ?), F α → G α
- preserves_pure' : ∀ {α : Type ?} (x : α), c.app α (has_pure.pure x) = has_pure.pure x
- preserves_seq' : ∀ {α β : Type ?} (x : F (α → β)) (y : F α), c.app β (x <*> y) = c.app (α → β) x <*> c.app α y
Equations
- applicative_transformation.has_coe_to_fun F G = {F := λ (_x : applicative_transformation F G), Π {α : Type u}, F α → G α, coe := λ (a : applicative_transformation F G), a.app}
- to_functor : functor t
- traverse : Π {m : Type ? → Type ?} [_inst_1 : applicative m] {α β : Type ?}, (α → m β) → t α → m (t β)
Instances
- bitraversable.traversable
- id.traversable
- option.traversable
- list.traversable
- sum.traversable
- tactic.interactive.ac_mono_ctx'.traversable
- vector.traversable
- array.traversable
- free_magma.traversable
- free_add_magma.traversable
- free_semigroup.traversable
- free_add_semigroup.traversable
- buffer.traversable
- dlist.traversable
- lazy_list.traversable
Equations
- to_is_lawful_functor : is_lawful_functor t
- id_traverse : ∀ {α : Type ?} (x : t α), traversable.traverse id.mk x = x
- comp_traverse : ∀ {F G : Type ? → Type ?} [_inst_1_1 : applicative F] [_inst_2 : applicative G] [_inst_3 : is_lawful_applicative F] [_inst_4 : is_lawful_applicative G] {α β γ : Type ?} (f : β → F γ) (g : α → G β) (x : t α), traversable.traverse (functor.comp.mk ∘ functor.map f ∘ g) x = functor.comp.mk (traversable.traverse f <$> traversable.traverse g x)
- traverse_eq_map_id : ∀ {α β : Type ?} (f : α → β) (x : t α), traversable.traverse (id.mk ∘ f) x = id.mk (f <$> x)
- naturality : ∀ {F G : Type ? → Type ?} [_inst_1_1 : applicative F] [_inst_2 : applicative G] [_inst_3 : is_lawful_applicative F] [_inst_4 : is_lawful_applicative G] (η : applicative_transformation F G) {α β : Type ?} (f : α → F β) (x : t α), ⇑η (traversable.traverse f x) = traversable.traverse (⇑η ∘ f) x
Instances
- bitraversable.is_lawful_traversable
- id.is_lawful_traversable
- option.is_lawful_traversable
- list.is_lawful_traversable
- sum.is_lawful_traversable
- vector.is_lawful_traversable
- array.is_lawful_traversable
- free_magma.is_lawful_traversable
- free_add_magma.is_lawful_traversable
- free_semigroup.is_lawful_traversable
- free_add_semigroup.is_lawful_traversable
- buffer.is_lawful_traversable
- dlist.is_lawful_traversable
- lazy_list.is_lawful_traversable
Equations
- id.traversable = {to_functor := applicative.to_functor monad.to_applicative, traverse := λ (_x : Type u_1 → Type u_1) (_x_1 : applicative _x) (_x_1 _x_2 : Type u_1), id}
Equations
- id.is_lawful_traversable = {to_is_lawful_functor := id.is_lawful_traversable._proof_1, id_traverse := id.is_lawful_traversable._proof_2, comp_traverse := id.is_lawful_traversable._proof_3, traverse_eq_map_id := id.is_lawful_traversable._proof_4, naturality := id.is_lawful_traversable._proof_5}
Equations
Equations
Equations
- sum.traverse f (sum.inr x) = sum.inr <$> f x
- sum.traverse f (sum.inl x) = has_pure.pure (sum.inl x)