module Cat.Functor.Closed where
When taken as a (bi)category,
the collection of (pre)categories is, in a suitably weak sense, Cartesian
closed: there is an equivalence between the functor categories
and
.
We do not define the full equivalence here, leaving the natural
isomorphisms aside and focusing on the inverse functors themselves:
Curry
and Uncurry
.
The two conversion functions act on objects essentially in the same way as currying and uncurrying behave on functions: the difference is that we must properly stage the action on morphisms. Currying a functor fixes a morphism , and we must show that is natural in . It follows from a bit of calculation using the functoriality of .
: Functor (C ×ᶜ D) E → Functor C Cat[ D , E ]
Curry {C = C} {D = D} {E = E} F = curried where
Curry open import Cat.Functor.Bifunctor {C = C} {D = D} {E = E} F
: Functor C Cat[ D , E ]
curried .F₀ = Right
curried .F₁ x→y = NT (λ f → first x→y) λ x y f →
curried (F-∘ F _ _)
sym (F₁ F) (Σ-pathp (C .idr _ ∙ sym (C .idl _)) (D .idl _ ∙ sym (D .idr _)))
·· ap _ _
·· F-∘ F .F-id = Nat-path λ x → F-id F
curried .F-∘ f g = Nat-path λ x →
curried (λ x → F₁ F (_ , x)) (sym (D .idl _)) ∙ F-∘ F _ _
ap
: Functor C Cat[ D , E ] → Functor (C ×ᶜ D) E
Uncurry {C = C} {D = D} {E = E} F = uncurried where
Uncurry import Cat.Reasoning C as C
import Cat.Reasoning D as D
import Cat.Reasoning E as E
module F = Functor F
: Functor (C ×ᶜ D) E
uncurried .F₀ (c , d) = F₀ (F.₀ c) d
uncurried .F₁ (f , g) = F.₁ f .η _ E.∘ F₁ (F.₀ _) g uncurried
The other direction must do slightly more calculation: Given a functor into functor-categories, and a pair of arguments, we must apply it twice: but at the level of morphisms, this involves composition in the codomain category, which throws a fair bit of complication into the functoriality constraints.
.F-id {x = x , y} = path where abstract
uncurried : E ._∘_ (F.₁ (C .id) .η y) (F₁ (F.₀ x) (D .id)) ≡ E .id
path =
path .₁ C.id .η y E.∘ F₁ (F.₀ x) D.id ≡⟨ E.elimr (F-id (F.₀ x)) ⟩
F.₁ C.id .η y ≡⟨ (λ i → F.F-id i .η y) ⟩
F.id ∎
E
.F-∘ (f , g) (f' , g') = path where abstract
uncurried : uncurried .F₁ (f C.∘ f' , g D.∘ g')
path .F₁ (f , g) E.∘ uncurried .F₁ (f' , g')
≡ uncurried =
path .₁ (f C.∘ f') .η _ E.∘ F₁ (F.₀ _) (g D.∘ g') ≡˘⟨ E.pulll (λ i → F.F-∘ f f' (~ i) .η _) ⟩
F.₁ f .η _ E.∘ F.₁ f' .η _ E.∘ ⌜ F₁ (F.₀ _) (g D.∘ g') ⌝ ≡⟨ ap! (F-∘ (F.₀ _) _ _) ⟩
F.₁ f .η _ E.∘ F.₁ f' .η _ E.∘ F₁ (F.₀ _) g E.∘ F₁ (F.₀ _) g' ≡⟨ cat! E ⟩
F.₁ f .η _ E.∘ ⌜ F.₁ f' .η _ E.∘ F₁ (F.₀ _) g ⌝ E.∘ F₁ (F.₀ _) g' ≡⟨ ap! (F.₁ f' .is-natural _ _ _) ⟩
F.₁ f .η _ E.∘ (F₁ (F.₀ _) g E.∘ F.₁ f' .η _) E.∘ F₁ (F.₀ _) g' ≡⟨ cat! E ⟩
F((F.₁ f .η _ E.∘ F₁ (F.₀ _) g) E.∘ (F.₁ f' .η _ E.∘ F₁ (F.₀ _) g')) ∎