open import Algebra.Prelude

open import Cat.Instances.Sets.Cocomplete
open import Cat.Instances.Functor.Limits
open import Cat.Instances.Slice.Presheaf
open import Cat.Instances.Sets.Complete
open import Cat.Diagram.Everything
open import Cat.Functor.Everything
open import Cat.Instances.Elements
open import Cat.Instances.Slice
open import Cat.Instances.Lift

import Cat.Functor.Bifunctor as Bifunctor
import Cat.Reasoning

module Topoi.Base where

Grothendieck topoi🔗

Topoi are an abstraction introduced by Alexander Grothendieck in the 1960s as a generalisation of topological spaces, suitable for his work in algebraic geometry. Later (in the work of William Lawvere, and even later Myles Tierney), topoi found a new life as “categories with a nice internal logic”, which mirrors that of the category Sets\sets. Perhaps surprisingly, every Grothendieck topos is also a topos in their logical conception (called elementary); Since elementary topoi are very hard to come by predicatively, we formalise a particular incarnation of Grothendieck’s notion here.

Grothendieck described his topoi by first defining a notion of site, which generalises the (poset of) open subsets of a topological space, and then describing what “sheaves on a site” should be, the corresponding generalisation of sheaves on a space. Instead of directly defining the notion of site, we will leave it implicitly, and define a topos as follows:

A topos T\ca{T} is a full subcategory of a presheaf category [Cop,Sets][\ca{C}\op,\sets] (the category C\ca{C} is part of the definition) such that the inclusion functor ι:T[Cop,Sets]\iota : \ca{T} \mono [\ca{C}\op,\sets] admits a left adjoint, and this left adjoint preserves finite limits. We summarise this situation in the diagram below, where “lex” (standing for “left exact”) is old-timey speak for “finite limit preserving”.

In type theory, we must take care about the specific universes in which everything lives. Now, much of Grothendieck topos theory generalises to arbitrary “base” topoi, via the use of bounded geometric morphisms, but the “main” definition talks about Sets\sets-topoi. In particular, every universe κ\kappa generates a theory of Setsκ\sets_\kappa-topoi, the categories of κ\kappa-small sheaves on κ\kappa-small sites.

Fix a universe level κ\kappa, and consider the category Setsκ\sets_\kappa: A topos T\ca{T} might be a large category (i.e. it might have a space of objects oo with o>κo > \kappa), but it is essentially locally small, since it admits a fully-faithful functor into [Cop,Setsκ][\ca{C}\op,\sets_\kappa], which have homs at level κ\kappa. Hence, even if we allowed the category T\ca{T} to have hom\homs at a level \ell, we would have no more information there than fits in κ\kappa.

For this reason, a topos C\ca{C} only has two levels: The level oo of its objects, and the level κ\kappa of the category of Sets relative to which C\ca{C} is a topos.

record Topos {o} κ (𝓣 : Precategory o κ) : Type (lsuc (o  κ)) where
  field
    site : Precategory κ κ

    ι : Functor 𝓣 (PSh κ site)
    has-ff : is-fully-faithful ι

    L : Functor (PSh κ site) 𝓣
    L-lex : is-lex L

    L⊣ι : L  ι

  module ι = Functor ι
  module L = Functor L
  module L⊣ι = _⊣_ L⊣ι

As generalised spaces🔗

I’ll echo here the standard definition of topological space, but due to personal failings I can’t motivate it. A topological space (X,τ)(X, \tau) consists of a set of points XX, and a topology τ\tau, a subset of its powerset which is closed under arbitrary unions and finite intersections.

Let’s reword that using category-theoretic language: Recall that the powerset of XX is the poset [X,Props][X,\props]. It is — being a functor category into a complete and cocomplete (thin) category — also complete and cocomplete, so all joins and finite meets (unions and intersections) exist; Furthermore, the distributive law

S(iFi)=i(SFi) S \cap \left(\bigcup_i F_i\right) = \bigcup_i (S \cap F_i)

holds for any subset SS and family of subsets FF. A lattice satisfying these properties (finite meets, arbitrary joins, the distributive law) is called a frame. A map of frames f:ABf : A \to B is defined to be a map of their underlying sets preserving finite meets and arbitrary joins — by abstract nonsense, this is the same as a function f:BAf^* : B \to A which preserves finite meets and has a right adjoint f:ABf_* : A \to B1.

We can prove that a topology τ\tau on XX is the same thing as a subframe of its powerset [X,Props][X,\props] — a collection of subsets of XX, closed under finite meets and arbitrary joins.

Now, the notion of “topos” as a generalisation of that of “topological space” is essentially self-evident: A topos T\ca{T} is a sub-topos of a presheaf category [Cop,Sets][\ca{C}\op,\sets]. We have essentially categorified “subframe” into “subtopos”, and Props\props into Sets\sets. Note that, while this seems circular (“a topos is a sub-topos of…”), the notion of subtopos — or rather, of geometric embedding — makes no mention of actual topoi, and makes sense for any pair of categories.

As categories of spaces🔗

Another perspective on topoi is that they are categories of spaces, rather than spaces themselves. We start by looking at presheaf topoi, [Cop,Sets][\ca{C}^op,\sets]. The coyoneda lemma tells us that every presheaf is a colimit of representables, which can be restated in less abstract terms by saying that presheaves are instructions for gluing together objects of C\ca{C}. The objects of C\ca{C} are then interpreted as “primitive shapes”, and the maps in C\ca{C} are interpreted as “maps to glue against”.

Let’s make this more concrete by considering an example: Take C=\ca{C} = \bull \rightrightarrows \bull, the category with two points — let’s call them VV and EE — and two arrows s,t:VEs, t : V \to E. A presheaf FF on this category is given by a set F0(V)F_0(V), a set F0(E)F_0(E), and two functions F1(s),F1(t):F0(E)F0(V)F_1(s), F_1(t) : F_0(E) \to F_0(V). We call F0(V)F_0(V) the vertex set, and F0(E)F_0(E) the edge set. Indeed, a presheaf on this category is a directed multigraph, and maps between presheaves preserve adjacency.

Our statement about “gluing primitive shapes” now becomes the rather pedestrian statement “graphs are made up of vertices and edges”. For instance, the graph \bull \to \bull \to \bull can be considered as a disjoint union of two edges, which is then glued together in a way such that the target of the first is the source of the second. The maps s,t:VEs, t : V \to E exhibit the two ways that a vertex can be considered “part of” an edge.

As “logically nice” categories🔗

The definition of topos implies that any topos T\ca{T} enjoys many of the same categorical properties of the category Sets\sets (see below). Topoi are complete and cocomplete, cartesian closed (even locally so) — colimits are stable under pullback, coproducts are disjoint, and equivalence relations are effective.

These properties, but especially local cartesian closure, imply that the internal logic of a topos is an incredibly nice form of type theory. Dependent sums and products exist, there is an object of natural numbers, the poset of subobjects is a complete lattice (even a Heyting algebra), including existential and universal quantification. Effectivity of congruences means that two points in a quotient are identified if, and only if, they are related by the congruence.

In fact, this is essentially the definition of a topos. The actual definition, as a lex reflective subcategory of a presheaf category, essentially ensures that the category T\ca{T} inherits the nice logical properties of Sets\sets (through the presheaf category, which is also very nice logically).

Terminology: As was implicitly mentioned above, for a topos L:TPSh(C)L : \ca{T} \xadj{}{} \psh(\ca{C}), we call the category C\ca{C} the site of definition. Objects in T\ca{T} are called sheaves, and the functor LL is called sheafification. Maps between topoi are called geometric morphisms, and will be defined below. We denote the 2-category of topoi, geometric morphisms and natural transformations by Top\topos, following Johnstone. When PSh(C)\psh(\ca{C}) is regarded as a topos unto itself, rather than an indirection in the definition of a sheaf topos, we call it the topos of C\ca{C}-sets.

Examples🔗

The “trivial” example of topoi is the category Sets\sets, which is equivalently the category [op,Sets][*\op,\sets] of presheaves on the terminal category. This is, in fact, the terminal object in the 2-category Top\topos of topoi (morphisms are described below), so we denote it by 𝟙.

𝟙 :  {κ}  Topos κ (Sets κ)
𝟙 {κ} = sets where
  open Topos
  open Functor
  open _⊣_
  open is-lex

The inclusion functor SetsPSh()\sets \mono \psh(*) is given by the “constant presheaf” functor, which sends each set XX to the constant functor with value XX.

  incl : Functor (Sets κ) (PSh κ (Lift-cat κ κ ⊤Cat))
  incl .F₀ x .F₀ _    = x
  incl .F₀ x .F₁ _ f  = f
  incl .F₀ x .F-id    = refl
  incl .F₀ x .F-∘ f g = refl

  incl .F₁ f = NT  _  f)  _ _ _  refl)

  incl .F-id    = Nat-path λ _  refl
  incl .F-∘ f g = Nat-path λ _  refl

  sets : Topos κ (Sets κ)
  sets .site = Lift-cat κ κ ⊤Cat

  sets .ι = incl

This functor is fully faithful since a natural transformation between presheaves on the point is entirely determined by its component at.. well, the point. Hence, the map ηη\eta \mapsto \eta_* exhibits an inverse to the inclusion functor’s action on morphisms, which is sufficient (and necessary) to conclude fully-faithfulness.

  sets .has-ff {x} {y} = is-iso→is-equiv isic where
    open is-iso
    isic : is-iso (incl .F₁ {x} {y})
    isic .inv nt         = η nt _

    isic .rinv nt i .η _ = η nt _
    isic .rinv nt i .is-natural _ _ f j x =
      y .is-tr _ _  j  nt .η _ x)  j  nt .is-natural _ _ f j x) i j

    isic .linv x i = x

The “sheafification” left adjoint is given by evaluating a presheaf FF at its sole section F()F(*), and similarly by evaluating a natural transformation η:FG\eta : F \To G at its one component η:F()G()\eta_* : F(*) \to G(*).

  sets .L .F₀ F    = F₀ F _
  sets .L .F₁ nt   = η nt _
  sets .L .F-id    = refl
  sets .L .F-∘ f g = refl
These functors actually define an adjoint equivalence of categories, LL is continuous and, in particular, lex. Rather than appealing to that theorem, though, we define the preservation of finite limits directly for efficiency concerns.
  sets .L-lex .pres-⊤ {T} psh-terminal set = contr (cent .η _) uniq where
    func = incl .F₀ set
    cent = psh-terminal func .centre
    uniq :  f  cent .η _  f
    uniq f = ap  e  e .η _) (psh-terminal func .paths f′) where
      f′ : _ => _
      f′ .η _ = f
      f′ .is-natural _ _ _ = funext λ x  happly (sym (F-id T)) _

  sets .L-lex .pres-pullback {P} {X} {Y} {Z} pb = pb′ where
    open is-pullback
    pb′ : is-pullback (Sets κ) _ _ _ _
    pb′ .square = ap  e  η e _) (pb .square)
    pb′ .limiting {P'} {p₁' = p₁'} {p₂' = p₂'} p =
      η (pb .limiting {P′ = incl .F₀ P'} {p₁' = p1'} {p₂' = p2'}
          (Nat-path λ _  p)) _
      where
        p1' : _ => _
        p1' .η _ = p₁'
        p1' .is-natural x y f i o = F-id X (~ i) (p₁' o)
        p2' : _ => _
        p2' .η _ = p₂'
        p2' .is-natural x y f i o = F-id Y (~ i) (p₂' o)
    pb′ .p₁∘limiting = ap  e  η e _) (pb .p₁∘limiting)
    pb′ .p₂∘limiting = ap  e  η e _) (pb .p₂∘limiting)
    pb′ .unique {P′} {lim' = lim'} p1 p2 =
      ap  e  η e _) (pb .unique {lim' = l′}
        (Nat-path λ _  p1) (Nat-path λ _  p2))
      where
        l′ : incl .F₀ P′ => P
        l′ .η _ = lim'
        l′ .is-natural x y f i o = F-id P (~ i) (lim' o)

Similarly, we can construct the adjunction unit and counit by looking at components and constructing identity natural transformations.

  sets .L⊣ι .unit .η _ .η _ f            = f
  sets .L⊣ι .unit .η F .is-natural _ _ _ = F-id F
  sets .L⊣ι .unit .is-natural _ _ _      = Nat-path λ _  refl

  sets .L⊣ι .counit .η _ x            = x
  sets .L⊣ι .counit .is-natural _ _ _ = refl

  sets .L⊣ι .zig = refl
  sets .L⊣ι .zag = Nat-path λ _  refl

More canonical examples are given by any presheaf category, where both the inclusion and sheafification functors are the identity. Examples of presheaf topoi are abundant in abstract homotopy theory (the categories of cubical and simplicial sets) — which also play an important role in modelling homotopy type theory; Another example of a presheaf topos is the category of quivers (directed multigraphs).

Presheaf :  {κ} (C : Precategory κ κ)  Topos κ (PSh κ C)
Presheaf {κ} C = psh where
  open Functor
  open Topos
  psh : Topos _ _
  psh .site = C
  psh .ι = Id
  psh .has-ff = id-equiv
  psh .L = Id
  psh .L-lex .is-lex.pres-⊤ c = c
  psh .L-lex .is-lex.pres-pullback pb = pb
  psh .L⊣ι = adj where
    open _⊣_
    adj : Id  Id
    adj .unit = NT  _  idnt) λ x y f  Nat-path λ _  refl
    adj .counit = NT  _  idnt)  x y f  Nat-path λ _  refl)
    adj .zig = Nat-path λ _  refl
    adj .zag = Nat-path λ _  refl

Properties of topoi🔗

The defining property of a topos T\ca{T} is that it admits a geometric embedding into a presheaf category, meaning the adjunction L:TPSh(C):ιL : \ca{T} \xadj{}{} \psh(\ca{C}) : \iota is very special indeed: Since the right adjoint is fully faithful, the adjunction is monadic, meaning that it exhibits T\ca{T} as the category of algebras for a (lex, idempotent) monad: the “sheafification monad”. This gives us completeness in T\ca{T} for “free” (really, it’s because presheaf categories are complete, and those are complete because Sets\sets is.)

module _ {o κ} {𝓣 : Precategory o κ} (T : Topos κ 𝓣) where
  open Topos T

  Sheafify : Monad (PSh κ site)
  Sheafify = Adjunction→Monad L⊣ι

  Sheafify-monadic : is-monadic L⊣ι
  Sheafify-monadic = is-reflective→is-monadic L⊣ι has-ff

  Topos-is-complete : is-complete κ κ 𝓣
  Topos-is-complete = equivalence→complete
    (is-equivalence.inverse-equivalence Sheafify-monadic)
    (Eilenberg-Moore-is-complete
      (Functor-cat-is-complete (Sets-is-complete {ι = κ} {κ} {κ})))

Furthermore, since LL preserves colimits (it is a left adjoint), then we can compute the colimit of some diagram F:JTF : J \to \ca{T} as the colimit (in PSh(C)\psh(\ca{C})) of ιF\iota F — which exists because Sets\sets is cocomplete — then apply LL to get a colimiting cocone for LιFL \iota F. But the counit of the adjunction ε:LιId\eps : L \iota \To \id{Id} is a natural isomorphism, so we have a colimiting cocone for FF.

  Topos-is-cocomplete : is-cocomplete κ κ 𝓣
  Topos-is-cocomplete F =
    Colimit-ap-iso _
      (F∘-iso-id-l (is-reflective→counit-iso L⊣ι has-ff))
      sheafified
      where
      psh-colim : Colimit (ι F∘ F)
      psh-colim = Functor-cat-is-cocomplete (Sets-is-cocomplete {ι = κ} {κ} {κ}) _

      sheafified : Colimit ((L F∘ ι) F∘ F)
      sheafified = subst Colimit F∘-assoc $
        left-adjoint-colimit L⊣ι psh-colim

Since the reflector is left exact, and thus in particular preserves finite products, a theorem of Johnstone (Elephant A4.3.1) implies the topos T\ca{T} is an exponential ideal in PSh(C)\psh(\ca{C}): If YY is a sheaf, and XX is any presheaf, then the internal hom [X,Y][X,Y] is a sheaf: topoi are cartesian closed.

Since topoi are defined as embedding faithfully into a category of presheaves, it follows that they have a small generating set, in the same sense that representables generate presheaves: In fact, the generating set of a Grothendieck topos is exactly the set of representables of its site!

Elaborating, letting T\ca{T} be a topos, two maps f,g:XYTf, g : X \to Y \in \ca{T} are equal if and only if they are equal under ι\iota, i.e. as maps of presheaves. But it follows from the coyoneda lemma that maps of presheaves are equal if and only if they are equal on all representables. Consequently, two maps in a T\ca{T} are equal if and only if they agree on all generalised elements with domain the sheafification of a representable:

  Representables-generate
    :  {X Y} {f g : C.Hom X Y}
     (  {A} (h : C.Hom (L.₀ (よ₀ ct.site A)) X)
       f C.∘ h  g C.∘ h )
     f  g
  Representables-generate {f = f} {g} sep =
    fully-faithful→faithful {F = ct.ι} ct.has-ff $
      Representables-generate-presheaf ct.site λ h 
        ι.₁ f PSh.∘ h                                    ≡⟨ mangle 
        ι.₁ (f C.∘ counit.ε _ C.∘ L.₁ h) PSh.∘ unit.η _  ≡⟨ ap (PSh._∘ _) (ap ι.₁ (sep _)) 
        ι.₁ (g C.∘ counit.ε _ C.∘ L.₁ h) PSh.∘ unit.η _  ≡˘⟨ mangle ≡˘
        ι.₁ g PSh.∘ h                                    

Finally, the property of “being a topos” is invariant under slicing. More specifically, since a given topos can be presented by different sites, a presentation TPSh(C)\ca{T} \xadj{}{} \psh(\ca{C}) can be sliced under an object XTX \in \ca{T} to present T/X\ca{T}/X as a sheaf topos, with site of definition PSh(ι(X))\psh(\int \iota(X)). This follows from a wealth of pre-existing theorems:

  • A pair LRL \dashv R of adjoint functors can be sliced under an object XX, giving an adjunction ΣεL/R(X)R/X\Sigma \epsilon L/R(X) \dashv R/X; Slicing a functor preserves full-faithfulness and left exactness, hence slicing a geometric embedding gives a geometric embedding.
  • The category PSh(C)/ι(X)\psh(\ca{C})/\iota(X) is equivalent to the category PSh(ι(X))\psh(\int \iota(X)), hence “being a presheaf topos is stable under slicing”. Furthermore, this equivalence is part of an ambidextrous adjunction, hence both functors preserve limits.
  • Dependent sum Σf\Sigma f along an isomorphism is an equivalence of categories; But since a topos is presented by a reflective subcategory inclusion, the counit is an isomorphism.

We build the geometric embedding presenting T/X\ca{T}/X as a topos by composing the adjunctions ε!(L/ι(X))ι/X\epsilon_!(L/\iota(X)) \dashv \iota/X and FF1F \dashv F^{-1} — where FF is the equivalence PSh(C)/XPSh(X)\psh(\ca{C})/X \to \psh(\int X). The right adjoint is fully faithful because it composes two fully faithful functors (a slice of ι\iota and an equivalence), the left adjoint preserves finite limits because it is a composite of two equivalences (hence two right adjoints) and a lex functor.

  Slice-topos : Topos  (Slice C X)
  Slice-topos .site =  T.site (T.ι.F₀ X)
  Slice-topos .ι = slice→total F∘ Sliced (T.ι) X
  Slice-topos .has-ff = ∙-is-equiv (Sliced-ff {F = T.ι} (T.has-ff)) slice→total-is-ff
  Slice-topos .L = (Σf (T .L⊣ι.counit.ε _) F∘ Sliced (T.L) (T.ι.F₀ X)) F∘ total→slice

  Slice-topos .L-lex = F∘-is-lex
    (F∘-is-lex
      (right-adjoint→lex
        (is-equivalence.F⁻¹⊣F
          (Σ-iso-equiv (C.iso→invertible
            (is-reflective→counit-is-iso T.L⊣ι T.has-ff)))))
      (Sliced-lex T.L-lex))
    (right-adjoint→lex (slice→total-is-equiv .is-equivalence.F⊣F⁻¹))

  Slice-topos .L⊣ι = LF⊣GR (is-equivalence.F⁻¹⊣F slice→total-is-equiv)
                           (Sliced-adjoints T.L⊣ι)

Geometric morphisms🔗

The definition of a topos as a generalisation of topological space leads us to look for a categorification of “continuous map” to functors between topoi. In the same way that a continuous function f:XYf : X \to Y may be seen as a homomorphism of frames f:O(Y)O(X)f^* : O(Y) \to O(X), with defining feature the preservation of finite meets and arbitrary joins, we shall define a geometric morphism Top(X,Y)\topos(X,Y) to be a functor f:YXf^* : Y \to X which is left exact and admits a right adjoint.

Why the right adjoint? Well, right adjoints preserve limits, but by duality, left adjoints preserve colimits. This embodies the “arbitrary joins” part of a map of frames, whereas the “finite meets” come from left exactness.

record Geom[_,_] (E : Precategory o ) (F : Precategory o′ ℓ′) : Type (lvl E F) where
  no-eta-equality
  field
    Inv[_]  : Functor F E
    Dir[_]  : Functor E F
    Inv-lex : is-lex Inv[_]
    Inv⊣Dir : Inv[_]  Dir[_]

open Geom[_,_] public

Computation establishes that the identity functor is left exact, and self adjoint, so it is in particular both the direct and inverse image parts of a geometric morphism TT\ca{T} \to \ca{T}.

Idg : Geom[ E , E ]
Idg {E = E} = record { Inv[_] = Id ; Dir[_] = Id
                     ; Inv-lex = lex ; Inv⊣Dir = adj }

Since adjunctions compose, geometric morphisms do, too. Observe that the composite of inverse images and the composite of direct images go in different directions! Fortunately, this matches the convention for composing adjunctions, where the functors “swap sides”: LFGRLF \dashv GR.

_G∘_ : Geom[ F , G ]  Geom[ E , F ]  Geom[ E , G ]
f G∘ g = mk where
  module f = Geom[_,_] f
  module g = Geom[_,_] g
  open is-lex

  mk : Geom[ _ , _ ]
  Inv[ mk ] = Inv[ g ] F∘ Inv[ f ]
  Dir[ mk ] = Dir[ f ] F∘ Dir[ g ]
  mk .Inv⊣Dir = LF⊣GR f.Inv⊣Dir g.Inv⊣Dir

The composition of two left-exact functors is again left-exact, so there’s no impediment to composition there, either.

  mk .Inv-lex .pres-⊤ term ob = g.Inv-lex .pres-⊤ (f.Inv-lex .pres-⊤ term) _
  mk .Inv-lex .pres-pullback pb = g.Inv-lex .pres-pullback (f.Inv-lex .pres-pullback pb)

An important class of geometric morphism is the geometric embedding, which we’ve already met as the very definition of Topos. These are the geometric morphisms with fully faithful direct image. These are again closed under composition, so if we want to exhibit that a certain category C\ca{C} is a topos, it suffices to give a geometric embedding CT\ca{C} \to \ca{T}, where T\ca{T} is some topos which is convenient for this application.

record Geom[_↪_] (E : Precategory o ) (F : Precategory o′ ℓ′) : Type (lvl E F) where
  field
    morphism : Geom[ E , F ]
    has-ff : is-fully-faithful Dir[ morphism ]

Geometric-embeddings-compose : Geom[ F  G ]  Geom[ E  F ]  Geom[ E  G ]
Geometric-embeddings-compose f g =
  record { morphism = f .morphism G∘ g .morphism
         ; has-ff = ∙-is-equiv (g .has-ff) (f .has-ff) }
  where open Geom[_↪_]

Topos→geometric-embedding : (T : Topos κ E)  Geom[ E  PSh κ (T .Topos.site) ]
Topos→geometric-embedding T = emb where
  emb : Geom[ _  _ ]
  Inv[ emb .Geom[_↪_].morphism ] = T .Topos.L
  Dir[ emb .Geom[_↪_].morphism ] = T .Topos.ι
  emb .Geom[_↪_].morphism .Inv-lex = T .Topos.L-lex
  emb .Geom[_↪_].morphism .Inv⊣Dir = T .Topos.L⊣ι
  emb .Geom[_↪_].has-ff = T .Topos.has-ff

  1. By the adjoint functor theorem, any map between posets which preserves arbitrary joins has a right adjoint; Conversely, every map which has a right adjoint preserves arbitrary joins.↩︎