open import Cat.Diagram.Limit.Finite open import Cat.Diagram.Limit.Base open import Cat.Diagram.Terminal open import Cat.Prelude module Cat.Instances.Sets.Complete where
Sets is complete🔗
We prove that the category of -sets is -complete for any universe levels and . Inverting this to speak of maxima rather than ordering, to admit all -limits, we must be in at least the category of -sets, but any extra adjustment is also acceptable. So, suppose we have an indexing category and a diagram ; Let’s build a limit for it!
Sets-is-complete : ∀ {ι κ o} → is-complete ι κ (Sets (ι ⊔ κ ⊔ o)) Sets-is-complete {D = D} F = lim where module D = Precategory D module F = Functor F comm-prop : ∀ f → is-prop (∀ x y (g : D.Hom x y) → F.₁ g (f x) ≡ (f y)) comm-prop f = Π-is-hlevel 1 λ _ → Π-is-hlevel 1 λ _ → Π-is-hlevel 1 λ _ → F.₀ _ .is-tr _ _
Since Set is closed under (arbitrary) products
, we can build the limit of an arbitrary diagram — which we will write — by first taking the product (which is a set of dependent functions), then restricting ourselves to the subset of those for which , i.e., those which are cones over .
f-apex : Set _ f-apex .∣_∣ = Σ[ f ∈ ((j : D.Ob) → ∣ F.₀ j ∣) ] (∀ x y (g : D.Hom x y) → F.₁ g (f x) ≡ (f y)) f-apex .is-tr = Σ-is-hlevel 2 (Π-is-hlevel 2 (λ x → F.₀ x .is-tr)) (λ f → is-prop→is-set (comm-prop f))
To form a cone, given an object , and an inhabitant of the type underlying f-apex, we must cough up (for ψ) an object of ; But this is exactly what gives us. Similarly, since witnesses that commutes, we can project it directly.
open Cone cone : Cone F cone .Cone.apex = f-apex cone .ψ x = λ { (f , p) → f x } cone .commutes o = funext λ { (_ , p) → p _ _ o }
Given some other cone , to build a cone homomorphism , recall that comes equipped with its own function , which we can simply flip around to get a function ; This function is in the subset carved out by since is a cone, hence , as required.
open Terminal lim : Limit F lim .top = cone lim .has⊤ K = contr map map-unique where module K = Cone K open Cone-hom map : Cone-hom F K cone map .hom x = (λ j → K.ψ j x) , λ x y f → happly (K.commutes f) _ map .commutes _ = refl map-unique : ∀ m → map ≡ m map-unique m = Cone-hom-path _ (funext λ x → Σ-prop-path comm-prop (funext λ y i → m .commutes y (~ i) x))
Finite set-limits🔗
For expository reasons, we present the computation of the most famous shapes of finite limit (terminal objects, products, pullbacks, and equalisers) in the category of sets. All the definitions below are redundant, since finite limits are always small, and thus the category of sets of any level admits them.
Sets-terminal : Terminal (Sets ℓ) Sets-terminal .top = Lift _ ⊤ , is-prop→is-set (λ _ _ → refl) Sets-terminal .has⊤ _ = fun-is-hlevel 0 (contr (lift tt) λ x i → lift tt)
Products are given by product sets:
Sets-products : (A B : Set ℓ) → Product A B Sets-products A B .apex = (∣ A ∣ × ∣ B ∣) , ×-is-hlevel 2 (A .is-tr) (B .is-tr) Sets-products A B .π₁ = fst Sets-products A B .π₂ = snd Sets-products A B .has-is-product .⟨_,_⟩ f g x = f x , g x Sets-products A B .has-is-product .π₁∘factor = refl Sets-products A B .has-is-product .π₂∘factor = refl Sets-products A B .has-is-product .unique o p q i x = p i x , q i x
Equalisers are given by carving out the subset of where and agree using :
Sets-equalisers : (f g : Hom A B) → Equaliser {A = A} {B = B} f g Sets-equalisers {A = A} {B = B} f g = eq where eq : Equaliser f g eq .apex = Σ[ x ∈ ∣ A ∣ ] (f x ≡ g x) , Σ-is-hlevel 2 (A .is-tr) λ _ → is-prop→is-set (B .is-tr _ _) eq .equ = fst eq .has-is-eq .equal = funext snd eq .has-is-eq .limiting {e′ = e′} p x = e′ x , happly p x eq .has-is-eq .universal = refl eq .has-is-eq .unique {p = p} q = funext λ x → Σ-prop-path (λ _ → B .is-tr _ _) (happly (sym q) x)
Pullbacks are the same, but carving out a subset of .
Sets-pullbacks : ∀ {A B C} (f : Hom A C) (g : Hom B C) → Pullback {X = A} {Y = B} {Z = C} f g Sets-pullbacks {A = A} {B = B} {C = C} f g = pb where pb : Pullback f g pb .apex .∣_∣ = Σ[ x ∈ ∣ A ∣ ] Σ[ y ∈ ∣ B ∣ ] (f x ≡ g y) pb .apex .is-tr = Σ-is-hlevel 2 (A .is-tr) λ _ → Σ-is-hlevel 2 (B .is-tr) λ _ → is-prop→is-set (C .is-tr _ _) pb .p₁ (x , _ , _) = x pb .p₂ (_ , y , _) = y pb .has-is-pb .square = funext (snd ⊙ snd) pb .has-is-pb .limiting {p₁' = p₁'} {p₂'} p a = p₁' a , p₂' a , happly p a pb .has-is-pb .p₁∘limiting = refl pb .has-is-pb .p₂∘limiting = refl pb .has-is-pb .unique {p = p} {lim' = lim'} q r i x = q i x , r i x , λ j → is-set→squarep (λ i j → C .is-tr) (λ j → f (q j x)) (λ j → lim' x .snd .snd j) (happly p x) (λ j → g (r j x)) i j
Hence, Sets is finitely complete:
open Finitely-complete Sets-finitely-complete : Finitely-complete (Sets ℓ) Sets-finitely-complete .terminal = Sets-terminal Sets-finitely-complete .products = Sets-products Sets-finitely-complete .equalisers = Sets-equalisers Sets-finitely-complete .pullbacks = Sets-pullbacks