Functors #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Defines a functor between categories, extending a prefunctor
between quivers.
Introduces notation C ⥤ D
for the type of all functors from C
to D
.
(Unfortunately the ⇒
arrow (\functor
) is taken by core,
but in mathlib4 we should switch to this.)
The prefunctor between the underlying quivers.
Instances for category_theory.functor.to_prefunctor
- obj : C → D
- map : Π {X Y : C}, (X ⟶ Y) → (self.obj X ⟶ self.obj Y)
- map_id' : (∀ (X : C), self.map (𝟙 X) = 𝟙 (self.obj X)) . "obviously"
- map_comp' : (∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), self.map (f ≫ g) = self.map f ≫ self.map g) . "obviously"
functor C D
represents a functor between categories C
and D
.
To apply a functor F
to an object use F.obj X
, and to a morphism use F.map f
.
The axiom map_id
expresses preservation of identities, and
map_comp
expresses functoriality.
See https://stacks.math.columbia.edu/tag/001B.
Instances for category_theory.functor
- category_theory.functor.has_sizeof_inst
- category_theory.functor.inhabited
- category_theory.functor.category
- category_theory.limits.functor_category_has_limits_of_shape
- category_theory.limits.functor_category_has_colimits_of_shape
- category_theory.limits.functor_category_has_limits_of_size
- category_theory.limits.functor_category_has_colimits_of_size
- category_theory.monoidal.functor_category_monoidal
- category_theory.monoidal.functor_category_braided
- category_theory.monoidal.functor_category_symmetric
- category_theory.monoidal.right_rigid_functor_category
- category_theory.monoidal.left_rigid_functor_category
- category_theory.monoidal.rigid_functor_category
- category_theory.limits.category_theory.functor.has_zero_morphisms
- category_theory.limits.has_zero_object.category_theory.functor.has_zero_object
- category_theory.functor_category_preadditive
- category_theory.limits.functor_category_has_finite_limits
- category_theory.limits.functor_category_has_finite_colimits
- category_theory.abelian.functor_category_abelian
- category_theory.abelian.functor_category_abelian'
- category_theory.functor_category_linear
- Action.category_theory.functor.category_theory.right_rigid_category
- Action.category_theory.functor.category_theory.left_rigid_category
- Action.category_theory.functor.category_theory.rigid_category
- category_theory.functor.monoidal_closed
- category_theory.idempotents.functor_category_is_idempotent_complete
- category_theory.coe_monad
- category_theory.coe_comonad
- category_theory.functor.abelian
- category_theory.functor.finitary_extensive
- category_theory.functor.limits.has_finite_products
- category_theory.functor.cartesian_closed
𝟭 C
is the identity functor on a category C
.
Equations
Instances for category_theory.functor.id
- category_theory.full.id
- category_theory.faithful.id
- category_theory.functor.is_equivalence_refl
- category_theory.functor.id.functor.corepresentable
- category_theory.limits.id_preserves_limits
- category_theory.limits.id_preserves_colimits
- category_theory.limits.id_reflects_limits
- category_theory.limits.id_reflects_colimits
- category_theory.id_creates_limits
- category_theory.id_creates_colimits
- category_theory.limits.id_preserves_finite_limits
- category_theory.limits.id_preserves_finite_colimits
- category_theory.functor.id.additive
- category_theory.functor.id.linear
- category_theory.representably_flat.id
- category_theory.functor.id.exponential_ideal
- category_theory.localization.lifting.id
Equations
- category_theory.functor.inhabited C = {default := 𝟭 C _inst_1}
F ⋙ G
is the composition of a functor F
and a functor G
(F
first, then G
).
Equations
Instances for category_theory.functor.comp
- category_theory.functor.final.comp_has_colimit
- category_theory.functor.initial.comp_has_limit
- category_theory.faithful.comp
- category_theory.full.comp
- category_theory.functor.is_equivalence_trans
- category_theory.adjunction.left_adjoint_of_comp
- category_theory.adjunction.right_adjoint_of_comp
- category_theory.functor.preserves_monomorphisms_comp
- category_theory.functor.preserves_epimorphisms_comp
- category_theory.functor.reflects_monomorphisms_comp
- category_theory.functor.reflects_epimorphisms_comp
- category_theory.functor.comp.reflects_isomorphisms
- category_theory.limits.has_limit_equivalence_comp
- category_theory.limits.has_colimit_equivalence_comp
- category_theory.limits.comp_preserves_limit
- category_theory.limits.comp_preserves_limits_of_shape
- category_theory.limits.comp_preserves_limits
- category_theory.limits.comp_preserves_colimit
- category_theory.limits.comp_preserves_colimits_of_shape
- category_theory.limits.comp_preserves_colimits
- category_theory.limits.comp_reflects_limit
- category_theory.limits.comp_reflects_limits_of_shape
- category_theory.limits.comp_reflects_limits
- category_theory.limits.comp_reflects_colimit
- category_theory.limits.comp_reflects_colimits_of_shape
- category_theory.limits.comp_reflects_colimits
- category_theory.comp_creates_limit
- category_theory.comp_creates_limits_of_shape
- category_theory.comp_creates_limits
- category_theory.comp_creates_colimit
- category_theory.comp_creates_colimits_of_shape
- category_theory.comp_creates_colimits
- category_theory.functor.comp.additive
- category_theory.functor.comp.linear
- category_theory.reflective.comp
- category_theory.comp_comparison_forget_has_limit
- category_theory.comp_comparison_has_limit
- CompHaus_op_to_Frame.faithful
- CompHaus_to_Locale.faithful
- category_theory.limits.comp_preserves_filtered_colimits
- category_theory.limits.comp_preserves_cofiltered_limits
- category_theory.representably_flat.comp
- Top.presheaf.stalk_functor_preserves_mono
- category_theory.Sheaf.category_theory.Sheaf_forget.category_theory.is_right_adjoint
- category_theory.localization.lifting.comp_right
- algebraic_geometry.SheafedSpace.is_open_immersion.has_limit_cospan_forget_of_left
- algebraic_geometry.SheafedSpace.is_open_immersion.has_limit_cospan_forget_of_right
- algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_PresheafedSpace_preserves_pullback_of_left
- algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_Top_preserves_pullback_of_left
- algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_PresheafedSpace_preserves_pullback_of_right
- algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_PresheafedSpace_reflects_pullback_of_left
- algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_PresheafedSpace_reflects_pullback_of_right
- algebraic_geometry.is_open_immersion.has_limit_cospan_forget_of_left
- algebraic_geometry.is_open_immersion.has_limit_cospan_forget_of_right