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 #
- Traversabletype class - exposes the- traversefunction
- sequence- based on- traverse, turns a collection of effects into an effect returning a collection
- LawfulTraversable- laws for a traversable functor
- ApplicativeTransformation- the notion of a natural transformation for applicative functors
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
A transformation between applicative functors.  It is a natural
transformation such that app preserves the Pure.pure and
Functor.map (<*>) operations. See
ApplicativeTransformation.preserves_map for naturality.
- app : (α : Type u) → F α → G αThe function on objects defined by an ApplicativeTransformation.
- An - ApplicativeTransformationpreserves- pure.
- preserves_seq' : ∀ {α β : Type u} (x : F (α → β)) (y : F α), self.app β (Seq.seq x fun (x : Unit) => y) = Seq.seq (self.app (α → β) x) fun (x : Unit) => self.app α yAn ApplicativeTransformationintertwinesseq.
Instances For
An ApplicativeTransformation preserves pure.
An ApplicativeTransformation intertwines seq.
Equations
- ApplicativeTransformation.instCoeFunForallForall F G = { coe := fun (η : ApplicativeTransformation F G) {α : Type u} => η.app α }
The identity applicative transformation from an applicative functor to itself.
Equations
Instances For
Equations
- ApplicativeTransformation.instInhabited = { default := ApplicativeTransformation.idTransformation }
The composition of applicative transformations.
Equations
Instances For
A traversable functor is a functor along with a way to commute
with all applicative functors (see sequence).  For example, if t
is the traversable functor List and m is the applicative functor
IO, then given a function f : α → IO β, the function Functor.map f is
List α → List (IO β), but traverse f is List α → IO (List β).
- map : {α β : Type u} → (α → β) → t α → t β
- mapConst : {α β : Type u} → α → t β → t α
- traverse : {m : Type u → Type u} → [inst : Applicative m] → {α β : Type u} → (α → m β) → t α → m (t β)The function commuting a traversable functor twith an arbitrary applicative functorm.
Instances
A traversable functor is lawful if its traverse satisfies a
number of additional properties.  It must send pure : α → Id α to pure,
send the composition of applicative functors to the composition of the
traverse of each, send each function f to fun x ↦ f <$> x, and
satisfy a naturality condition with respect to applicative
transformations.
- map_const : ∀ {α β : Type u}, Functor.mapConst = Functor.map ∘ Function.const β
- traverseplays well with- pureof the identity monad
- comp_traverse : ∀ {F G : Type u → Type u} [inst : Applicative F] [inst_1 : Applicative G] [inst_2 : LawfulApplicative F] [inst_3 : LawfulApplicative G] {α β γ : Type u} (f : β → F γ) (g : α → G β) (x : t α), traverse (Functor.Comp.mk ∘ Functor.map f ∘ g) x = Functor.Comp.mk (traverse f <$> traverse g x)traverseplays well with composition of applicative functors.
- naturality : ∀ {F G : Type u → Type u} [inst : Applicative F] [inst_1 : Applicative G] [inst_2 : LawfulApplicative F] [inst_3 : LawfulApplicative G] (η : ApplicativeTransformation F G) {α β : Type u} (f : α → F β) (x : t α), (fun {α : Type u} => η.app α) (traverse f x) = traverse ((fun {α : Type u} => η.app α) ∘ f) xThe naturality axiom explaining how lawful traversable functors should play with lawful applicative functors. 
Instances
traverse plays well with pure of the identity monad
traverse plays well with composition of applicative functors.
The naturality axiom explaining how lawful traversable functors should play with lawful applicative functors.
Equations
- instTraversableId = Traversable.mk fun {m : Type u_1 → Type u_1} [Applicative m] {α β : Type u_1} => id
Equations
- instTraversableOption = Traversable.mk fun {m : Type u_1 → Type u_1} [Applicative m] {α β : Type u_1} => Option.traverse
Equations
- instTraversableList = Traversable.mk fun {m : Type u_1 → Type u_1} [Applicative m] {α β : Type u_1} => List.traverse
Defines a traverse function on the second component of a sum type.
This is used to give a Traversable instance for the functor σ ⊕ -.
Equations
Instances For
Equations
- instTraversableSum = Traversable.mk (@Sum.traverse σ)