cat/src/Cat/Category/Product.agda

61 lines
2.1 KiB
Agda
Raw Normal View History

2018-02-05 13:59:53 +00:00
module Cat.Category.Product where
open import Agda.Primitive
open import Cubical
2018-02-05 15:35:33 +00:00
open import Data.Product as P hiding (_×_)
open import Cat.Category
open Category
module _ { ' : Level} ( : Category ') {A B obj : Object } where
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
2018-02-05 15:35:33 +00:00
∃![ x ] ( [ π₁ x ] x₁ P.× [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency:
2018-02-05 15:35:33 +00:00
-- record IsProduct {a b : Level} ( : Category a b)
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (a ⊔ b) where
-- field
2018-02-05 15:35:33 +00:00
-- issProduct : ∀ {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
-- → ∃![ x ] ( [ π₁ ∘ x ] ≡ x₁ P.× [ π₂ ∘ x ] ≡ x₂)
-- open IsProduct
record Product { ' : Level} { : Category '} (A B : Object ) : Set ( ') where
no-eta-equality
field
obj : Object
proj₁ : [ obj , A ]
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
-- | Arrow product
2018-02-05 15:35:33 +00:00
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
2018-02-05 15:35:33 +00:00
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field
product : (A B : Object ) Product { = } A B
open Product hiding (obj)
module _ (A B : Object ) where
open Product (product A B)
_×_ : Object
_×_ = obj
-- | Parallel product of arrows
--
-- The product mentioned in awodey in Def 6.1 is not the regular product of
-- arrows. It's a "parallel" product
module _ {A A' B B' : Object } where
open Product (product A B) hiding (_P[_×_]) renaming (proj₁ to fst ; proj₂ to snd)
_|×|_ : [ A , A' ] [ B , B' ] [ A × B , A' × B' ]
a |×| b = product A' B'
P[ [ a fst ]
× [ b snd ]
2018-02-05 15:35:33 +00:00
]