Use dotted expression in Cat

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-08 11:20:51 +01:00
parent 5ad506a09f
commit 48672b01bd

View file

@ -17,7 +17,6 @@ open import Cat.Category.NaturalTransformation
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
open Functor using (fmap ; omap)
open Category using (Object ; 𝟙) open Category using (Object ; 𝟙)
-- The category of categories -- The category of categories
@ -169,14 +168,19 @@ module _ { ' : Level} (unprovable : IsCategory (RawCat ')) where
-- Basically proves that `Cat ` is cartesian closed. -- Basically proves that `Cat ` is cartesian closed.
module CatExponential { : Level} ( 𝔻 : Category ) where module CatExponential { : Level} ( 𝔻 : Category ) where
open Data.Product private
open import Cat.Categories.Fun open Data.Product
open import Cat.Categories.Fun
module = Category
module 𝔻 = Category 𝔻
Category = Category Category = Category
open Fun 𝔻 renaming (identity to idN) open Fun 𝔻 renaming (identity to idN)
private private
:omap: : Functor 𝔻 × Object Object 𝔻 :omap: : Functor 𝔻 × Object Object 𝔻
:omap: (F , A) = omap F A :omap: (F , A) = F.omap A
where
module F = Functor F
prodObj : Category prodObj : Category
prodObj = Fun prodObj = Fun
@ -193,28 +197,31 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
B : Object B : Object
B = proj₂ cod B = proj₂ cod
module F = Functor F
module G = Functor G
:fmap: : (pobj : NaturalTransformation F G × [ A , B ]) :fmap: : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ omap F A , omap G B ] 𝔻 [ F.omap A , G.omap B ]
:fmap: ((θ , θNat) , f) = result :fmap: ((θ , θNat) , f) = result
where where
θA : 𝔻 [ omap F A , omap G A ] θA : 𝔻 [ F.omap A , G.omap A ]
θA = θ A θA = θ A
θB : 𝔻 [ omap F B , omap G B ] θB : 𝔻 [ F.omap B , G.omap B ]
θB = θ B θB = θ B
F→f : 𝔻 [ omap F A , omap F B ] F→f : 𝔻 [ F.omap A , F.omap B ]
F→f = fmap F f F→f = F.fmap f
G→f : 𝔻 [ omap G A , omap G B ] G→f : 𝔻 [ G.omap A , G.omap B ]
G→f = fmap G f G→f = G.fmap f
l : 𝔻 [ omap F A , omap G B ] l : 𝔻 [ F.omap A , G.omap B ]
l = 𝔻 [ θB Ff ] l = 𝔻 [ θB F.fmap f ]
r : 𝔻 [ omap F A , omap G B ] r : 𝔻 [ F.omap A , G.omap B ]
r = 𝔻 [ Gf θA ] r = 𝔻 [ G.fmap f θA ]
-- There are two choices at this point, -- There are two choices at this point,
-- but I suppose the whole point is that -- but I suppose the whole point is that
-- by `θNat f` we have `l ≡ r` -- by `θNat f` we have `l ≡ r`
-- lem : 𝔻 [ θ B ∘ F .fmap f ] ≡ 𝔻 [ G .fmap f ∘ θ A ] -- lem : 𝔻 [ θ B ∘ F .fmap f ] ≡ 𝔻 [ G .fmap f ∘ θ A ]
-- lem = θNat f -- lem = θNat f
result : 𝔻 [ omap F A , omap G B ] result : 𝔻 [ F.omap A , G.omap B ]
result = l result = l
open CatProduct renaming (obj to _×p_) using () open CatProduct renaming (obj to _×p_) using ()
@ -237,23 +244,27 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
:ident: = begin :ident: = begin
:fmap: {c} {c} (𝟙 (prodObj ×p ) {c}) ≡⟨⟩ :fmap: {c} {c} (𝟙 (prodObj ×p ) {c}) ≡⟨⟩
:fmap: {c} {c} (idN F , 𝟙 ) ≡⟨⟩ :fmap: {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C fmap F (𝟙 )] ≡⟨⟩ 𝔻 [ identityTrans F C F.fmap (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 fmap F (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity 𝔻 [ 𝟙 𝔻 F.fmap (𝟙 )] ≡⟨ proj₂ 𝔻.isIdentity
fmap F (𝟙 ) ≡⟨ F.isIdentity F.fmap (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻 𝟙 𝔻
where where
open module 𝔻 = Category 𝔻
open module F = Functor F open module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where module _ {F×A G×B H×C : Functor 𝔻 × Object } where
F = F×A .proj₁ private
A = F×A .proj₂ F = F×A .proj₁
G = G×B .proj₁ A = F×A .proj₂
B = G×B .proj₂ G = G×B .proj₁
H = H×C .proj₁ B = G×B .proj₂
C = H×C .proj₂ H = H×C .proj₁
-- Not entirely clear what this is at this point: C = H×C .proj₂
_P⊕_ = Category._∘_ (prodObj ×p ) {F×A} {G×B} {H×C} module F = Functor F
module G = Functor G
module H = Functor H
-- Not entirely clear what this is at this point:
_P⊕_ = Category._∘_ (prodObj ×p ) {F×A} {G×B} {H×C}
module _ module _
-- NaturalTransformation F G × .Arrow A B -- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × [ A , B ]} {θ×f : NaturalTransformation F G × [ A , B ]}
@ -279,31 +290,28 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
ηθNat = proj₂ ηθNT ηθNat = proj₂ ηθNT
:isDistributive: : :isDistributive: :
𝔻 [ 𝔻 [ η C θ C ] fmap F ( [ g f ] ) ] 𝔻 [ 𝔻 [ η C θ C ] F.fmap ( [ g f ] ) ]
𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ θ B fmap F f ] ] 𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
:isDistributive: = begin :isDistributive: = begin
𝔻 [ (ηθ C) fmap F ( [ g f ]) ] 𝔻 [ (ηθ C) F.fmap ( [ g f ]) ]
≡⟨ ηθNat ( [ g f ]) ≡⟨ ηθNat ( [ g f ])
𝔻 [ fmap H ( [ g f ]) (ηθ A) ] 𝔻 [ H.fmap ( [ g f ]) (ηθ A) ]
≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive) ≡⟨ cong (λ φ 𝔻 [ φ ηθ A ]) (H.isDistributive)
𝔻 [ 𝔻 [ fmap H g fmap H f ] (ηθ A) ] 𝔻 [ 𝔻 [ H.fmap g H.fmap f ] (ηθ A) ]
≡⟨ sym isAssociative ≡⟨ sym 𝔻.isAssociative
𝔻 [ fmap H g 𝔻 [ fmap H f ηθ A ] ] 𝔻 [ H.fmap g 𝔻 [ H.fmap f ηθ A ] ]
≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) isAssociative ≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) 𝔻.isAssociative
𝔻 [ fmap H g 𝔻 [ 𝔻 [ fmap H f η A ] θ A ] ] 𝔻 [ H.fmap g 𝔻 [ 𝔻 [ H.fmap f η A ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f))) ≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (cong (λ φ 𝔻 [ φ θ A ]) (sym (ηNat f)))
𝔻 [ fmap H g 𝔻 [ 𝔻 [ η B fmap G f ] θ A ] ] 𝔻 [ H.fmap g 𝔻 [ 𝔻 [ η B G.fmap f ] θ A ] ]
≡⟨ cong (λ φ 𝔻 [ fmap H g φ ]) (sym isAssociative) ≡⟨ cong (λ φ 𝔻 [ H.fmap g φ ]) (sym 𝔻.isAssociative)
𝔻 [ fmap H g 𝔻 [ η B 𝔻 [ fmap G f θ A ] ] ] 𝔻 [ H.fmap g 𝔻 [ η B 𝔻 [ G.fmap f θ A ] ] ]
≡⟨ isAssociative ≡⟨ 𝔻.isAssociative
𝔻 [ 𝔻 [ fmap H g η B ] 𝔻 [ fmap G f θ A ] ] 𝔻 [ 𝔻 [ H.fmap g η B ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ fmap G f θ A ] ]) (sym (ηNat g)) ≡⟨ cong (λ φ 𝔻 [ φ 𝔻 [ G.fmap f θ A ] ]) (sym (ηNat g))
𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ fmap G f θ A ] ] 𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ G.fmap f θ A ] ]
≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C fmap G g ] φ ]) (sym (θNat f)) ≡⟨ cong (λ φ 𝔻 [ 𝔻 [ η C G.fmap g ] φ ]) (sym (θNat f))
𝔻 [ 𝔻 [ η C fmap G g ] 𝔻 [ θ B fmap F f ] ] 𝔻 [ 𝔻 [ η C G.fmap g ] 𝔻 [ θ B F.fmap f ] ]
where
open Category 𝔻
module H = Functor H
eval : Functor (CatProduct.obj prodObj ) 𝔻 eval : Functor (CatProduct.obj prodObj ) 𝔻
-- :eval: : Functor (prodObj ×p ) 𝔻 -- :eval: : Functor (prodObj ×p ) 𝔻