cat/src/Cat/Product.agda

51 lines
2 KiB
Agda
Raw Normal View History

module Cat.Product where
open import Agda.Primitive
open import Data.Product
open import Cubical
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 ])
∃![ x ] ( [ π₁ x ] x₁ × [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency:
-- record IsProduct { ' : Level} ( : Category {} {'})
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (') where
-- field
-- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
-- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂)
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₂
arrowProduct : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
arrowProduct π₁ π₂ = proj₁ (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field
product : (A B : Object ) Product { = } A B
open Product
objectProduct : (A B : Object ) Object
objectProduct A B = Product.obj (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product
parallelProduct : {A A' B B' : Object } [ A , A' ] [ B , B' ]
[ objectProduct A B , objectProduct A' B' ]
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
( [ a (product A B) .proj₁ ])
( [ b (product A B) .proj₂ ])