The full subcategory associated to a property of objects #
Given a category C and P : ObjectProperty C, we define
a category structure on the type P.FullSubcategory
of objects in C satisfying P.
A subtype-like structure for full subcategories. Morphisms just ignore the property. We don't use
actual subtypes since the simp-normal form ↑X of X.val does not work well for full
subcategories.
- obj : C
The category of which this is a full subcategory
- property : P self.obj
The predicate satisfied by all objects in this subcategory
Instances For
Equations
- One or more equations did not get rendered due to their size.
The forgetful functor from a full subcategory into the original category ("forgetting" the condition).
Instances For
Alias of CategoryTheory.ObjectProperty.FullSubcategory.id_hom.
Alias of CategoryTheory.ObjectProperty.FullSubcategory.comp_hom.
Constructor for morphisms in a full subcategory.
Equations
- CategoryTheory.ObjectProperty.homMk f = { hom := f }
Instances For
The inclusion of a full subcategory is fully faithful.
Equations
- P.fullyFaithfulι = { preimage := fun {X Y : P.FullSubcategory} (f : P.ι.obj X ⟶ P.ι.obj Y) => CategoryTheory.ObjectProperty.homMk f, map_preimage := ⋯, preimage_map := ⋯ }
Instances For
Constructor for isomorphisms in P.FullSubcategory when
P : ObjectProperty C.
Equations
- P.isoMk e = { hom := CategoryTheory.ObjectProperty.homMk e.hom, inv := CategoryTheory.ObjectProperty.homMk e.inv, hom_inv_id := ⋯, inv_hom_id := ⋯ }
Instances For
If P and P' are properties of objects such that P ≤ P', there is
an induced functor P.FullSubcategory ⥤ P'.FullSubcategory.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If h : P ≤ P', then ιOfLE h is fully faithful.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If h : P ≤ P' is an inequality of properties of objects,
this is the obvious isomorphism ιOfLE h ⋙ P'.ι ≅ P.ι.
Equations
Instances For
A functor which maps objects to objects satisfying a certain property induces a lift through the full subcategory of objects satisfying that property.
Equations
Instances For
Composing the lift of a functor through a full subcategory with the inclusion yields the original functor. This is actually true definitionally.
Equations
- P.liftCompιIso F hF = CategoryTheory.Iso.refl ((P.lift F hF).comp P.ι)
Instances For
When h : P ≤ Q, this is the canonical isomorphism
P.lift F hF ⋙ ιOfLE h ≅ Q.lift F _.
Equations
- P.liftCompιOfLEIso F hF h = CategoryTheory.Iso.refl ((P.lift F hF).comp (CategoryTheory.ObjectProperty.ιOfLE h))