Functoriality of the symmetry of equivalences #
Using the calculus of mates in CategoryTheory.Adjunction.Mates
, we prove that passing
to the symmetric equivalence defines an equivalence between C ≌ D
and (D ≌ C)ᵒᵖ
,
and provides the definition of the functor that takes an equivalence to its inverse.
Main definitions #
Equivalence.symmEquiv C D
: the equivalence(C ≌ D) ≌ (D ≌ C)ᵒᵖ
obtained by takingEquivalence.symm
on objects, andconjugateEquiv
on maps.Equivalence.inverseFunctor C D
: The functor(C ≌ D) ⥤ (D ⥤ C)ᵒᵖ
sending an equivalencee
to the functore.inverse
.congrLeftFunctor C D E
: the functor (C ≌ D) ⥤ ((C ⥤ E) ≌ (D ⥤ E))ᵒᵖ that appliesEquivalence.congrLeft
on objects, and whiskers left byconjugateEquiv
on maps.
The forward functor of the equivalence (C ≌ D) ≌ (D ≌ C)ᵒᵖ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The inverse functor of the equivalence (C ≌ D) ≌ (D ≌ C)ᵒᵖ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Taking the symmetric of an equivalence induces an equivalence of categories
(C ≌ D) ≌ (D ≌ C)ᵒᵖ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The inverse
functor that sends a functor to its inverse.
Equations
Instances For
The inverse
functor sends an equivalence to its inverse.
Equations
Instances For
We can compare the way we obtain a natural isomorphism e.inverse ≅ f.inverse
from
an isomorphism e ≌ f
via inverseFunctor
with the way we get one through
Iso.isoInverseOfIsoFunctor
.
An "unopped" version of the equivalence `inverseFunctorObj'.
Equations
Instances For
Promoting Equivalence.congrLeft
to a functor.
Equations
- One or more equations did not get rendered due to their size.