module Cat.Diagram.Coseparator {o ℓ} (C : Precategory o ℓ) where
open Cat.Reasoning C
open _=>_
Coseparating objects🔗
A coseparating (or cogenerating) object is formally the dual of a separator.
: Ob → Type _
is-coseparator =
is-coseparator c ∀ {x y} {f g : Hom x y}
→ (∀ (m : Hom y c) → m ∘ f ≡ m ∘ g)
→ f ≡ g
Equivalently, an object is a coseparator if the hom functor is faithful.
: ∀ {s} → is-coseparator s → is-faithful (よ₀ C s)
coseparator→faithful = cos (happly p)
coseparator→faithful cos p
: ∀ {s} → is-faithful (よ₀ C s) → is-coseparator s
faithful→coseparator = faithful (ext p) faithful→coseparator faithful p
Coseparating families🔗
Likewise, a coseparating family is dual to a separating family.
: ∀ {ℓi} {Idx : Type ℓi} → (Idx → Ob) → Type _
is-coseparating-family =
is-coseparating-family s ∀ {x y} {f g : Hom x y}
→ (∀ {i} (mᵢ : Hom y (s i)) → mᵢ ∘ f ≡ mᵢ ∘ g )
→ f ≡ g
Coseparators and powers🔗
Equivalently to approximating objects with separators and copowers, we may approximate them with coseparators and powers.
module _ (powers : (I : Set ℓ) → has-products-indexed-by C ∣ I ∣) where
open Powers powers
coseparator→mono: ∀ {s x} → is-coseparator s → is-monic (⋔!.tuple (Hom x s) s λ f → f)
{s} {x} cosep f g p = cosep λ m →
coseparator→mono (sym $ ⋔!.commute _ _) ⟩
m ∘ f ≡⟨ pushl .π _ _ m ∘ (⋔!.tuple _ _ λ f → f) ∘ f ≡⟨ refl⟩∘⟨ p ⟩
⋔!.π _ _ m ∘ (⋔!.tuple _ _ λ f → f) ∘ g ≡⟨ pulll $ ⋔!.commute _ _ ⟩
⋔!
m ∘ g ∎
mono→coseparator: ∀ {s} → (∀ {x} → is-monic (⋔!.tuple (Hom x s) s λ f → f)) → is-coseparator s
{f = f} {g = g} p = monic f g $ ⋔!.unique₂ _ _ λ m →
mono→coseparator monic _ _ _ ∙ p _ ∙ sym (assoc _ _ _)
assoc
coseparating-family→mono: ∀ (Idx : Set ℓ) (sᵢ : ∣ Idx ∣ → Ob)
→ is-coseparating-family sᵢ
→ ∀ {x} → is-monic (∏!.tuple (Σ[ i ∈ ∣ Idx ∣ ] Hom x (sᵢ i)) (sᵢ ⊙ fst) snd )
= cosep λ {i} mᵢ →
coseparating-family→mono Idx sᵢ cosep f g p (sym $ ∏!.commute _ _) ⟩
mᵢ ∘ f ≡⟨ pushl .π _ _ (i , mᵢ) ∘ (∏!.tuple _ _ snd) ∘ f ≡⟨ refl⟩∘⟨ p ⟩
∏!.π _ _ (i , mᵢ) ∘ (∏!.tuple _ _ snd) ∘ g ≡⟨ pulll $ ∏!.commute _ _ ⟩
∏!
mᵢ ∘ g ∎
coseparating-family→make-mono: ∀ (Idx : Set ℓ) (sᵢ : ∣ Idx ∣ → Ob)
→ is-coseparating-family sᵢ
→ ∀ {x} → x ↪ ∏!.ΠF (Σ[ i ∈ ∣ Idx ∣ ] Hom x (sᵢ i)) (sᵢ ⊙ fst)
= make-mono _ $
coseparating-family→make-mono Idx sᵢ cosep
coseparating-family→mono Idx sᵢ cosep
mono→coseparating-family: ∀ (Idx : Set ℓ)
→ (sᵢ : ∣ Idx ∣ → Ob)
→ (∀ {x} → is-monic (∏!.tuple (Σ[ i ∈ ∣ Idx ∣ ] Hom x (sᵢ i)) (sᵢ ⊙ fst) snd))
→ is-coseparating-family sᵢ
{f = f} {g = g} p =
mono→coseparating-family Idx sᵢ monic .unique₂ _ _ λ (i , mᵢ) →
monic f g $ ∏!_ _ _ ∙ p _ ∙ sym (assoc _ _ _) assoc