module Cat.Displayed.Instances.Gluing
{o ℓ o' ℓ'} {C : Precategory o ℓ} {D : Precategory o' ℓ'} {F : Functor D C}
(ccart : Cartesian-category C)
(dcart : Cartesian-category D)
(pres : Cartesian-functor F dcart ccart)
where
Artin gluing🔗
private
module F = Fr F
module C = Cartesian-category ccart
module D = Cartesian-category dcart
open Cartesian-functor pres renaming (module preserved to is)
open /-Obj public
open Displayed
Let and be Cartesian categories, and be a Cartesian functor between them. The (Artin) gluing along is the pullback of fundamental fibration along i.e. the left display map in the diagram
We write out the definition of gluing in components to get slightly better definitional behaviour. Note that, in some cases of interest, will not have pullbacks, meaning that neither or will be Cartesian fibrations.
: Displayed D (o ⊔ ℓ) ℓ
Gl = Change-of-base F (Slices C) Gl
module Gl = Displayed Gl
open is-terminal-over
open TerminalP
To be explicit, an object of over is a pair of an object and a morphism A map over is a map such that the square
commutes.
Displayed Cartesian structure🔗
We will now show that is a family of Cartesian categories over In particular, it has a total terminal object as soon as preserves the terminal object, and total products as soon as preserves products. The structure is “componentwise”, in that the domain of (e.g.) the total terminal object is the terminal object in while the morphisms are induced by comparison maps.
: TerminalP Gl D.terminal
Gl-terminal .top' = cut {dom = C.top} (pres-terminal _ .centre)
Gl-terminal .has⊤' .!' = record
Gl-terminal { map = C.!
; com = is-contr→is-prop (pres-terminal _) _ _
}
.has⊤' .!-unique' h = Slice-pathp (C.!-unique _) Gl-terminal
open is-product-over
open ProductP
The case for products is similar: the domain of is and the map into is induced by the preserved product structure on
: has-products-over Gl D.products
Gl-products {x} {y} a b = done module Gl-products where
Gl-products : Gl ʻ (x D.⊗₀ y)
apx .dom = a .dom C.⊗₀ b .dom
apx .map = is.⟨ a .map C.∘ C.π₁ , b .map C.∘ C.π₂ ⟩
apx
: ProductP Gl _ a b
done .apex' = apx
done .π₁' = record { com = sym is.π₁∘⟨⟩ }
done .π₂' = record { com = sym is.π₂∘⟨⟩ }
done .has-is-product' .⟨_,_⟩' {f = f} {a' = a'} {g = g} f' g' =
done record { com = coh }
where abstract
: apx .map C.∘ C.⟨ f' .map , g' .map ⟩ ≡ F.₁ (D.⟨ f , g ⟩) C.∘ a' .map
coh = C.pullr (C.⟨⟩∘ _) ∙ sym (is.unique
coh (F.pulll D.π₁∘⟨⟩ ∙ sym (f' .com) ∙ sym (C.pullr C.π₁∘⟨⟩))
(F.pulll D.π₂∘⟨⟩ ∙ sym (g' .com) ∙ sym (C.pullr C.π₂∘⟨⟩)))
.has-is-product' .π₁∘⟨⟩' = Slice-pathp C.π₁∘⟨⟩
done .has-is-product' .π₂∘⟨⟩' = Slice-pathp C.π₂∘⟨⟩
done .has-is-product' .unique' p q = Slice-pathp $ C.⟨⟩-unique
done (λ i → p i .map)
(λ i → q i .map)
open Gl-products renaming (apx to _×Gl_) using () public
open Cartesian-over
: Cartesian-over Gl dcart
Gl-cartesian .terminal' = Gl-terminal
Gl-cartesian .products' = Gl-products Gl-cartesian
Cartesian closed structure🔗
The construction in this section is adapted to the displayed setting from that given by Johnstone (2002, A2.1.12).
Now suppose that each of and are Cartesian closed, and that additionally admits pullbacks. With no further assumptions on the Artin gluing is made into a Cartesian closed category over in that it has total exponential objects.
Gl-closed: Cartesian-closed C ccart
→ (dcl : Cartesian-closed D dcart)
→ has-pullbacks C
→ Cartesian-closed-over Gl Gl-cartesian dcl
{A} {B} A' B' = done where Gl-closed ccl dcl pullbacks
module Cλ = Cartesian-closed ccl
module Dλ = Cartesian-closed dcl
open is-exponential-over
open ExponentialP
open /-Obj A' renaming (dom to X ; map to f)
open /-Obj B' renaming (dom to Y ; map to g)
To explain the construction of the exponential for and 1 let us pretend that is the syntactic category of some theory with function types, is a category of semantic objects, and the functor embeds syntactic objects into semantic ones. In this situation, we think of the maps in each object of as mapping semantic things to witnesses for their definability in terms of syntax. In general, the semantic exponential will contain many “exotic” operations which can not be syntactically defined, for example those defined by adversarial knowledge of the internal structures of and so we have no hope of writing a map
Instead, we observe that the commutativity condition for morphisms in described above can be internalised using a pullback in we want to restrict the exponential to those maps which fit into a commutative square with a map from To that end, we define the map as the exponential transpose of where the middle map is the inverse of product comparison map. We therefore think of it as turning syntactic functions to ones which take a semantic argument, by passing to the “definability witnesses” of the domain.
: C.Hom (F.₀ Dλ.[ A , B ]) Cλ.[ X , F.₀ B ]
ψ = Cλ.ƛ (F.₁ Dλ.ev C.∘ comparison.inv C.∘ (C.id C.⊗₁ f)) ψ
The map is much more simply defined as the functorial action of on It therefore takes semantic functions to ones producing a syntactic result, by passing to the “definability witnesses” of each result. The pullback
can thus be thought of as consisting of semantic functions paired with syntactic terms which “pointwise” witnesses their definability, in the sense that if witnesses the definability of some then the syntactic application is a definability witness of As indicated in the diagram, the projection onto is precisely the exponential in
: C.Hom Cλ.[ X , Y ] Cλ.[ X , F.₀ B ]
φ = Cλ.ƛ (g C.∘ Cλ.ev) φ
module pb = Pullback (pullbacks ψ φ)
: Gl ʻ Dλ.[ A , B ]
pow .dom = pb.apex
pow .map = pb.p₁ pow
The evaluation map lying over evaluation map, is defined in terms of the projection onto and evaluation map, explicitly as the composition
: Gl.Hom[ Dλ.ev ] (pow ×Gl A') B'
evm .map = Cλ.ev C.∘ pb.p₂ C.⊗₁ C.id evm
We must also calculate that this evaluation map is actually displayed
over
Finally, if we have some other
the exponential transpose of a map
over
has semantic component
and syntactic component the composition
which is forced on us by the definition of maps in
We omit the verification that this satisfies the pullback condition, and
that it is indeed an exponential transpose.
.com = Equiv.injective (_ , Cλ.lambda-is-equiv) $
evm .ƛ (g C.∘ Cλ.ev C.∘ (pb.p₂ C.⊗₁ C.id))
Cλ.ƛ (C.pulll refl) ⟩
≡⟨ ap Cλ.ƛ ((g C.∘ Cλ.ev) C.∘ (pb.p₂ C.⊗₁ C.id))
Cλ.ƛ-∘' _ _ _ ⟩
≡⟨ Cλ.ƛ (Cλ.ev C.∘ (C.id C.⊗₁ C.id)) C.∘ Cλ.ƛ (g C.∘ Cλ.ev) C.∘ pb.p₂
Cλ.eliml (ap Cλ.ƛ (C.elimr (C.×-functor .Functor.F-id)) ∙ Cλ.lambda-ev) ⟩
≡⟨ C.ƛ (g C.∘ Cλ.ev) C.∘ pb.p₂
Cλ.square ⟩
≡⟨ sym pb.ƛ (F.₁ Dλ.ev C.∘ comparison.inv C.∘ (C.id C.⊗₁ f)) C.∘ pb.p₁
Cλ.ƛ-∘-idr _ _ ⟩
≡˘⟨ Cλ.ƛ ((F.₁ Dλ.ev C.∘ comparison.inv C.∘ (C.id C.⊗₁ f)) C.∘ (pb.p₁ C.⊗₁ C.id))
Cλ.ƛ (C.pullr (C.pullr (Fr.collapse C.×-functor (C.idl _ ,ₚ C.idr _)))) ⟩
≡⟨ ap Cλ.ƛ (F.₁ Dλ.ev C.∘ comparison.inv C.∘ (pb.p₁ C.⊗₁ f))
Cλ
∎
: ExponentialP Gl Gl-cartesian _ A' B'
done .B^A' = pow
done .ev' = evm
done .has-is-exponential' .ƛ' {Γ' = Γ} {m = mβ} m .map = alpha module alpha where
done abstract
: ψ C.∘ F.₁ (Dλ.ƛ mβ) C.∘ Γ .map ≡ φ C.∘ Cλ.ƛ (m .map)
coh = Equiv.injective₂ (Equiv.inverse (_ , Cλ.lambda-is-equiv))
coh (
.unlambda (Cλ.ƛ (F.₁ Dλ.ev C.∘ comparison.inv C.∘ (C.id C.⊗₁ f)) C.∘ F.₁ (Dλ.ƛ mβ) C.∘ Γ .map)
Cλ.unlambda-∘ _ _ ⟩
≡⟨ Cλ.unlambda (Cλ.ƛ (F.₁ Dλ.ev C.∘ comparison.inv C.∘ (C.id C.⊗₁ f))) C.∘ ((F.₁ (Dλ.ƛ mβ) C.∘ Γ .map) C.⊗₁ C.id)
Cλ.pushl (Cλ.commutes _) ⟩
≡⟨ C.₁ Dλ.ev C.∘ (comparison.inv C.∘ (C.id C.⊗₁ f)) C.∘ ((F.₁ (Dλ.ƛ mβ) C.∘ Γ .map) C.⊗₁ C.id)
F._∘_ refl (C.pullr (Fr.weave C.×-functor (C.pulll (C.idl _) ,ₚ C.elimr refl ∙ C.introl F.F-id))) ⟩
≡⟨ ap₂ C.₁ Dλ.ev C.∘ comparison.inv C.∘ (F.₁ (Dλ.ƛ mβ) C.⊗₁ F.₁ D.id) C.∘ (Γ .map C.⊗₁ f)
F._∘_ refl (C.extendl (comparison-nat _ _ _)) ⟩
≡⟨ ap₂ C.₁ Dλ.ev C.∘ F.₁ (Dλ.ƛ mβ D.⊗₁ D.id) C.∘ comparison.inv C.∘ (Γ .map C.⊗₁ f)
F.pulll (F.collapse (Dλ.commutes _)) ⟩
≡⟨ C.₁ mβ C.∘ comparison.inv C.∘ (Γ .map C.⊗₁ f)
F.com ⟩
≡˘⟨ m .∘ m .map ∎)
g C(
.unlambda (Cλ.ƛ (g C.∘ Cλ.ev) C.∘ Cλ.ƛ (m .map))
Cλ.unlambda-∘ _ _ ⟩
≡⟨ Cλ.unlambda (Cλ.ƛ (g C.∘ Cλ.ev)) C.∘ (Cλ.ƛ (m .map) C.⊗₁ C.id)
Cλ.pushl (Cλ.commutes _) ⟩
≡⟨ C.∘ Cλ.unlambda (Cλ.ƛ (m .map))
g C._∘_ refl (Cλ.commutes _) ⟩
≡⟨ ap₂ C.∘ m .map ∎)
g C
= pb.universal {p₁' = F.₁ (Dλ.ƛ mβ) C.∘ Γ .map} {p₂' = Cλ.ƛ (m .map)} coh
alpha
.has-is-exponential' .ƛ' m .com = pb.p₁∘universal
done .has-is-exponential' .commutes' m = Slice-pathp comm1 where abstract
done : (Cλ.ev C.∘ (pb.p₂ C.⊗₁ C.id)) C.∘ (alpha.alpha m C.⊗₁ C.id) ≡ m .map
comm1 = C.pullr (sym (Bifunctor.first∘first C.×-functor))
comm1 (Cλ.ev C.∘_) (ap₂ C._⊗₁_ pb.p₂∘universal refl)
∙∙ ap .commutes _
∙∙ Cλ
.has-is-exponential' .unique' {Γ' = Γ} {m = mβ} {m'β} {p} {m' = m} m' q =
done (pb.unique coh₁ coh₂)
Slice-pathp where
: pb.p₁ C.∘ m' .map ≡ F.₁ (Dλ.ƛ mβ) C.∘ Γ .map
coh₁ =
coh₁ .p₁ C.∘ m' .map
pb.com ⟩
≡⟨ m' .₁ (m'β) C.∘ Γ .map
F._∘_ (ap F.₁ (Dλ.unique _ p)) refl ⟩
≡⟨ ap₂ C.₁ (Dλ.ƛ mβ) C.∘ Γ .map
F
∎
: pb.p₂ C.∘ m' .map ≡ Cλ.ƛ (m .map)
coh₂ = Cλ.unique _ $
coh₂ .ev C.∘ (pb.p₂ C.∘ m' .map) C.⊗₁ C.id
Cλ._∘_ refl (Bifunctor.first∘first C.×-functor) ⟩
≡⟨ ap₂ C.ev C.∘ pb.p₂ C.⊗₁ C.id C.∘ m' .map C.⊗₁ C.id
Cλ.pulll refl ∙ (λ i → q i .map) ⟩
≡⟨ C.map
m ∎
In this section, for notational convenience, we will refer to the objects of metonymically by their maps.↩︎
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.