Formatting

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-21 13:25:24 +01:00
parent 890154a81d
commit cd3514c8cf
1 changed files with 59 additions and 96 deletions

View File

@ -18,8 +18,6 @@ open import Cat.Category.NaturalTransformation
open import Cat.Equality
open Equality.Data.Product
open Category using (Object ; 𝟙)
-- The category of categories
module _ ( ' : Level) where
private
@ -35,19 +33,18 @@ module _ ( ' : Level) where
ident-l = Functor≡ refl
RawCat : RawCategory (lsuc ( ')) ( ')
RawCat =
record
{ Object = Category '
; Arrow = Functor
; 𝟙 = identity
; _∘_ = F[_∘_]
}
RawCategory.Object RawCat = Category '
RawCategory.Arrow RawCat = Functor
RawCategory.𝟙 RawCat = identity
RawCategory._∘_ RawCat = F[_∘_]
private
open RawCategory RawCat
isAssociative : IsAssociative
isAssociative {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
ident : IsIdentity identity
ident = ident-l , ident-r
isIdentity : IsIdentity identity
isIdentity = ident-l , ident-r
-- NB! `ArrowsAreSets RawCat` is *not* provable. The type of functors,
-- however, form a groupoid! Therefore there is no (1-)category of
@ -55,7 +52,7 @@ module _ ( ' : Level) where
--
-- Because of this there is no category of categories.
Cat : (unprovable : IsCategory RawCat) Category (lsuc ( ')) ( ')
Category.raw (Cat _) = RawCat
Category.raw (Cat _) = RawCat
Category.isCategory (Cat unprovable) = unprovable
-- | In the following we will pretend there is a category of categories when
@ -72,28 +69,31 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
module = Category
module 𝔻 = Category 𝔻
Obj = Object × Object 𝔻
Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
𝟙' : {o : Obj} Arr o o
𝟙' = 𝟙 , 𝟙 𝔻
_∘_ :
{a b c : Obj}
Arr b c
Arr a b
Arr a c
_∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
module _ where
private
Obj = .Object × 𝔻.Object
Arr : Obj Obj Set '
Arr (c , d) (c' , d') = [ c , c' ] × 𝔻 [ d , d' ]
𝟙 : {o : Obj} Arr o o
𝟙 = .𝟙 , 𝔻.𝟙
_∘_ :
{a b c : Obj}
Arr b c
Arr a b
Arr a c
_∘_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) [ bc∈C ab∈C ] , 𝔻 [ bc∈D ab∈D ]}
rawProduct : RawCategory '
RawCategory.Object rawProduct = Obj
RawCategory.Arrow rawProduct = Arr
RawCategory.𝟙 rawProduct = 𝟙'
RawCategory._∘_ rawProduct = _∘_
open RawCategory rawProduct
rawProduct : RawCategory '
RawCategory.Object rawProduct = Obj
RawCategory.Arrow rawProduct = Arr
RawCategory.𝟙 rawProduct = 𝟙
RawCategory._∘_ rawProduct = _∘_
open RawCategory rawProduct
arrowsAreSets : ArrowsAreSets
arrowsAreSets = setSig {sA = .arrowsAreSets} {sB = λ x 𝔻.arrowsAreSets}
isIdentity : IsIdentity 𝟙'
isIdentity : IsIdentity 𝟙
isIdentity
= Σ≡ (fst .isIdentity) (fst 𝔻.isIdentity)
, Σ≡ (snd .isIdentity) (snd 𝔻.isIdentity)
@ -189,102 +189,65 @@ module CatExponential { : Level} ( 𝔻 : Category ) where
Category = Category
open Fun 𝔻 renaming (identity to idN)
omap : Functor 𝔻 × Object Object 𝔻
omap : Functor 𝔻 × .Object 𝔻.Object
omap (F , A) = Functor.omap F A
-- The exponential object
object : Category
object = Fun
module _ {dom cod : Functor 𝔻 × Object } where
module _ {dom cod : Functor 𝔻 × .Object} where
open Σ dom renaming (proj₁ to F ; proj₂ to A)
open Σ cod renaming (proj₁ to G ; proj₂ to B)
private
F : Functor 𝔻
F = proj₁ dom
A : Object
A = proj₂ dom
G : Functor 𝔻
G = proj₁ cod
B : Object
B = proj₂ cod
module F = Functor F
module G = Functor G
fmap : (pobj : NaturalTransformation F G × [ A , B ])
𝔻 [ F.omap A , G.omap B ]
fmap ((θ , θNat) , f) = result
where
θA : 𝔻 [ F.omap A , G.omap A ]
θA = θ A
θB : 𝔻 [ F.omap B , G.omap B ]
θB = θ B
F→f : 𝔻 [ F.omap A , F.omap B ]
F→f = F.fmap f
G→f : 𝔻 [ G.omap A , G.omap B ]
G→f = G.fmap f
l : 𝔻 [ F.omap A , G.omap B ]
l = 𝔻 [ θB F.fmap f ]
r : 𝔻 [ F.omap A , G.omap B ]
r = 𝔻 [ G.fmap f θA ]
result : 𝔻 [ F.omap A , G.omap B ]
result = l
fmap ((θ , θNat) , f) = 𝔻 [ θ B F.fmap f ]
-- Alternatively:
--
-- fmap ((θ , θNat) , f) = 𝔻 [ G.fmap f ∘ θ A ]
--
-- Since they are equal by naturality of θ.
open CatProduct renaming (object to _⊗_) using ()
module _ {c : Functor 𝔻 × Object } where
private
F : Functor 𝔻
F = proj₁ c
C : Object
C = proj₂ c
module _ {c : Functor 𝔻 × .Object} where
open Σ c renaming (proj₁ to F ; proj₂ to C)
ident : fmap {c} {c} (NT.identity F , 𝟙 {A = proj₂ c}) 𝟙 𝔻
ident : fmap {c} {c} (NT.identity F , .𝟙 {A = proj₂ c}) 𝔻.𝟙
ident = begin
fmap {c} {c} (𝟙 (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , 𝟙 ) ≡⟨⟩
𝔻 [ identityTrans F C F.fmap (𝟙 )] ≡⟨⟩
𝔻 [ 𝟙 𝔻 F.fmap (𝟙 )] ≡⟨ 𝔻.leftIdentity
F.fmap (𝟙 ) ≡⟨ F.isIdentity
𝟙 𝔻
fmap {c} {c} (Category.𝟙 (object ) {c}) ≡⟨⟩
fmap {c} {c} (idN F , .𝟙) ≡⟨⟩
𝔻 [ identityTrans F C F.fmap .𝟙 ] ≡⟨⟩
𝔻 [ 𝔻.𝟙 F.fmap .𝟙 ] ≡⟨ 𝔻.leftIdentity
F.fmap .𝟙 ≡⟨ F.isIdentity
𝔻.𝟙
where
module F = Functor F
module _ {F×A G×B H×C : Functor 𝔻 × Object } where
module _ {F×A G×B H×C : Functor 𝔻 × .Object} where
open Σ F×A renaming (proj₁ to F ; proj₂ to A)
open Σ G×B renaming (proj₁ to G ; proj₂ to B)
open Σ H×C renaming (proj₁ to H ; proj₂ to C)
private
F = F×A .proj₁
A = F×A .proj₂
G = G×B .proj₁
B = G×B .proj₂
H = H×C .proj₁
C = H×C .proj₂
module F = Functor F
module G = Functor G
module H = Functor H
module _
-- NaturalTransformation F G × .Arrow A B
{θ×f : NaturalTransformation F G × [ A , B ]}
{η×g : NaturalTransformation G H × [ B , C ]} where
open Σ θ×f renaming (proj₁ to θNT ; proj₂ to f)
open Σ θNT renaming (proj₁ to θ ; proj₂ to θNat)
open Σ η×g renaming (proj₁ to ηNT ; proj₂ to g)
open Σ ηNT renaming (proj₁ to η ; proj₂ to ηNat)
private
θ : Transformation F G
θ = proj₁ (proj₁ θ×f)
θNat : Natural F G θ
θNat = proj₂ (proj₁ θ×f)
f : [ A , B ]
f = proj₂ θ×f
η : Transformation G H
η = proj₁ (proj₁ η×g)
ηNat : Natural G H η
ηNat = proj₂ (proj₁ η×g)
g : [ B , C ]
g = proj₂ η×g
ηθNT : NaturalTransformation F H
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
ηθ = proj₁ ηθNT
ηθNat = proj₂ ηθNT
ηθNT = NT[_∘_] {F} {G} {H} ηNT θNT
open Σ ηθNT renaming (proj₁ to ηθ ; proj₂ to ηθNat)
isDistributive :
𝔻 [ 𝔻 [ η C θ C ] F.fmap ( [ g f ] ) ]