Skip to content

Commit

Permalink
refactor use case
Browse files Browse the repository at this point in the history
  • Loading branch information
felixwellen committed Feb 20, 2024
1 parent 2a54073 commit 3ee4974
Show file tree
Hide file tree
Showing 3 changed files with 102 additions and 102 deletions.
85 changes: 6 additions & 79 deletions Cubical/HITs/SmashProduct/SymmetricMonoidal.agda
Original file line number Diff line number Diff line change
@@ -1,42 +1,23 @@
{-# OPTIONS --safe #-}

{-
This file contians a proof that the smash product turns the universe
of pointed types into a symmetric monoidal precategory. The pentagon
and hexagon are proved in separate files due to the length of the
proofs. The remaining identities and the main result are proved here.
-}

module Cubical.HITs.SmashProduct.SymmetricMonoidal where

open import Cubical.HITs.SmashProduct.Base
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Pointed
open import Cubical.Foundations.Isomorphism
open import Cubical.HITs.Pushout.Base
open import Cubical.Data.Unit
open import Cubical.Data.Sigma
open import Cubical.HITs.Wedge
open import Cubical.Foundations.GroupoidLaws
open import Cubical.Foundations.Pointed.Homogeneous
open import Cubical.Foundations.Path
open import Cubical.Foundations.Function
open import Cubical.Foundations.Equiv

open import Cubical.Data.Unit
open import Cubical.Data.Sigma
open import Cubical.Data.Bool
open import Cubical.HITs.Wedge
open import Cubical.HITs.Pushout.Base
open import Cubical.HITs.SmashProduct.Base
open import Cubical.HITs.SmashProduct.Pentagon
open import Cubical.HITs.SmashProduct.Hexagon
open import Cubical.Categories.Category.Precategory
renaming (_×_ to _×'_)
open import Cubical.Data.Bool

open Precategory hiding (_∘_)
open Prefunctor
open isSymmetricPrecategory
open isMonoidalPrecategory
open PreNatIso
open PreNatTrans
open preIsIso


private
variable
Expand Down Expand Up @@ -527,57 +508,3 @@ snd ⋀lIdEquiv∙ = refl
cong (push (inl (snd A)) ∙_)
(sym (rUnit _) ∙ (λ i j push (push tt (~ i)) (~ j)))
∙ rCancel _


-- ⋀ as a functor
⋀F : {ℓ} Prefunctor (PointedCat ℓ ×' PointedCat ℓ) (PointedCat ℓ)
F-ob ⋀F (A , B) = A ⋀∙ B
F-hom ⋀F (f , g) = f ⋀→∙ g
F-id ⋀F = ⋀→∙-idfun
F-seq ⋀F (f , g) (f' , g') = ⋀→∙-comp f f' g g'

⋀lUnitNatIso : PreNatIso (PointedCat ℓ) (PointedCat ℓ)
(restrFunctorₗ ⋀F Bool*∙) (idPrefunctor (PointedCat ℓ))
N-ob (trans ⋀lUnitNatIso) X = ≃∙map ⋀lIdEquiv∙
N-hom (trans ⋀lUnitNatIso) f = ⋀lId-sq f
inv' (isIs ⋀lUnitNatIso c) = ≃∙map (invEquiv∙ ⋀lIdEquiv∙)
sect (isIs (⋀lUnitNatIso {ℓ = ℓ}) c) =
≃∙→ret/sec∙ (⋀lIdEquiv∙ {ℓ = ℓ} {A = c}) .snd
retr (isIs ⋀lUnitNatIso c) =
≃∙→ret/sec∙ ⋀lIdEquiv∙ .fst

makeIsIso-Pointed : {ℓ} {A B : Pointed ℓ} {f : A →∙ B}
isEquiv (fst f) preIsIso {C = PointedCat ℓ} f
inv' (makeIsIso-Pointed {f = f} eq) = ≃∙map (invEquiv∙ ((fst f , eq) , snd f))
sect (makeIsIso-Pointed {f = f} eq) = ≃∙→ret/sec∙ ((fst f , eq) , snd f) .snd
retr (makeIsIso-Pointed {f = f} eq) = ≃∙→ret/sec∙ ((fst f , eq) , snd f) .fst

restrₗᵣ : PreNatIso (PointedCat ℓ) (PointedCat ℓ)
(restrFunctorᵣ ⋀F Bool*∙) (restrFunctorₗ ⋀F Bool*∙)
N-ob (trans restrₗᵣ) X = ⋀comm→∙
N-hom (trans restrₗᵣ) f = ⋀comm-sq f (idfun∙ Bool*∙)
isIs restrₗᵣ c = makeIsIso-Pointed (isoToIsEquiv ⋀CommIso)

-- main result
⋀Symm : {ℓ} isSymmetricPrecategory (PointedCat ℓ)
_⊗_ (isMonoidal ⋀Symm) = ⋀F
𝟙 (isMonoidal ⋀Symm) = Bool*∙
N-ob (trans (⊗assoc (isMonoidal ⋀Symm))) (A , B , C) = ≃∙map SmashAssocEquiv∙
N-hom (trans (⊗assoc (isMonoidal ⋀Symm))) (f , g , h) = ⋀assoc-⋀→∙ f g h
inv' (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙map (invEquiv∙ SmashAssocEquiv∙)
sect (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙→ret/sec∙ SmashAssocEquiv∙ .snd
retr (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙→ret/sec∙ SmashAssocEquiv∙ .fst
⊗lUnit (isMonoidal ⋀Symm) = ⋀lUnitNatIso
⊗rUnit (isMonoidal ⋀Symm) = compPreNatIso _ _ _ restrₗᵣ ⋀lUnitNatIso
triang (isMonoidal (⋀Symm {ℓ})) X Y = ⋀triang
⊗pentagon (isMonoidal ⋀Symm) X Y Z W =
(∘∙-assoc assc₅∙ assc₄∙ assc₃∙) ∙ pentagon∙
N-ob (trans (Braid ⋀Symm)) X = ⋀comm→∙
N-hom (trans (Braid ⋀Symm)) (f , g) = ⋀comm-sq f g
isIs (Braid ⋀Symm) _ = makeIsIso-Pointed (isoToIsEquiv ⋀CommIso)
isSymmetricPrecategory.hexagon ⋀Symm a b c = hexagon∙
symBraiding ⋀Symm X Y =
ΣPathP ((funExt (Iso.rightInv ⋀CommIso)) , (sym (rUnit refl)))
96 changes: 96 additions & 0 deletions Cubical/HITs/SmashProduct/SymmetricMonoidalCat.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,96 @@
{-# OPTIONS --safe #-}

{-
This file contians a proof that the smash product turns the universe
of pointed types into a symmetric monoidal wild category. The pentagon
and hexagon are proved in separate files due to the length of the
proofs. The remaining identities and the main result are proved here.
-}

module Cubical.HITs.SmashProduct.SymmetricMonoidalCat where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Pointed
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.GroupoidLaws
open import Cubical.Foundations.Path
open import Cubical.Foundations.Equiv

open import Cubical.Data.Unit
open import Cubical.Data.Sigma using (ΣPathP)
open import Cubical.Data.Bool
open import Cubical.HITs.SmashProduct.Base
open import Cubical.HITs.SmashProduct.Pentagon
open import Cubical.HITs.SmashProduct.Hexagon
open import Cubical.HITs.SmashProduct.SymmetricMonoidal

open import Cubical.WildCat.Base
open import Cubical.WildCat.Functor
open import Cubical.WildCat.Product
open import Cubical.WildCat.BraidedSymmetricMonoidal
open import Cubical.WildCat.Instances.Pointed

open WildCat
open WildFunctor
open isSymmetricWildCat
open isMonoidalWildCat
open WildNatIso
open WildNatTrans
open wildIsIso

private
variable
ℓ ℓ' : Level

-- ⋀ as a functor
⋀F : {ℓ} WildFunctor (PointedCat ℓ × PointedCat ℓ) (PointedCat ℓ)
F-ob ⋀F (A , B) = A ⋀∙ B
F-hom ⋀F (f , g) = f ⋀→∙ g
F-id ⋀F = ⋀→∙-idfun
F-seq ⋀F (f , g) (f' , g') = ⋀→∙-comp f f' g g'

⋀lUnitNatIso : WildNatIso (PointedCat ℓ) (PointedCat ℓ)
(restrFunctorₗ ⋀F Bool*∙) (idWildFunctor (PointedCat ℓ))
N-ob (trans ⋀lUnitNatIso) X = ≃∙map ⋀lIdEquiv∙
N-hom (trans ⋀lUnitNatIso) f = ⋀lId-sq f
inv' (isIs ⋀lUnitNatIso c) = ≃∙map (invEquiv∙ ⋀lIdEquiv∙)
sect (isIs (⋀lUnitNatIso {ℓ = ℓ}) c) =
≃∙→ret/sec∙ (⋀lIdEquiv∙ {ℓ = ℓ} {A = c}) .snd
retr (isIs ⋀lUnitNatIso c) =
≃∙→ret/sec∙ ⋀lIdEquiv∙ .fst

makeIsIso-Pointed : {ℓ} {A B : Pointed ℓ} {f : A →∙ B}
isEquiv (fst f) wildIsIso {C = PointedCat ℓ} f
inv' (makeIsIso-Pointed {f = f} eq) = ≃∙map (invEquiv∙ ((fst f , eq) , snd f))
sect (makeIsIso-Pointed {f = f} eq) = ≃∙→ret/sec∙ ((fst f , eq) , snd f) .snd
retr (makeIsIso-Pointed {f = f} eq) = ≃∙→ret/sec∙ ((fst f , eq) , snd f) .fst

restrₗᵣ : WildNatIso (PointedCat ℓ) (PointedCat ℓ)
(restrFunctorᵣ ⋀F Bool*∙) (restrFunctorₗ ⋀F Bool*∙)
N-ob (trans restrₗᵣ) X = ⋀comm→∙
N-hom (trans restrₗᵣ) f = ⋀comm-sq f (idfun∙ Bool*∙)
isIs restrₗᵣ c = makeIsIso-Pointed (isoToIsEquiv ⋀CommIso)

-- main result
⋀Symm : {ℓ} isSymmetricWildCat (PointedCat ℓ)
_⊗_ (isMonoidal ⋀Symm) = ⋀F
𝟙 (isMonoidal ⋀Symm) = Bool*∙
N-ob (trans (⊗assoc (isMonoidal ⋀Symm))) (A , B , C) = ≃∙map SmashAssocEquiv∙
N-hom (trans (⊗assoc (isMonoidal ⋀Symm))) (f , g , h) = ⋀assoc-⋀→∙ f g h
inv' (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙map (invEquiv∙ SmashAssocEquiv∙)
sect (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙→ret/sec∙ SmashAssocEquiv∙ .snd
retr (isIs (⊗assoc (isMonoidal ⋀Symm)) (A , B , C)) =
≃∙→ret/sec∙ SmashAssocEquiv∙ .fst
⊗lUnit (isMonoidal ⋀Symm) = ⋀lUnitNatIso
⊗rUnit (isMonoidal ⋀Symm) = compWildNatIso _ _ _ restrₗᵣ ⋀lUnitNatIso
triang (isMonoidal (⋀Symm {ℓ})) X Y = ⋀triang
⊗pentagon (isMonoidal ⋀Symm) X Y Z W =
(∘∙-assoc assc₅∙ assc₄∙ assc₃∙) ∙ pentagon∙
N-ob (trans (Braid ⋀Symm)) X = ⋀comm→∙
N-hom (trans (Braid ⋀Symm)) (f , g) = ⋀comm-sq f g
isIs (Braid ⋀Symm) _ = makeIsIso-Pointed (isoToIsEquiv ⋀CommIso)
isSymmetricWildCat.hexagon ⋀Symm a b c = hexagon∙
symBraiding ⋀Symm X Y =
ΣPathP ((funExt (Iso.rightInv ⋀CommIso)) , (sym (rUnit refl)))
23 changes: 0 additions & 23 deletions Cubical/WildCat/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,6 @@
module Cubical.WildCat.Base where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv

open import Cubical.Data.Sigma renaming (_×_ to _×'_)

Expand Down Expand Up @@ -68,28 +67,6 @@ idIso {C = C} = wildiso (C .id ) (C .id) (C .⋆IdL (C .id)) (C .⋆IdL (C .id))
pathToIso : {C : WildCat ℓ ℓ'} {x y : C .ob} (p : x ≡ y) WildCatIso C x y
pathToIso {C = C} p = J (λ z _ WildCatIso C _ z) idIso p

-- Equivalences in wild categories (analogous to HoTT-terminology for maps between types)
record WildCatEquiv (C : WildCat ℓ ℓ') (x y : C .ob) : Type ℓ' where
no-eta-equality
constructor wildequiv
field
mor : C [ x , y ]
linv : C [ y , x ]
rinv : C [ y , x ]
isLinv : linv ⋆⟨ C ⟩ mor ≡ C .id
isRinv : mor ⋆⟨ C ⟩ rinv ≡ C .id

idWildEquiv : {C : WildCat ℓ ℓ'} {x : C .ob} WildCatEquiv C x x
idWildEquiv {C = C} = wildequiv (C .id ) (C .id ) (C .id) (C .⋆IdL (C .id)) (C .⋆IdL (C .id))

pathToEquiv : {C : WildCat ℓ ℓ'} {x y : C .ob} (p : x ≡ y) WildCatEquiv C x y
pathToEquiv {C = C} p = J (λ z _ WildCatEquiv C _ z) idWildEquiv p

-- Univalent Categories
record isUnivalent (C : WildCat ℓ ℓ') : Type (ℓ-max ℓ ℓ') where
field
univ : (x y : C .ob) isEquiv (pathToEquiv {C = C} {x = x} {y = y})

-- Natural isomorphisms
module _ {C : WildCat ℓ ℓ'}
{x y : C .ob} (f : Hom[_,_] C x y) where
Expand Down

0 comments on commit 3ee4974

Please sign in to comment.