Preserving (co)equalizers #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Constructions to relate the notions of preserving (co)equalizers and reflecting (co)equalizers to concrete (co)forks.
In particular, we show that equalizer_comparison f g G is an isomorphism iff G preserves
the limit of the parallel pair f,g, as well as the dual result.
The map of a fork is a limit iff the fork consisting of the mapped morphisms is a limit. This
essentially lets us commute fork.of_ι with functor.map_cone.
Equations
- category_theory.limits.is_limit_map_cone_fork_equiv G w = (category_theory.limits.is_limit.postcompose_hom_equiv (category_theory.limits.diagram_iso_parallel_pair (category_theory.limits.parallel_pair f g ⋙ G)) (G.map_cone (category_theory.limits.fork.of_ι h w))).symm.trans (category_theory.limits.is_limit.equiv_iso_limit (category_theory.limits.fork.ext (category_theory.iso.refl ((category_theory.limits.cones.postcompose (category_theory.limits.diagram_iso_parallel_pair (category_theory.limits.parallel_pair f g ⋙ G)).hom).obj (G.map_cone (category_theory.limits.fork.of_ι h w))).X) _))
The property of preserving equalizers expressed in terms of forks.
The property of reflecting equalizers expressed in terms of forks.
If G preserves equalizers and C has them, then the fork constructed of the mapped morphisms of
a fork is a limit.
If the equalizer comparison map for G at (f,g) is an isomorphism, then G preserves the
equalizer of (f,g).
Equations
- category_theory.limits.preserves_equalizer.of_iso_comparison G f g = category_theory.limits.preserves_limit_of_preserves_limit_cone (category_theory.limits.equalizer_is_equalizer f g) (⇑((category_theory.limits.is_limit_map_cone_fork_equiv G _).symm) (category_theory.limits.limit.is_limit (category_theory.limits.parallel_pair (G.map f) (G.map g))).of_point_iso)
If G preserves the equalizer of (f,g), then the equalizer comparison map for G at (f,g) is
an isomorphism.
The map of a cofork is a colimit iff the cofork consisting of the mapped morphisms is a colimit.
This essentially lets us commute cofork.of_π with functor.map_cocone.
Equations
- category_theory.limits.is_colimit_map_cocone_cofork_equiv G w = (category_theory.limits.is_colimit.precompose_inv_equiv (category_theory.limits.diagram_iso_parallel_pair (category_theory.limits.parallel_pair f g ⋙ G)) (G.map_cocone (category_theory.limits.cofork.of_π h w))).symm.trans (category_theory.limits.is_colimit.equiv_iso_colimit (category_theory.limits.cofork.ext (category_theory.iso.refl ((category_theory.limits.cocones.precompose (category_theory.limits.diagram_iso_parallel_pair (category_theory.limits.parallel_pair f g ⋙ G)).inv).obj (G.map_cocone (category_theory.limits.cofork.of_π h w))).X) _))
The property of preserving coequalizers expressed in terms of coforks.
The property of reflecting coequalizers expressed in terms of coforks.
If G preserves coequalizers and C has them, then the cofork constructed of the mapped morphisms
of a cofork is a colimit.
If the coequalizer comparison map for G at (f,g) is an isomorphism, then G preserves the
coequalizer of (f,g).
Equations
- category_theory.limits.of_iso_comparison G f g = category_theory.limits.preserves_colimit_of_preserves_colimit_cocone (category_theory.limits.coequalizer_is_coequalizer f g) (⇑((category_theory.limits.is_colimit_map_cocone_cofork_equiv G _).symm) (category_theory.limits.colimit.is_colimit (category_theory.limits.parallel_pair (G.map f) (G.map g))).of_point_iso)
If G preserves the coequalizer of (f,g), then the coequalizer comparison map for G at (f,g)
is an isomorphism.
Any functor preserves coequalizers of split pairs.
Equations
- category_theory.limits.preserves_split_coequalizers G f g = category_theory.limits.preserves_colimit_of_preserves_colimit_cocone (category_theory.has_split_coequalizer.is_split_coequalizer f g).is_coequalizer (⇑((category_theory.limits.is_colimit_map_cocone_cofork_equiv G _).symm) ((category_theory.has_split_coequalizer.is_split_coequalizer f g).map G).is_coequalizer)