cat/src/Cat/Category/Product.agda

69 lines
2.2 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Category.Product where
open import Agda.Primitive
open import Cubical
open import Data.Product as P hiding (_×_ ; proj₁ ; proj₂)
open import Cat.Category hiding (module Propositionality)
module _ {a b : Level} ( : Category a b) where
open Category
module _ (A B : Object) where
record RawProduct : Set (a b) where
no-eta-equality
field
object : Object
proj₁ : [ object , A ]
proj₂ : [ object , B ]
-- FIXME Not sure this is actually a proposition - so this name is
-- misleading.
record IsProduct (raw : RawProduct) : Set (a b) where
open RawProduct raw public
field
isProduct : {X : Object} (f : [ X , A ]) (g : [ X , B ])
∃![ f×g ] ( [ proj₁ f×g ] f P.× [ proj₂ f×g ] g)
-- | Arrow product
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , object ]
_P[_×_] π₁ π₂ = P.proj₁ (isProduct π₁ π₂)
record Product : Set (a b) where
field
raw : RawProduct
isProduct : IsProduct raw
open IsProduct isProduct public
record HasProducts : Set (a b) where
field
product : (A B : Object) Product A B
_×_ : Object Object Object
A × B = Product.object (product A B)
-- | 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' ]
f |×| g = product A' B'
P[ [ f fst ]
× [ g snd ]
]
module Propositionality {a b : Level} { : Category a b} {A B : Category.Object } where
-- TODO I'm not sure this is actually provable. Check with Thierry.
propProduct : isProp (Product A B)
propProduct = {!!}
propHasProducts : isProp (HasProducts )
propHasProducts = {!!}