Documentation

Mathlib.Control.Traversable.Basic

Traversable type class #

Type classes for traversing collections. The concepts and laws are taken from

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:

Main definitions #

Tags #

traversable iterator functor applicative

References #

structure ApplicativeTransformation (F : Type u → Type v) [Applicative F] (G : Type u → Type w) [Applicative G] :
Type (max (u + 1) v w)

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.

Instances For
    theorem ApplicativeTransformation.preserves_pure' {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (self : ApplicativeTransformation F G) {α : Type u} (x : α) :
    self.app α (pure x) = pure x

    An ApplicativeTransformation preserves pure.

    theorem ApplicativeTransformation.preserves_seq' {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (self : ApplicativeTransformation F G) {α : Type u} {β : 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 α y

    An ApplicativeTransformation intertwines seq.

    instance ApplicativeTransformation.instCoeFunForallForall (F : Type u → Type v) [Applicative F] (G : Type u → Type w) [Applicative G] :
    CoeFun (ApplicativeTransformation F G) fun (x : ApplicativeTransformation F G) => {α : Type u} → F αG α
    Equations
    theorem ApplicativeTransformation.app_eq_coe {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) :
    η.app = fun {α : Type u} => η.app α
    @[simp]
    theorem ApplicativeTransformation.coe_mk {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (f : (α : Type u) → F αG α) (pp : ∀ {α : Type u} (x : α), f α (pure x) = pure x) (ps : ∀ {α β : Type u} (x : F (αβ)) (y : F α), f β (Seq.seq x fun (x : Unit) => y) = Seq.seq (f (αβ) x) fun (x : Unit) => f α y) :
    (fun {α : Type u} => { app := f, preserves_pure' := pp, preserves_seq' := ps }.app α) = f
    theorem ApplicativeTransformation.congr_fun {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) (η' : ApplicativeTransformation F G) (h : η = η') {α : Type u} (x : F α) :
    (fun {α : Type u} => η.app α) x = (fun {α : Type u} => η'.app α) x
    theorem ApplicativeTransformation.congr_arg {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) {α : Type u} {x : F α} {y : F α} (h : x = y) :
    (fun {α : Type u} => η.app α) x = (fun {α : Type u} => η.app α) y
    theorem ApplicativeTransformation.coe_inj {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] ⦃η : ApplicativeTransformation F G ⦃η' : ApplicativeTransformation F G (h : (fun {α : Type u} => η.app α) = fun {α : Type u} => η'.app α) :
    η = η'
    theorem ApplicativeTransformation.ext {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] ⦃η : ApplicativeTransformation F G ⦃η' : ApplicativeTransformation F G (h : ∀ (α : Type u) (x : F α), (fun {α : Type u} => η.app α) x = (fun {α : Type u} => η'.app α) x) :
    η = η'
    theorem ApplicativeTransformation.ext_iff {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] {η : ApplicativeTransformation F G} {η' : ApplicativeTransformation F G} :
    η = η' ∀ (α : Type u) (x : F α), (fun {α : Type u} => η.app α) x = (fun {α : Type u} => η'.app α) x
    theorem ApplicativeTransformation.preserves_pure {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) {α : Type u} (x : α) :
    (fun {α : Type u} => η.app α) (pure x) = pure x
    theorem ApplicativeTransformation.preserves_seq {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) {α : Type u} {β : Type u} (x : F (αβ)) (y : F α) :
    (fun {α : Type u} => η.app α) (Seq.seq x fun (x : Unit) => y) = Seq.seq ((fun {α : Type u} => η.app α) x) fun (x : Unit) => (fun {α : Type u} => η.app α) y
    theorem ApplicativeTransformation.preserves_map {F : Type u → Type v} [Applicative F] [LawfulApplicative F] {G : Type u → Type w} [Applicative G] [LawfulApplicative G] (η : ApplicativeTransformation F G) {α : Type u} {β : Type u} (x : αβ) (y : F α) :
    (fun {α : Type u} => η.app α) (x <$> y) = x <$> (fun {α : Type u} => η.app α) y
    theorem ApplicativeTransformation.preserves_map' {F : Type u → Type v} [Applicative F] [LawfulApplicative F] {G : Type u → Type w} [Applicative G] [LawfulApplicative G] (η : ApplicativeTransformation F G) {α : Type u} {β : Type u} (x : αβ) :
    (fun {α : Type u} => η.app α) Functor.map x = Functor.map x fun {α : Type u} => η.app α

    The identity applicative transformation from an applicative functor to itself.

    Equations
    • ApplicativeTransformation.idTransformation = { app := fun (α : Type u) => id, preserves_pure' := , preserves_seq' := }
    Instances For
      Equations
      • ApplicativeTransformation.instInhabited = { default := ApplicativeTransformation.idTransformation }

      The composition of applicative transformations.

      Equations
      • η'.comp η = { app := fun (α : Type u) (x : F α) => (fun {α : Type u} => η'.app α) ((fun {α : Type u} => η.app α) x), preserves_pure' := , preserves_seq' := }
      Instances For
        @[simp]
        theorem ApplicativeTransformation.comp_apply {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] {H : Type u → Type s} [Applicative H] (η' : ApplicativeTransformation G H) (η : ApplicativeTransformation F G) {α : Type u} (x : F α) :
        (fun {α : Type u} => (η'.comp η).app α) x = (fun {α : Type u} => η'.app α) ((fun {α : Type u} => η.app α) x)
        theorem ApplicativeTransformation.comp_assoc {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] {H : Type u → Type s} [Applicative H] {I : Type u → Type t} [Applicative I] (η'' : ApplicativeTransformation H I) (η' : ApplicativeTransformation G H) (η : ApplicativeTransformation F G) :
        (η''.comp η').comp η = η''.comp (η'.comp η)
        @[simp]
        theorem ApplicativeTransformation.comp_id {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) :
        η.comp ApplicativeTransformation.idTransformation = η
        @[simp]
        theorem ApplicativeTransformation.id_comp {F : Type u → Type v} [Applicative F] {G : Type u → Type w} [Applicative G] (η : ApplicativeTransformation F G) :
        ApplicativeTransformation.idTransformation.comp η = η
        class Traversable (t : Type u → Type u) extends Functor :
        Type (u + 1)

        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 t with an arbitrary applicative functor m.

        Instances
          def sequence {t : Type u → Type u} {α : Type u} {f : Type u → Type u} [Applicative f] [Traversable t] :
          t (f α)f (t α)

          A traversable functor commutes with all applicative functors.

          Equations
          Instances For
            class LawfulTraversable (t : Type u → Type u) [Traversable t] extends LawfulFunctor :

            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.

            Instances
              theorem LawfulTraversable.id_traverse {t : Type u → Type u} [Traversable t] [self : LawfulTraversable t] {α : Type u} (x : t α) :
              traverse pure x = x

              traverse plays well with pure of the identity monad

              theorem LawfulTraversable.comp_traverse {t : Type u → Type u} [Traversable t] [self : LawfulTraversable t] {F : Type u → Type u} {G : Type u → Type u} [Applicative F] [Applicative G] [LawfulApplicative F] [LawfulApplicative G] {α : Type u} {β : Type u} {γ : 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)

              traverse plays well with composition of applicative functors.

              theorem LawfulTraversable.traverse_eq_map_id {t : Type u → Type u} [Traversable t] [self : LawfulTraversable t] {α : Type u} {β : Type u} (f : αβ) (x : t α) :
              traverse (pure f) x = id.mk (f <$> x)

              An axiom for traverse involving pure : β → Id β.

              theorem LawfulTraversable.naturality {t : Type u → Type u} [Traversable t] [self : LawfulTraversable t] {F : Type u → Type u} {G : Type u → Type u} [Applicative F] [Applicative G] [LawfulApplicative F] [LawfulApplicative G] (η : ApplicativeTransformation F G) {α : Type u} {β : Type u} (f : αF β) (x : t α) :
              (fun {α : Type u} => η.app α) (traverse f x) = traverse ((fun {α : Type u} => η.app α) f) x

              The naturality axiom explaining how lawful traversable functors should play with lawful applicative functors.

              Equations
              Equations
              Equations
              def Sum.traverse {σ : Type u} {F : Type u → Type u} [Applicative F] {α : Type u_1} {β : Type u} (f : αF β) :
              σ αF (σ β)

              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
                instance instTraversableSum {σ : Type u} :
                Equations