module Cat.Displayed.Bifibration
{o ℓ o' ℓ'} {ℬ : Precategory o ℓ} (ℰ : Displayed ℬ o' ℓ') where
Note that a bifibration is not the same as a “profunctor
valued in categories”. Those are a distinct concept, called
two-sided fibrations.
record is-bifibration : Type (o ⊔ ℓ ⊔ o' ⊔ ℓ') where
field
fibration : Cartesian-fibration
opfibration : Cocartesian-fibration
module fibration = Cartesian-fibration fibration
module opfibration = Cocartesian-fibration opfibration
If
E
is a bifibration, then its opreindexing functors are left adjoints to its reindexing functors.
To show this, it will suffice to construct natural isomorphism between
Ey(u∗(−),−)
and
Ex(−,u∗(−)).
However, we have already shown that
Ey(u∗(−),−)
and
Ex(−,u∗(−))
are both naturally isomorphic to
Eu(−,−), so all we need to do is compose
these natural isomorphisms!
module _ (bifib : is-bifibration) where
open is-bifibration bifib
open Cat.Displayed.Cartesian.Indexing ℰ fibration
open Cat.Displayed.Cocartesian.Indexing ℰ opfibration
cobase-change⊣base-change
: ∀ {x y} (f : Hom x y)
→ cobase-change f ⊣ base-change f
cobase-change⊣base-change {x} {y} f =
hom-natural-iso→adjoints $
fibration→hom-iso fibration f ∘ni opfibration→hom-iso opfibration f ni⁻¹
In fact, if
EB
is a cartesian fibration where every reindexing functor has a left
adjoint, then
E
is a bifibration!
Since
E
is a fibration, every
u:x→y
in
B
induces a natural isomorphism
Eu(x′,−)≃Ex(x′,u∗(−)),
by the universal property of cartesian lifts. If
u∗
additionally has a left adjoint
Lu,
we have natural isomorphisms
Eu(x′,−)≃Ex(x′,u∗(−))≃Ey(Lu(x′)−),
which implies
E
is a weak opfibration
; and any
weak opfibration that’s also a fibration is a proper opfibration.
module _ (fib : Cartesian-fibration) where
open Cartesian-fibration fib
open Cat.Displayed.Cartesian.Indexing ℰ fib
left-adjoint-base-change→opfibration
: (L : ∀ {x y} → (f : Hom x y) → Functor (Fibre ℰ x) (Fibre ℰ y))
→ (∀ {x y} → (f : Hom x y) → (L f ⊣ base-change f))
→ Cocartesian-fibration
left-adjoint-base-change→opfibration L adj =
fibration+weak-opfibration→opfibration fib $
hom-iso→weak-opfibration L λ u →
adjunct-hom-iso-from (adj u) _ ni⁻¹ ∘ni fibration→hom-iso-from fib u
Let
f:X→Y
be a morphism in
B,
and let
L:EX→EY
be a functor. If we are given a natural transformation
η:Id→f∗∘L
with
f∘η
cocartesian, then
L
is a left adjoint to
f∗.
cocartesian→left-adjoint-base-change
: ∀ {x y} {L : Functor (Fibre ℰ x) (Fibre ℰ y)} {f : Hom x y}
→ (L-unit : Id => base-change f F∘ L)
→ (∀ x → is-cocartesian (f ∘ id) (π* f (L .F₀ x) ∘' L-unit .η x))
→ L ⊣ base-change f
We will construct the adjunction by constructing a natural
equivalence of
Hom-sets
EXLA,B≃EYA,f∗B.
The map
v↦f∗(v)∘η
gives us the forward direction of this equivalence, so all that remains
is to find an inverse. Since
f∘η
is cocartesian, it satisfies a mapping-out universal property:
if
v:A→f∗B
is a vertical map in
E,
we can construct a vertical map
LA→B
by factoring
f∘v
through the cocartesian
f∘η;
The actual universal property says that this factorising
process is an equivalence.
cocartesian→left-adjoint-base-change {x = x} {y = y} {L = L} {f = f} L-unit cocart =
hom-iso→adjoints (λ v → base-change f .F₁ v Fib.∘ L-unit .η _)
precompose-equiv equiv-natural where
module cocart x = is-cocartesian (cocart x)
module f* = Functor (base-change f)
precompose-equiv
: ∀ {x' : Ob[ x ]} {y' : Ob[ y ]}
→ is-equiv {A = Hom[ id ] (L .F₀ x') y'} (λ v → f*.₁ v Fib.∘ L-unit .η x')
precompose-equiv {x'} {y'} = is-iso→is-equiv $ iso
(λ v → cocart.universalv _ (π* f _ ∘' v))
(λ v → π*.uniquep₂ _ _ refl _ _
(Fib.pulllf (π*.commutesp id-comm _)
∙[] symP (assoc' _ _ _)
∙[] cocart.commutesv x' _)
refl)
(λ v → symP $ cocart.uniquep x' _ _ _ _ $
assoc' _ _ _
∙[] Fib.pushlf (symP $ π*.commutesp id-comm _))
Furthermore, this equivalence is natural, but that’s a very tedious
proof.
equiv-natural
: hom-iso-natural {L = L} {R = base-change f} (λ v → f*.₁ v Fib.∘ L-unit .η _)
equiv-natural g h k =
π*.uniquep₂ _ _ _ _ _
(Fib.pulllf (π*.commutesp id-comm _)
∙[] pushl[] _ (pushl[] _ (to-pathp⁻ (smashr _ _))))
(Fib.pulllf (π*.commutesp id-comm _)
∙[] extendr[] _ (Fib.pulllf (Fib.pulllf (π*.commutesp id-comm _)))
∙[] extendr[] _ (pullr[] _ (to-pathp (L-unit .is-natural _ _ h)))
∙[] pullr[] _ (Fib.pulllf (extendr[] _ (π*.commutesp id-comm _))))