cat/src/Cat/Category/Product.agda

124 lines
4.1 KiB
Agda
Raw Normal View History

2018-03-08 09:45:15 +00:00
{-# OPTIONS --allow-unsolved-metas #-}
2018-02-05 13:59:53 +00:00
module Cat.Category.Product where
open import Agda.Primitive
open import Cubical
open import Cubical.NType.Properties using (lemPropF)
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)
2018-03-08 09:23:37 +00:00
module _ {a b : Level} ( : Category a b) where
2018-03-08 09:23:37 +00:00
open Category
2018-03-08 09:28:05 +00:00
module _ (A B : Object) where
record RawProduct : Set (a b) where
no-eta-equality
field
2018-03-08 09:45:15 +00:00
object : Object
proj₁ : [ object , A ]
proj₂ : [ object , B ]
2018-03-08 09:28:05 +00:00
-- FIXME Not sure this is actually a proposition - so this name is
-- misleading.
2018-03-08 09:28:05 +00:00
record IsProduct (raw : RawProduct) : Set (a b) where
open RawProduct raw public
field
ump : {X : Object} (f : [ X , A ]) (g : [ X , B ])
∃![ f×g ] ( [ proj₁ f×g ] f P.× [ proj₂ f×g ] g)
2018-03-08 09:28:05 +00:00
-- | Arrow product
_P[_×_] : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
2018-03-08 09:45:15 +00:00
[ X , object ]
_P[_×_] π₁ π₂ = P.proj₁ (ump π₁ π₂)
2018-03-08 09:28:05 +00:00
record Product : Set (a b) where
field
raw : RawProduct
isProduct : IsProduct raw
open IsProduct isProduct public
2018-03-08 09:20:29 +00:00
2018-03-08 09:23:37 +00:00
record HasProducts : Set (a b) where
2018-03-08 09:20:29 +00:00
field
2018-03-08 09:23:37 +00:00
product : (A B : Object) Product A B
2018-03-08 09:20:29 +00:00
2018-03-08 09:30:35 +00:00
_×_ : Object Object Object
2018-03-08 09:45:15 +00:00
A × B = Product.object (product A B)
2018-03-08 09:20:29 +00:00
-- | 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
2018-03-08 09:23:37 +00:00
module _ {A A' B B' : Object} where
2018-03-08 09:20:29 +00:00
open Product
open Product (product A B) hiding (_P[_×_]) renaming (proj₁ to fst ; proj₂ to snd)
_|×|_ : [ A , A' ] [ B , B' ] [ A × B , A' × B' ]
2018-03-08 09:30:35 +00:00
f |×| g = product A' B'
P[ [ f fst ]
× [ g snd ]
2018-03-08 09:20:29 +00:00
]
module _ {a b : Level} { : Category a b} {A B : Category.Object } where
private
open Category
module _ (raw : RawProduct A B) where
module _ (x y : IsProduct A B raw) where
private
module x = IsProduct x
module y = IsProduct y
module _ {X : Object} (f : [ X , A ]) (g : [ X , B ]) where
prodAux : x.ump f g y.ump f g
prodAux = {!!}
propIsProduct' : x y
propIsProduct' i = record { ump = λ f g prodAux f g i }
propIsProduct : isProp (IsProduct A B raw)
propIsProduct = propIsProduct'
Product≡ : {x y : Product A B} (Product.raw x Product.raw y) x y
Product≡ {x} {y} p i = record { raw = p i ; isProduct = q i }
where
q : (λ i IsProduct A B (p i)) [ Product.isProduct x Product.isProduct y ]
q = lemPropF propIsProduct p
module _ {a b : Level} { : Category a b} {A B : Category.Object } where
open Category
private
module _ (x y : HasProducts ) where
private
module x = HasProducts x
module y = HasProducts y
module _ (A B : Object) where
module pX = Product (x.product A B)
module pY = Product (y.product A B)
objEq : pX.object pY.object
objEq = {!!}
proj₁Eq : (λ i [ objEq i , A ]) [ pX.proj₁ pY.proj₁ ]
proj₁Eq = {!!}
proj₂Eq : (λ i [ objEq i , B ]) [ pX.proj₂ pY.proj₂ ]
proj₂Eq = {!!}
rawEq : pX.raw pY.raw
RawProduct.object (rawEq i) = objEq i
RawProduct.proj₁ (rawEq i) = {!!}
RawProduct.proj₂ (rawEq i) = {!!}
isEq : (λ i IsProduct A B (rawEq i)) [ pX.isProduct pY.isProduct ]
isEq = {!!}
appEq : x.product A B y.product A B
appEq = Product≡ rawEq
productEq : x.product y.product
productEq i = λ A B appEq A B i
propHasProducts' : x y
propHasProducts' i = record { product = productEq i }
propHasProducts : isProp (HasProducts )
propHasProducts = propHasProducts'