module Cat.Functor.Joint where
Families of functors🔗
This module generalizes properties of functors (fullness, faithfulness, conservativity, etc.) to families of functors For the rest of this section we will fix a family of functors
: Functor K Cat[ C , D ] → Functor C Cat[ K , D ]
Swap .F₀ c .F₀ k = F .F₀ k .F₀ c
Swap F .F₀ c .F₁ f = F .F₁ f .η c
Swap F .F₀ c .F-id = F .F-id ηₚ c
Swap F .F₀ c .F-∘ f g = F .F-∘ f g ηₚ c
Swap F .F₁ f .η k = F .F₀ k .F₁ f
Swap F .F₁ f .is-natural x y g = sym (F .F₁ g .is-natural _ _ f)
Swap F .F-id = ext λ k → F .F₀ k .F-id
Swap F .F-∘ f g = ext λ k → F .F₀ k .F-∘ f g
Swap F
module _
{oc ℓc od ℓd}
{C : Precategory oc ℓc}
{D : Precategory od ℓd}
where
private
module C = Cat.Reasoning C
module D = Cat.Reasoning D
A family of functors is jointly faithful when: - For any if for every then
: (Idx → Functor C D) → Type _
is-jointly-faithful =
is-jointly-faithful Fᵢ ∀ {x y} {f g : C.Hom x y} → (∀ i → Fᵢ i .F₁ f ≡ Fᵢ i .F₁ g) → f ≡ g
The canonical example of a family of jointly faithful functors are the family of hom-functors indexed by objects of this is a restatment of the coyoneda lemma.
Note that every functor induces a family of functors via the mapping on objects: this family is jointly faithful precisely when is faithful.
swap-faithful→jointly-faithful: (F : Functor K Cat[ C , D ])
→ is-faithful (Swap F)
→ is-jointly-faithful (F .F₀)
= faithful (ext p)
swap-faithful→jointly-faithful F faithful p
jointly-faithful→swap-faithful: (F : Functor K Cat[ C , D ])
→ is-jointly-faithful (F .F₀)
→ is-faithful (Swap F)
= joint (λ i → p ηₚ i) jointly-faithful→swap-faithful F joint p
Jointly conservative functors🔗
A family of functors is jointly conservative when: - For if is an iso for every then is an iso.
: (Idx → Functor C D) → Type _
is-jointly-conservative =
is-jointly-conservative Fᵢ ∀ {x y} {f : C.Hom x y} → (∀ i → D.is-invertible (Fᵢ i .F₁ f)) → C.is-invertible f
We can also rephrase joint-conservativity as a property of a diagram
swap-conservative→jointly-conservative: (F : Functor K Cat[ C , D ])
→ is-conservative (Swap F)
→ is-jointly-conservative (F .F₀)
=
swap-conservative→jointly-conservative F reflect-iso isos (invertible→invertibleⁿ (Swap F .F₁ _) isos)
reflect-iso
jointly-conservative→swap-conservative: (F : Functor K Cat[ C , D ])
→ is-jointly-conservative (F .F₀)
→ is-conservative (Swap F)
=
jointly-conservative→swap-conservative F reflect-iso isos (λ i → is-invertibleⁿ→is-invertible isos i) reflect-iso
Jointly full functors🔗
A diagram of functors is jointly full when the functor is full. Explicitly, is jointly full if a family of morphisms that is natural in implies the mere existence of a with
Note that joint fullness must be a property of a diagram of functors, due to the naturality constraint. This
: Functor K Cat[ C , D ] → Type _
is-jointly-full = is-full (Swap F)
is-jointly-full F
jointly-full-fibre: ∀ {x y}
→ (F : Functor K Cat[ C , D ])
→ is-jointly-full F
→ (gᵢ : ∀ i → D.Hom (F .F₀ i .F₀ x) (F .F₀ i .F₀ y))
→ (∀ {i j} (f : K.Hom i j) → gᵢ j D.∘ F .F₁ f .η x ≡ F .F₁ f .η y D.∘ gᵢ i)
→ ∃[ f ∈ C.Hom x y ] (∀ i → F .F₀ i .F₁ f ≡ gᵢ i)
=
jointly-full-fibre F joint-full gᵢ gᵢ-nat (Σ-map₂ λ p i → p ηₚ i) (joint-full (NT gᵢ λ i j f → gᵢ-nat f)) ∥-∥-map
Jointly fully faithful functors🔗
A diagram of functors is jointly fully faithful when the functor is fully faithful. Explicitly, is jointly faully faithful if there is an equivalence of natural transformations and morphisms
: Functor K Cat[ C , D ] → Type _
is-jointly-fully-faithful = is-fully-faithful (Swap F) is-jointly-fully-faithful F
If a diagram of functors is jointly fully and jointly faithful, then it is jointly fully faithful.
jointly-full+jointly-faithful→jointly-ff: (F : Functor K Cat[ C , D ])
→ is-jointly-full F
→ is-jointly-faithful (F .F₀)
→ is-jointly-fully-faithful F
{x} {y} .is-eqv α =
jointly-full+jointly-faithful→jointly-ff F full faithful
is-prop∙→is-contr img-is-prop $(λ _ → img-is-prop) (Σ-map₂ (λ p → ext p)) $
∥-∥-elim (λ i → α .η i) (λ f → α .is-natural _ _ f)
jointly-full-fibre F full where
: is-prop (Σ[ f ∈ C.Hom x y ] (Swap F .F₁ f ≡ α))
img-is-prop (f , p) (g , q) =
img-is-prop (λ f → Nat-is-set (Swap F .F₁ f) α)
Σ-prop-path (faithful (λ i → p ηₚ i ∙ sym (q ηₚ i)))