Functors with two arguments #
This file defines bifunctors.
A bifunctor is a function F : Type* → Type* → Type*
along with a bimap which turns F α β
into
F α' β'
given two functions α → α'
and β → β'
. It further
- respects the identity:
bimap id id = id
- composes in the obvious way:
(bimap f' g') ∘ (bimap f g) = bimap (f' ∘ f) (g' ∘ g)
Main declarations #
Bifunctor
: A typeclass for the bare bimap of a bifunctor.LawfulBifunctor
: A typeclass asserting this bimap respects the bifunctor laws.
Equations
- Bifunctor.const = { bimap := fun {α α' : Type ?u.14} {β β' : Type ?u.13} (f : α → α') (x : β → β') => f }
@[instance 10]
instance
Bifunctor.lawfulFunctor
{F : Type u₀ → Type u₁ → Type u₂}
[Bifunctor F]
[LawfulBifunctor F]
{α : Type u₀}
:
LawfulFunctor (F α)
instance
Function.bicompl.lawfulBifunctor
{F : Type u₀ → Type u₁ → Type u₂}
[Bifunctor F]
(G : Type u_1 → Type u₀)
(H : Type u_2 → Type u₁)
[Functor G]
[Functor H]
[LawfulFunctor G]
[LawfulFunctor H]
[LawfulBifunctor F]
:
LawfulBifunctor (bicompl F G H)
instance
Function.bicompr.bifunctor
{F : Type u₀ → Type u₁ → Type u₂}
[Bifunctor F]
(G : Type u₂ → Type u_1)
[Functor G]
:
Equations
- Function.bicompr.bifunctor G = { bimap := fun {_α α' : Type ?u.54} {_β β' : Type ?u.53} (f : _α → α') (f' : _β → β') (x : Function.bicompr G F _α _β) => bimap f f' <$> x }
instance
Function.bicompr.lawfulBifunctor
{F : Type u₀ → Type u₁ → Type u₂}
[Bifunctor F]
(G : Type u₂ → Type u_1)
[Functor G]
[LawfulFunctor G]
[LawfulBifunctor F]
:
LawfulBifunctor (bicompr G F)