module Cat.Functor.Subcategory whereSubcategoriesπ
A subcategory is specified by a predicate on objects, and a predicate on morphisms between objects within that is closed under identities and composites.
To start, we package up all the data required to define a subcategory up into a record. Note that we omit the requirement that the predicate on objects is a proposition; this tends to be ill-behaved unless the is univalent.
module _ {o β} (C : Precategory o β) where
open Cat.Reasoning C record Subcat (o' β' : Level) : Type (o β β β lsuc o' β lsuc β') where
no-eta-equality
field
is-ob : Ob β Type o'
is-hom : β {x y} (f : Hom x y) β is-ob x β is-ob y β Type β'
is-hom-prop
: β {x y} (f : Hom x y) (px : is-ob x) (py : is-ob y)
β is-prop (is-hom f px py)
is-hom-id : β {x} β (px : is-ob x) β is-hom id px px
is-hom-β
: β {x y z} {f : Hom y z} {g : Hom x y}
β {px : is-ob x} {py : is-ob y} {pz : is-ob z}
β is-hom f py pz β is-hom g px py
β is-hom (f β g) px pzMorphisms of wide subcategories are defined as morphisms in where holds.
module _ {o o' β β'} {C : Precategory o β} (subcat : Subcat C o' β') where
open Cat.Reasoning C
open Subcat subcat
record Subcat-hom (x y : Ξ£[ ob β Ob ] (is-ob ob)) : Type (β β β') where
no-eta-equality
constructor sub-hom
field
hom : Hom (x .fst) (y .fst)
witness : is-hom hom (x .snd) (y .snd)
open Subcat-hommodule _ {o β} {C : Precategory o β} where
private module C = Precategory C
instance
Membership-subcat-ob : β {o' β'} β Membership C.Ob (Subcat C o' β') _
Membership-subcat-ob = record { _β_ = Ξ» o S β o β S .Subcat.is-ob }
module _ {o o' β β'} {C : Precategory o β} {S : Subcat C o' β'} where
open Cat.Reasoning C
open Subcat S
Subcat-hom-pathp
: {x x' y y' : Ξ£[ ob β C ] (ob β S)}
β {f : Subcat-hom S x y} {g : Subcat-hom S x' y'}
β (p : x β‘ x') (q : y β‘ y')
β PathP (Ξ» i β Hom (p i .fst) (q i .fst)) (f .hom) (g .hom)
β PathP (Ξ» i β Subcat-hom S (p i) (q i)) f g
Subcat-hom-pathp p q r i .hom = r i
Subcat-hom-pathp {f = f} {g = g} p q r i .witness =
is-propβpathp (Ξ» i β is-hom-prop (r i) (p i .snd) (q i .snd)) (f .witness) (g .witness) i
instance
Extensional-subcat-hom
: β {βr x y} β¦ sa : Extensional (Hom (x .fst) (y .fst)) βr β¦
β Extensional (Subcat-hom S x y) βr
Extensional-subcat-hom β¦ sa β¦ = injectionβextensional!
(Subcat-hom-pathp refl refl) sa
Funlike-Subcat-hom
: β {β β'} {A : Type β} {B : A β Type β'} {x y}
β β¦ _ : Funlike (Hom (x .fst) (y .fst)) A B β¦ β Funlike (Subcat-hom S x y) A B
Funlike-Subcat-hom β¦ i β¦ = record { _Β·_ = Ξ» f x β apply (f .hom) x }
H-Level-Subcat-hom : β {x y n} β H-Level (Subcat-hom S x y) (2 + n)
H-Level-Subcat-hom = basic-instance 2 $ Isoβis-hlevel 2 eqv $
Ξ£-is-hlevel 2 (Hom-set _ _) Ξ» _ β
is-hlevel-suc 1 (is-hom-prop _ _ _)
where unquoteDecl eqv = declare-record-iso eqv (quote Subcat-hom)We can then use this data to construct a category.
module _ {o o' β β'} {C : Precategory o β} (subcat : Subcat C o' β') where
open Cat.Reasoning C
open Subcat subcat Subcategory : Precategory (o β o') (β β β')
Subcategory .Precategory.Ob = β«β subcat
Subcategory .Precategory.Hom = Subcat-hom subcat
Subcategory .Precategory.Hom-set _ _ = hlevel 2
Subcategory .Precategory.id .hom = id
Subcategory .Precategory.id .witness = is-hom-id _
Subcategory .Precategory._β_ f g .hom = f .hom β g .hom
Subcategory .Precategory._β_ f g .witness = is-hom-β (f .witness) (g .witness)
Subcategory .Precategory.idr f = ext (idr _)
Subcategory .Precategory.idl f = ext (idl _)
Subcategory .Precategory.assoc f g h = ext (assoc _ _ _)From pseudomonic functorsπ
There is another way of representing subcategories: By giving a faithful functor
module _
{o o' β β'} {C : Precategory o β} {D : Precategory o' β'}
{F : Functor C D}
(faithful : is-faithful F)
where
open Functor F
private
module C = Cat.Reasoning C
module D = Cat.Reasoning DWe construct a subcategory from a faithful functor by restricting to the objects in the essential image of and restricting the morphisms to those that lie in the image of
Faithful-subcat : Subcat D (o β β') (β β β')
Faithful-subcat .Subcat.is-ob x = Essential-fibre F x
Faithful-subcat .Subcat.is-hom f (y , y-es) (z , z-es) =
Ξ£[ g β C.Hom y z ] (D.to z-es D.β Fβ g D.β D.from y-es β‘ f)
Faithful-subcat .Subcat.is-hom-prop f (y , y-es) (z , z-es) (g , p) (h , q) =
Ξ£-prop-path! $
faithful $
D.isoβepic (y-es D.Isoβ»ΒΉ) _ _ $
D.isoβmonic z-es _ _ $
p β sym q
Faithful-subcat .Subcat.is-hom-id (y , y-es) =
C.id , apβ D._β_ refl (D.eliml F-id) β D.invl y-es
Faithful-subcat .Subcat.is-hom-β
{f = f} {g = g} {x , x-es} {y , y-es} {z , z-es} (h , p) (i , q) =
(h C.β i) ,
D.push-inner (F-β h i)
ββ D.insert-inner (D.invr y-es)
ββ apβ D._β_ (sym (D.assoc _ _ _) β p) qThere is an equivalence between canonical subcategory associated with and
Faithful-subcat-domain : Functor (Subcategory Faithful-subcat) C
Faithful-subcat-domain .Functor.Fβ (x , x-es) = x-es .fst
Faithful-subcat-domain .Functor.Fβ f = f .witness .fst
Faithful-subcat-domain .Functor.F-id = refl
Faithful-subcat-domain .Functor.F-β _ _ = refl
Faithful-subcat-domain-is-ff : is-fully-faithful Faithful-subcat-domain
Faithful-subcat-domain-is-ff {x = x , x' , x-es} {y = y , y' , y-es} =
is-isoβis-equiv $ iso
(Ξ» f β sub-hom (D.to y-es D.β Fβ f D.β D.from x-es) (f , refl))
(Ξ» _ β refl)
(Ξ» f β ext (f .witness .snd))
Faithful-subcat-domain-is-split-eso : is-split-eso Faithful-subcat-domain
Faithful-subcat-domain-is-split-eso x =
(Fβ x , x , D.id-iso) , C.id-isoThere is a faithful functor from a subcategory on to
module _ {o o' β β'} {C : Precategory o β} {S : Subcat C o' β'} where
open Cat.Reasoning C
private module Sub = Cat.Reasoning (Subcategory S)
open Subcat S Forget-subcat : Functor (Subcategory S) C
Forget-subcat .Functor.Fβ (x , _) = x
Forget-subcat .Functor.Fβ f = f .hom
Forget-subcat .Functor.F-id = refl
Forget-subcat .Functor.F-β _ _ = refl
is-faithful-Forget-subcat : is-faithful Forget-subcat
is-faithful-Forget-subcat = extFurthermore, if the subcategory contains all of the isomorphisms of then the forgetful functor is pseudomonic.
is-pseudomonic-Forget-subcat
: (β {x y} {f : Hom x y} {px : x β S} {py : y β S}
β is-invertible f β is-hom f px py)
β is-pseudomonic Forget-subcat
is-pseudomonic-Forget-subcat invert .is-pseudomonic.faithful =
is-faithful-Forget-subcat
is-pseudomonic-Forget-subcat invert .is-pseudomonic.isos-full f =
pure $ Sub.make-iso
(sub-hom (f .to) (invert (isoβinvertible f)))
(sub-hom (f .from) (invert (isoβinvertible (f Isoβ»ΒΉ))))
(ext (f .invl))
(ext (f .invr)) , ext reflUnivalent subcategoriesπ
Let be a univalent category. A subcategory of is univalent when the predicate on objects is a proposition.
subcat-isoβiso : β {x y : Ξ£[ x β Ob ] (x β S)} β x Sub.β
y β x .fst β
y .fst
subcat-isoβiso f = make-iso (Sub.to f .hom) (Sub.from f .hom)
(ap hom (Sub.invl f)) (ap hom (Sub.invr f))
subcat-is-category
: is-category C
β (β x β is-prop (x β S))
β is-category (Subcategory S)
subcat-is-category cat ob-prop .to-path {a , pa} {b , pb} f =
Ξ£-prop-path ob-prop (cat .to-path (subcat-isoβiso f))
subcat-is-category cat ob-prop .to-path-over p =
Sub.β
-pathp refl _ $
Subcat-hom-pathp refl _ $
apd (Ξ» _ β to) (cat .to-path-over (subcat-isoβiso p))