module Cat.Functor.Subcategory where
Subcategories🔗
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.
record Subcat (o' ℓ' : Level) : Type (o ⊔ ℓ ⊔ lsuc o' ⊔ lsuc ℓ') where
no-eta-equality
field
: Ob → Type o'
is-ob : ∀ {x y} (f : Hom x y) → is-ob x → is-ob y → Type ℓ'
is-hom
is-hom-prop: ∀ {x y} (f : Hom x y) (px : is-ob x) (py : is-ob y)
→ is-prop (is-hom f px py)
: ∀ {x} → (px : is-ob x) → is-hom id px px
is-hom-id
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 pz
Morphisms 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 (x .fst) (y .fst)
hom : subcat .is-hom hom (x .snd) (y .snd)
witness
open Subcat-hom
We can then use this data to construct a category.
: Precategory (o ⊔ o') (ℓ ⊔ ℓ')
Subcategory .Precategory.Ob = Σ[ ob ∈ Ob ] is-ob ob
Subcategory .Precategory.Hom = Subcat-hom subcat
Subcategory .Precategory.Hom-set _ _ = Subcat-hom-is-set
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 _ _ _) Subcategory
From pseudomonic functors🔗
There is another way of representing subcategories: By giving a faithful functor .
We 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 .
: 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) =
Faithful-subcat .Hom y z ] (D.to z-es D.∘ F₁ g D.∘ D.from y-es ≡ f)
Σ[ g ∈ C.Subcat.is-hom-prop f (y , y-es) (z , z-es) (g , p) (h , q) =
Faithful-subcat (λ _ → D.Hom-set _ _ _ _) $
Σ-prop-path
faithful $.iso→epic (y-es D.Iso⁻¹) _ _ $
D.iso→monic z-es _ _ $
D
p ∙ sym q.Subcat.is-hom-id (y , y-es) =
Faithful-subcat .id , ap₂ D._∘_ refl (D.eliml F-id) ∙ D.invl y-es
C.Subcat.is-hom-∘
Faithful-subcat {f = f} {g = g} {x , x-es} {y , y-es} {z , z-es} (h , p) (i , q) =
(h C.∘ i) ,
.push-inner (F-∘ h i)
D.insert-inner (D.invr y-es)
·· D._∘_ (sym (D.assoc _ _ _) ∙ p) q ·· ap₂ D
There is an equivalence between canonical subcategory associated with and .
: 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-fully-faithful Faithful-subcat-domain
Faithful-subcat-domain-is-ff {x = x , x' , x-es} {y = y , y' , y-es} =
Faithful-subcat-domain-is-ff
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))
: is-split-eso Faithful-subcat-domain
Faithful-subcat-domain-is-split-eso =
Faithful-subcat-domain-is-split-eso x (F₀ x , x , D.id-iso) , C.id-iso
There is a faithful functor from a subcategory on to .
: Functor (Subcategory subcat) 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
Forget-subcat
: is-faithful Forget-subcat
is-faithful-Forget-subcat = ext is-faithful-Forget-subcat
Furthermore, 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 : is-ob x} {py : is-ob y}
→ is-invertible f → subcat .is-hom f px py)
→ is-pseudomonic Forget-subcat
.is-pseudomonic.faithful =
is-pseudomonic-Forget-subcat invert
is-faithful-Forget-subcat.is-pseudomonic.isos-full f =
is-pseudomonic-Forget-subcat invert
pure $.make-iso
Sub(sub-hom (f .to) (invert (iso→invertible f)))
(sub-hom (f .from) (invert (iso→invertible (f Iso⁻¹))))
(ext (f .invl))
(ext (f .invr))
, ≅-path refl
Univalent subcategories🔗
Let be a univalent category. A subcategory of is univalent when the predicate on objects is a proposition.
: ∀ {x y : Σ[ x ∈ Ob ] is-ob x} → x Sub.≅ y → x .fst ≅ y .fst
subcat-iso→iso =
subcat-iso→iso f (Sub.to f .hom) (Sub.from f .hom)
make-iso (ap hom (Sub.invl f)) (ap hom (Sub.invr f))
subcat-is-category: is-category C
→ (∀ x → is-prop (is-ob x))
→ is-category (Subcategory subcat)
.to-path {a , pa} {b , pb} f =
subcat-is-category cat ob-prop (cat .to-path (subcat-iso→iso f))
Σ-prop-path ob-prop .to-path-over p =
subcat-is-category cat ob-prop .≅-pathp refl _ $
Sub_ $
Subcat-hom-pathp refl (λ _ → to) (cat .to-path-over (subcat-iso→iso p)) apd