cat/src/Cat/Category/Product.agda

358 lines
13 KiB
Agda
Raw Normal View History

{-# OPTIONS --cubical --caching #-}
2018-02-05 13:59:53 +00:00
module Cat.Category.Product where
2018-04-05 08:41:56 +00:00
open import Cat.Prelude as P hiding (_×_ ; fst ; snd)
open import Cat.Equivalence
open import Cat.Category
2018-03-08 09:23:37 +00:00
module _ {a b : Level} ( : Category a b) where
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
2018-04-05 08:41:56 +00:00
fst : [ object , A ]
snd : [ object , B ]
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 ])
2018-04-05 08:41:56 +00:00
∃![ f×g ] ( [ fst f×g ] f P.× [ snd 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 ]
2018-04-05 08:41:56 +00:00
_P[_×_] π₁ π₂ = P.fst (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-04-05 08:41:56 +00:00
open Product using (_P[_×_])
open Product (product A B) hiding (_P[_×_]) renaming (fst to fst ; snd to snd)
2018-03-08 09:20:29 +00:00
_|×|_ : [ 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
2018-04-03 10:40:20 +00:00
module _ (f×g : Arrow X y.object) where
2018-04-05 08:41:56 +00:00
help : isProp ({y} ( [ y.fst y ] f) P.× ( [ y.snd y ] g) f×g y)
2018-04-03 10:40:20 +00:00
help = propPiImpl (λ _ propPi (λ _ arrowsAreSets _ _))
res = ∃-unique (x.ump f g) (y.ump f g)
prodAux : x.ump f g y.ump f g
2018-04-03 10:40:20 +00:00
prodAux = lemSig ((λ f×g propSig (propSig (arrowsAreSets _ _) λ _ arrowsAreSets _ _) (λ _ help f×g))) _ _ res
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 Try0 {a b : Level} { : Category a b}
2018-05-01 09:41:45 +00:00
(let module = Category ) {𝒜 : .Object} where
2018-04-05 08:41:56 +00:00
open P
2018-03-29 22:12:01 +00:00
module _ where
raw : RawCategory _ _
raw = record
2018-05-01 09:41:45 +00:00
{ Object = Σ[ X .Object ] .Arrow X 𝒜 × .Arrow X
2018-04-25 06:19:36 +00:00
; Arrow = λ{ (A , a0 , a1) (B , b0 , b1)
Σ[ f .Arrow A B ]
[ b0 f ] a0
× [ b1 f ] a1
2018-03-29 22:12:01 +00:00
}
; identity = λ{ {X , f , g} .identity {X} , .rightIdentity , .rightIdentity}
2018-04-11 08:58:50 +00:00
; _<<<_ = λ { {_ , a0 , a1} {_ , b0 , b1} {_ , c0 , c1} (f , f0 , f1) (g , g0 , g1)
(f .<<< g)
2018-03-29 22:12:01 +00:00
, (begin
[ c0 [ f g ] ] ≡⟨ .isAssociative
[ [ c0 f ] g ] ≡⟨ cong (λ φ [ φ g ]) f0
[ b0 g ] ≡⟨ g0
a0
)
, (begin
[ c1 [ f g ] ] ≡⟨ .isAssociative
[ [ c1 f ] g ] ≡⟨ cong (λ φ [ φ g ]) f1
[ b1 g ] ≡⟨ g1
a1
)
}
}
2018-04-05 12:37:25 +00:00
module _ where
2018-04-13 13:26:46 +00:00
open RawCategory raw
2018-03-29 22:12:01 +00:00
2018-04-13 12:35:54 +00:00
propEqs : {X' : Object}{Y' : Object} (let X , xa , xb = X') (let Y , ya , yb = Y')
(xy : .Arrow X Y) isProp ( [ ya xy ] xa × [ yb xy ] xb)
2018-04-13 12:35:54 +00:00
propEqs xs = propSig (.arrowsAreSets _ _) (\ _ .arrowsAreSets _ _)
2018-04-26 08:20:57 +00:00
arrowEq : {X Y : Object} {f g : Arrow X Y} fst f fst g f g
arrowEq {X} {Y} {f} {g} p = λ i p i , lemPropF propEqs p {snd f} {snd g} i
2018-04-13 12:35:54 +00:00
private
isAssociative : IsAssociative
2018-04-26 08:20:57 +00:00
isAssociative {f = f , f0 , f1} {g , g0 , g1} {h , h0 , h1} = arrowEq .isAssociative
2018-03-29 22:12:01 +00:00
isIdentity : IsIdentity identity
2018-04-26 08:20:57 +00:00
isIdentity {AA@(A , a0 , a1)} {BB@(B , b0 , b1)} {f , f0 , f1} = arrowEq .leftIdentity , arrowEq .rightIdentity
arrowsAreSets : ArrowsAreSets
arrowsAreSets {X , x0 , x1} {Y , y0 , y1}
2018-04-26 08:20:57 +00:00
= sigPresSet .arrowsAreSets λ a propSet (propEqs _)
2018-04-05 12:37:25 +00:00
isPreCat : IsPreCategory raw
IsPreCategory.isAssociative isPreCat = isAssociative
IsPreCategory.isIdentity isPreCat = isIdentity
IsPreCategory.arrowsAreSets isPreCat = arrowsAreSets
open IsPreCategory isPreCat
2018-04-03 10:40:20 +00:00
2018-04-13 13:35:56 +00:00
module _ {𝕏 𝕐 : Object} where
2018-04-11 10:27:33 +00:00
open Σ 𝕏 renaming (fst to X ; snd to x)
open Σ x renaming (fst to xa ; snd to xb)
open Σ 𝕐 renaming (fst to Y ; snd to y)
open Σ y renaming (fst to ya ; snd to yb)
open import Cat.Equivalence using (composeIso) renaming (_≅_ to _≅_)
step0
: ((X , xa , xb) (Y , ya , yb))
2018-05-01 09:41:45 +00:00
(Σ[ p (X Y) ] (PathP (λ i .Arrow (p i) 𝒜) xa ya) × (PathP (λ i .Arrow (p i) ) xb yb))
step0
= (λ p cong fst p , cong-d (fst snd) p , cong-d (snd snd) p)
-- , (λ x → λ i → fst x i , (fst (snd x) i) , (snd (snd x) i))
, (λ{ (p , q , r) Σ≡ p λ i q i , r i})
2018-04-11 11:53:33 +00:00
, funExt (λ{ p refl})
, funExt (λ{ (p , q , r) refl})
step1
2018-05-01 09:41:45 +00:00
: (Σ[ p (X Y) ] (PathP (λ i .Arrow (p i) 𝒜) xa ya) × (PathP (λ i .Arrow (p i) ) xb yb))
Σ (X .≊ Y) (λ iso
let p = .isoToId iso
in
2018-05-01 09:41:45 +00:00
( PathP (λ i .Arrow (p i) 𝒜) xa ya)
× PathP (λ i .Arrow (p i) ) xb yb
)
step1
2018-04-13 12:18:28 +00:00
= symIso
2018-04-13 13:22:13 +00:00
(isoSigFst
{A = (X .≊ Y)}
2018-04-13 13:22:13 +00:00
{B = (X Y)}
(.groupoidObject _ _)
2018-05-01 09:41:45 +00:00
{Q = \ p (PathP (λ i .Arrow (p i) 𝒜) xa ya) × (PathP (λ i .Arrow (p i) ) xb yb)}
2018-04-13 13:22:13 +00:00
.isoToId
(symIso (_ , .asTypeIso {X} {Y}) .snd)
)
2018-04-13 12:18:28 +00:00
step2
: Σ (X .≊ Y) (λ iso
let p = .isoToId iso
in
2018-05-01 09:41:45 +00:00
( PathP (λ i .Arrow (p i) 𝒜) xa ya)
× PathP (λ i .Arrow (p i) ) xb yb
)
((X , xa , xb) (Y , ya , yb))
step2
= ( λ{ (iso@(f , f~ , inv-f) , p , q)
2018-04-26 08:20:57 +00:00
( f , sym (.domain-twist-sym iso p) , sym (.domain-twist-sym iso q))
, ( f~ , sym (.domain-twist iso p) , sym (.domain-twist iso q))
, arrowEq (fst inv-f)
, arrowEq (snd inv-f)
}
)
, (λ{ (f , f~ , inv-f , inv-f~)
let
iso : X .≊ Y
iso = fst f , fst f~ , cong fst inv-f , cong fst inv-f~
2018-04-13 07:40:09 +00:00
p : X Y
p = .isoToId iso
2018-05-01 09:41:45 +00:00
pA : .Arrow X 𝒜 .Arrow Y 𝒜
pA = cong (λ x .Arrow x 𝒜) p
pB : .Arrow X .Arrow Y
pB = cong (λ x .Arrow x ) p
k0 = begin
coe pB xb ≡⟨ .coe-dom iso
xb .<<< fst f~ ≡⟨ snd (snd f~)
yb
k1 = begin
coe pA xa ≡⟨ .coe-dom iso
xa .<<< fst f~ ≡⟨ fst (snd f~)
ya
in iso , coe-lem-inv k1 , coe-lem-inv k0})
2018-04-13 13:26:46 +00:00
, funExt (λ x lemSig
(λ x propSig prop0 (λ _ prop1))
_ _
2018-04-13 12:35:54 +00:00
(Σ≡ refl (.propIsomorphism _ _ _)))
2018-04-13 13:26:46 +00:00
, funExt (λ{ (f , _) lemSig propIsomorphism _ _ (Σ≡ refl (propEqs _ _ _))})
where
2018-05-01 09:41:45 +00:00
prop0 : {x} isProp (PathP (λ i .Arrow (.isoToId x i) 𝒜) xa ya)
prop0 {x} = pathJ (λ y p x isProp (PathP (λ i .Arrow (p i) 𝒜) xa x)) (λ x .arrowsAreSets _ _) Y (.isoToId x) ya
prop1 : {x} isProp (PathP (λ i .Arrow (.isoToId x i) ) xb yb)
prop1 {x} = pathJ (λ y p x isProp (PathP (λ i .Arrow (p i) ) xb x)) (λ x .arrowsAreSets _ _) Y (.isoToId x) yb
-- One thing to watch out for here is that the isomorphisms going forwards
-- must compose to give idToIso
iso
: ((X , xa , xb) (Y , ya , yb))
((X , xa , xb) (Y , ya , yb))
iso = step0 step1 step2
where
infixl 5 _⊙_
_⊙_ = composeIso
equiv1
: ((X , xa , xb) (Y , ya , yb))
((X , xa , xb) (Y , ya , yb))
equiv1 = _ , fromIso _ _ (snd iso)
2018-04-11 11:18:34 +00:00
univalent : Univalent
2018-04-13 13:35:56 +00:00
univalent = univalenceFrom≃ equiv1
2018-04-03 10:40:20 +00:00
2018-03-29 22:12:01 +00:00
isCat : IsCategory raw
2018-04-05 12:37:25 +00:00
IsCategory.isPreCategory isCat = isPreCat
IsCategory.univalent isCat = univalent
2018-03-29 22:12:01 +00:00
cat : Category _ _
cat = record
{ raw = raw
; isCategory = isCat
}
open Category cat
2018-05-01 09:41:45 +00:00
lemma : Terminal Product 𝒜
lemma = fromIsomorphism Terminal (Product 𝒜 ) (f , g , inv)
-- C-x 8 RET MATHEMATICAL BOLD SCRIPT CAPITAL A
-- 𝒜
where
2018-05-01 09:41:45 +00:00
f : Terminal Product 𝒜
2018-03-29 22:12:01 +00:00
f ((X , x0 , x1) , uniq) = p
where
2018-05-01 09:41:45 +00:00
rawP : RawProduct 𝒜
2018-03-29 22:12:01 +00:00
rawP = record
{ object = X
2018-04-05 08:41:56 +00:00
; fst = x0
; snd = x1
2018-03-29 22:12:01 +00:00
}
2018-04-05 08:41:56 +00:00
-- open RawProduct rawP renaming (fst to x0 ; snd to x1)
2018-05-01 09:41:45 +00:00
module _ {Y : .Object} (p0 : [ Y , 𝒜 ]) (p1 : [ Y , ]) where
2018-03-29 22:12:01 +00:00
uy : isContr (Arrow (Y , p0 , p1) (X , x0 , x1))
uy = uniq {Y , p0 , p1}
2018-04-05 08:41:56 +00:00
open Σ uy renaming (fst to Y→X ; snd to contractible)
open Σ Y→X renaming (fst to p0×p1 ; snd to cond)
2018-03-29 22:12:01 +00:00
ump : ∃![ f×g ] ( [ x0 f×g ] p0 P.× [ x1 f×g ] p1)
ump = p0×p1 , cond , λ {f} cond-f cong fst (contractible (f , cond-f))
2018-05-01 09:41:45 +00:00
isP : IsProduct 𝒜 rawP
2018-03-29 22:12:01 +00:00
isP = record { ump = ump }
2018-05-01 09:41:45 +00:00
p : Product 𝒜
2018-03-29 22:12:01 +00:00
p = record
{ raw = rawP
; isProduct = isP
}
2018-05-01 09:41:45 +00:00
g : Product 𝒜 Terminal
g p = 𝒳 , t
where
open Product p renaming (object to X ; fst to x₀ ; snd to x₁) using ()
2018-03-29 22:12:01 +00:00
module p = Product p
module isp = IsProduct p.isProduct
𝒳 : Object
𝒳 = X , x₀ , x₁
module _ {𝒴 : Object} where
open Σ 𝒴 renaming (fst to Y)
open Σ (snd 𝒴) renaming (fst to y₀ ; snd to y₁)
ump = p.ump y₀ y₁
open Σ ump renaming (fst to f')
open Σ (snd ump) renaming (fst to f'-cond)
𝒻 : Arrow 𝒴 𝒳
𝒻 = f' , f'-cond
contractible : (f : Arrow 𝒴 𝒳) 𝒻 f
contractible ff@(f , f-cond) = res
2018-03-29 22:12:01 +00:00
where
k : f' f
k = snd (snd ump) f-cond
prp : (a : .Arrow Y X) isProp
( ( [ x₀ a ] y₀)
× ( [ x₁ a ] y₁)
2018-03-29 22:12:01 +00:00
)
prp f f0 f1 = Σ≡
(.arrowsAreSets _ _ (fst f0) (fst f1))
(.arrowsAreSets _ _ (snd f0) (snd f1))
2018-03-29 22:12:01 +00:00
h :
( λ i
[ x₀ k i ] y₀
× [ x₁ k i ] y₁
) [ f'-cond f-cond ]
2018-03-29 22:12:01 +00:00
h = lemPropF prp k
res : (f' , f'-cond) (f , f-cond)
res = Σ≡ k h
t : IsTerminal 𝒳
t {𝒴} = 𝒻 , contractible
2018-03-29 22:12:01 +00:00
ve-re : x g (f x) x
ve-re x = Propositionality.propTerminal _ _
2018-04-03 10:40:20 +00:00
re-ve : p f (g p) p
re-ve p = Product≡ e
where
module p = Product p
-- RawProduct does not have eta-equality.
e : Product.raw (f (g p)) Product.raw p
RawProduct.object (e i) = p.object
2018-04-05 08:41:56 +00:00
RawProduct.fst (e i) = p.fst
RawProduct.snd (e i) = p.snd
2018-03-29 22:12:01 +00:00
inv : AreInverses f g
2018-04-11 11:53:33 +00:00
inv = funExt ve-re , funExt re-ve
2018-03-28 22:07:49 +00:00
2018-05-01 09:41:45 +00:00
propProduct : isProp (Product 𝒜 )
2018-04-03 10:40:20 +00:00
propProduct = equivPreservesNType {n = ⟨-1⟩} lemma Propositionality.propTerminal
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
productEq : x.product y.product
productEq = funExt λ A funExt λ B Try0.propProduct _ _
propHasProducts : isProp (HasProducts )
propHasProducts x y i = record { product = productEq x y i }
fmap≡ : {A : Set} {a0 a1 : A} {B : Set} (f : A B) Path a0 a1 Path (f a0) (f a1)
fmap≡ = cong