module Cat.Diagram.Omega whereSubobject classifiers🔗
In an arbitrary category a subobject of is, colloquially speaking, given by a monomorphism Formally, though, we must consider the object to be part of this data, since we can only speak of morphisms if we already know their domain and codomain. The generality in the definition means that the notion of subobject applies to completely arbitrary but in completely arbitrary situations, the notion might be very badly behaved.
There are several observations we can make about that “tame” the behaviour of For example, if it has pullbacks, then is a fibration, so that there is a universal way of re-indexing a subobject along a morphism and if it has images, it’s even a bifibration, so that each of these reindexings has a left adjoint, giving an interpretation of existential quantifiers. If is a regular category, existential quantifiers and substitution commute, restricting the behaviour of subobjects even further.
However, there is one assumption we can make about that rules them all: it has a generic subobject, so that is equivalent to a given in A generic subobject is a morphism so that, for any a subobject there is a unique arrow making the square
into a pullback. We can investigate this definition by instantiating
it in
which does have a generic subobject. Given an embedding
we have a family of propositions
which maps
to
the fibre of
over
The square above also computes a fibre, this time of
and saying that it is
means that
assigns
to precisely those
which are in
module _ {o ℓ} (C : Precategory o ℓ) where
  open Cat.Reasoning C
  open Subobjs C  record is-generic-subobject {Ω} (true : Subobject Ω) : Type (o ⊔ ℓ) where
    field
      name : ∀ {X} (m : Subobject X) → Hom X Ω
      classifies
        : ∀ {X} (m : Subobject X)
        → is-pullback-along C (m .map) (name m) (true .map)
      unique
        : ∀ {X} {m : Subobject X} {nm : Hom X Ω}
        → is-pullback-along C (m .map) nm (true .map)
        → nm ≡ name mWe follow (Johnstone 2002, A1.6) for our terminology: the morphism is called the generic subobject, and is the subobject classifier. This differs from the terminology in the nLab, where the morphism is called the subobject classifier.
  record Subobject-classifier : Type (o ⊔ ℓ) where
    field
      {Ω}     : Ob
      true    : Subobject Ω
      generic : is-generic-subobject true
    open is-generic-subobject generic publicWhile the definition of is-generic-subobject can be stated
without assuming that
has any limits at all, we’ll need to assume that
has pullbacks to get anything of value done.
Before we get started, however, we’ll prove a helper theorem that serves
as a “smart constructor” for Subobject-classifier, simplifying the
verification of the universal property in case
is univalent and has all finite limits (in
particular, a terminal object).
module _ {o ℓ} {C : Precategory o ℓ} (fc : Finitely-complete C) (cat : is-category C) where
  open Subobject-classifier
  open is-generic-subobject
  open Cat.Reasoning C
  open Subr (fc .Finitely-complete.pullbacks)
  open Terminal (fc .Finitely-complete.terminal) using (top)
  point→subobject : ∀ {A} (it : Hom top A) → Subobject A
  point→subobject it .dom = top
  point→subobject it .map = it
  point→subobject it .monic g h x = Terminal.!-unique₂ (fc .Finitely-complete.terminal) _ _Assuming that we have a terminal object, we no longer need to specify
the arrow
as an arbitrary element of
— we can instead ask for a point
instead. Since we have pullbacks, we also have the change-of-base
operation
for any
which allows us to simplify the condition that
“
is the pullback of
along
”,
which requires constructing a pullback square, to the simpler
The theorem is that it suffices to ask for:
- The point corresponding to the generic subobject,
 - An operation which maps subobjects to their names,
 - A witness that, for any we have and
 - A witness that, for any we have
 
These two conditions amount to saying that and form an equivalence between the sets and for all Note that we do not need to assume naturality for since it is an inverse equivalence to which is already natural.
  from-classification
    : ∀ {Ω} (true : Hom top Ω)
    → (name : ∀ {A} (m : Subobject A) → Hom A Ω)
    → (∀ {A} (m : Subobject A) → m ≅ₘ name m ^* point→subobject true)
    → (∀ {A} (h : Hom A Ω) → name (h ^* point→subobject true) ≡ h)
    → Subobject-classifier C
  from-classification tru nm invl invr = done where
    done : Subobject-classifier C
    done .Ω = _
    done .true = point→subobject tru
    done .generic .name = nm
    done .generic .classifies m = iso→is-pullback-along {m = m} {n = point→subobject tru} (invl m)Note that the uniqueness part of the universal property is satisfied
by the last constraint on
the is-pullback-along assumption
amounts to saying that
by univalence of
so we have
    done .generic .unique {m = m} {h'} p =
      let
        rem₁ : m ≡ h' ^* point→subobject tru
        rem₁ = Sub-is-category cat .to-path $
          is-pullback-along→iso {m = m} {n = point→subobject tru} p
      in sym (ap nm rem₁ ∙ invr _)  record make-subobject-classifier : Type (o ⊔ ℓ) where
    field
      {Ω}  : Ob
      true : Hom top Ω
      name : ∀ {A} (m : Subobject A) → Hom A Ω
      named-name : ∀ {A} (m : Subobject A) → m ≅ₘ name m ^* point→subobject true
      name-named : ∀ {A} (h : Hom A Ω) → name (h ^* point→subobject true) ≡ h
module _ where
  open make-subobject-classifier hiding (Ω)
  to-subobject-classifier : ∀ {o ℓ} {C : Precategory o ℓ} (fc : Finitely-complete C) (cat : is-category C) → make-subobject-classifier fc cat → Subobject-classifier C
  to-subobject-classifier fc cat mk = from-classification fc cat (mk .true) (mk .name) (mk .named-name) (mk .name-named)
Sets-subobject-classifier : ∀ {ℓ} → Subobject-classifier (Sets ℓ)
Sets-subobject-classifier {ℓ} =
  to-subobject-classifier Sets-finitely-complete Sets-is-category mk where
  open Subr (Sets-pullbacks {ℓ})
  open Cat.Prelude using (Ω)
  open make-subobject-classifier hiding (Ω)The prototypical category with a subobject classifier is the category of sets. Since it is finitely complete and univalent, our helper above will let us swiftly dispatch the construction. Our subobject classifier is given by the type of propositions, lifted to the right universe. The name of a subobject is the family of propositions Note that we must squash this fibre down so it fits in but this squashing is benign because is a monomorphism (hence, an embedding).
  unbox : ∀ {A : Set ℓ} (m : Subobject A) {x} → □ (fibre (m .map) x) → fibre (m .map) x
  unbox m = □-out (monic→is-embedding (hlevel 2) (λ {C} g h p → m .monic {C} g h p) _)
  mk : make-subobject-classifier _ _
  mk .make-subobject-classifier.Ω = el! (Lift _ Ω)
  mk .true _ = lift ⊤Ω
  mk .name m x = lift (elΩ (fibre (m .map) x))Showing that this name function
actually works takes a bit of fiddling, but it’s nothing outrageous.
  mk .named-name m = Sub-antisym
    record
      { map = λ w → m .map w , lift tt , ap lift (to-is-true (inc (w , refl)))
      ; com = refl
      }
    record { com = ext λ x _ p → sym (unbox m (from-is-true p) .snd )}
  mk .name-named h = ext λ a → Ω-ua
    (rec! λ x _ p y=a → from-is-true (ap h (sym y=a) ∙ p))
    (λ ha → inc ((a , lift tt , ap lift (to-is-true ha)) , refl))As generic objects🔗
We can connect the definition above with that of generic objects in a fibration: specifically, a generic subobject is a generic object in the fibration of subobjects. The similar naming is not an accident: the proof boils down to shuffling data around.
module _ {o ℓ} (C : Precategory o ℓ)  where
  open is-generic-subobject
  open is-pullback-along
  open Generic-object
  open Cat.Reasoning C using (pulll)
  open is-cartesian
  open Subobjs C  from-generic-subobject
    : ∀ {Ω} {tru : Subobject Ω} → is-generic-subobject C tru
    → Generic-object Subobjects tru
  from-generic-subobject gen .classify    = gen .name
  from-generic-subobject gen .classify' s = record { com = gen .classifies s .square }
  from-generic-subobject gen .classify-cartesian s = record
    { universal = λ {u} {u'} m a → record
      { map = gen .classifies s .universal (pulll refl ∙ a .com)
      ; com = sym (gen .classifies s .p₁∘universal)
      }
    ; commutes  = λ m h → prop!
    ; unique    = λ m h → prop!
    }open Subobject-classifier using (unique)
module props {o ℓ} {C : Precategory o ℓ} (pb : has-pullbacks C) (so : Subobject-classifier C) where
  open Subobject-classifier so renaming (Ω to Ω')
  open is-pullback-along
  open Cat.Reasoning C
  open is-invertible
  open is-pullback
  open Pullback
  open Subr pb
  named : ∀ {A} (f : Hom A Ω') → Subobject A
  named f = f ^* true
  named-name : ∀ {A} {m : Subobject A} → m ≅ₘ named (name m)
  named-name = is-pullback-along→iso (classifies _)
  name-named : ∀ {A} {f : Hom A Ω'} → name (named f) ≡ f
  name-named {f = f} = sym $ so .unique record
    { top = pb f (Subobjs.map true) .p₂
    ; has-is-pb = has-is-pb (pb f (true .map))
    }
  name-injective : ∀ {A} {m n : Subobject A} → name m ≡ name n → m ≅ₘ n
  name-injective {m = m} {n} p =
    m              Sub.≅⟨ named-name ⟩
    named (name m) Sub.≅⟨ path→iso (ap named p) ⟩
    named (name n) Sub.≅˘⟨ named-name ⟩
    n              Sub.≅∎
  name-ap : ∀ {A} {m n : Subobject A} → m ≅ₘ n → name m ≡ name n
  name-ap {m = m} im = so .unique record
    { top       = classifies m .top ∘ im .Sub.from .map
    ; has-is-pb = subst-is-pullback (sym (im .Sub.from .com) ∙ eliml refl) refl refl refl
        (is-pullback-iso (≅ₘ→iso im) (classifies m .has-is-pb))
    }
  named-injective : ∀ {A} {f g : Hom A Ω'} → named f ≅ₘ named g → f ≡ g
  named-injective {f = f} {g = g} p =
    f              ≡˘⟨ name-named ⟩
    name (named f) ≡⟨ name-ap p ⟩
    name (named g) ≡⟨ name-named ⟩
    g ∎
  Ω-unique₂
    : ∀ {A} {f g : Hom A Ω'}
    → is-pullback-along C (named g .map) f (true .map)
    → f ≡ g
  Ω-unique₂ {f = f} {g = g} pb = so .unique pb ∙ name-named
  is-total : ∀ {A} (f : Hom A Ω') → Type _
  is-total f = is-invertible (pb f (true .map) .p₁)
  factors→is-total : ∀ {A} {f : Hom A Ω'} → Factors f (true .map) → is-total f
  factors→is-total (h , p) .inv = pb _ _ .universal (idr _ ∙ p)
  factors→is-total (h , p) .inverses = record
    { invl = pb _ _ .p₁∘universal
    ; invr = Pullback.unique₂ (pb _ _) {p = pushl p}
      (pulll (pb _ _ .p₁∘universal) ∙ idl _)
      (pulll (pb _ _ .p₂∘universal))
      (idr _)
      (idr _ ∙ true .monic _ _ (sym (pulll (sym p) ∙ pb _ (true .map) .square)))
    }
  is-total→factors : ∀ {A} {f : Hom A Ω'} → is-total f → Factors f (true .map)
  is-total→factors {f = f} inv = record { snd = done } where
    rem₁ : is-pullback-along C id f (true .map)
    rem₁ = record { has-is-pb = record
      { square       = idr _ ∙ sym (pulll (sym (pb _ _ .square)) ∙ cancelr (inv .is-invertible.invl))
      ; universal    = λ {P'} {p₁'} _ → p₁'
      ; p₁∘universal = idl _
      ; p₂∘universal = λ {P'} {p₁'} {p₂'} {α} → true .monic _ _ $
          pulll (pulll (sym (pb _ _ .square)) ∙ cancelr (inv .is-invertible.invl))
        ∙ α
      ; unique       = λ p _ → introl refl ∙ p
      }}
    done =
      f                              ≡⟨ so .unique rem₁ ⟩
      name ⊤ₘ                        ≡⟨ intror refl ⟩
      name ⊤ₘ ∘ id                   ≡⟨ classifies ⊤ₘ .square ⟩
      true .map ∘ classifies ⊤ₘ .top ∎
  true-domain-is-terminal : is-terminal C (true .dom)
  true-domain-is-terminal X .centre  = classifies ⊤ₘ .top
  true-domain-is-terminal X .paths h = true .monic _ _ (sym (is-total→factors record
    { inv      = pb _ _ .universal (pullr refl)
    ; inverses = record
      { invl = pb _ _ .p₁∘universal
      ; invr = Pullback.unique₂ (pb _ _) {p = pullr refl}
        (pulll (pb _ _ .p₁∘universal)) (extendl (pb _ _ .p₂∘universal)) id-comm
        (true .monic _ _ (extendl (sym (pb _ _ .square)) ∙ pullr (ap (h ∘_) id-comm)))
      }
    } .snd))References
- Johnstone, Peter T. 2002. Sketches of an Elephant: a Topos Theory Compendium. Oxford Logic Guides. New York, NY: Oxford Univ. Press. https://cds.cern.ch/record/592033.