cat/src/Cat/Category/Product.agda

62 lines
2 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-03-08 09:20:29 +00:00
open import Data.Product as P hiding (_×_ ; proj₁ ; proj₂)
2018-03-08 09:20:29 +00:00
open import Cat.Category hiding (module Propositionality)
open Category
2018-03-08 09:20:29 +00:00
module _ {a b : Level} where
record RawProduct { : Category a b} (A B : Object ) : Set (a b) where
no-eta-equality
field
obj : Object
proj₁ : [ obj , A ]
proj₂ : [ obj , B ]
record IsProduct ( : Category a b) {A B : Object } (raw : RawProduct { = } A B) : Set (a b) where
open RawProduct raw public
field
isProduct : {X : Object } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ proj₁ x ] x₁ P.× [ proj₂ x ] x₂)
-- | Arrow product
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
_P[_×_] π₁ π₂ = P.proj₁ (isProduct π₁ π₂)
record Product { : Category a b} (A B : Object ) : Set (a b) where
field
raw : RawProduct { = } A B
isProduct : IsProduct {A} {B} raw
open IsProduct isProduct public
record HasProducts ( : Category a b) : Set (a b) where
field
product : (A B : Object ) Product { = } A B
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
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 ]
]
module Propositionality where
-- TODO `isProp (Product ...)`
-- TODO `isProp (HasProducts ...)`