aboutsummaryrefslogtreecommitdiff
path: root/Functor/Instance/Decorate.agda
diff options
context:
space:
mode:
Diffstat (limited to 'Functor/Instance/Decorate.agda')
-rw-r--r--Functor/Instance/Decorate.agda29
1 files changed, 13 insertions, 16 deletions
diff --git a/Functor/Instance/Decorate.agda b/Functor/Instance/Decorate.agda
index 30cec87..fedddba 100644
--- a/Functor/Instance/Decorate.agda
+++ b/Functor/Instance/Decorate.agda
@@ -1,4 +1,5 @@
{-# OPTIONS --without-K --safe #-}
+{-# OPTIONS --hidden-argument-puns #-}
open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory)
open import Categories.Functor.Monoidal.Symmetric using (module Lax)
@@ -19,6 +20,7 @@ module Functor.Instance.Decorate
import Categories.Category.Monoidal.Reasoning as βŠ—-Reasoning
import Categories.Diagram.Pushout as DiagramPushout
import Categories.Morphism.Reasoning as β‡’-Reasoning
+import Category.Diagram.Cospan π’ž as Cospan
open import Categories.Category using (Category; _[_,_]; _[_β‰ˆ_]; _[_∘_])
open import Categories.Category.Cocartesian using (module CocartesianMonoidal)
@@ -27,10 +29,10 @@ open import Categories.Category.Monoidal.Utilities using (module Shorthands)
open import Categories.Functor.Core using (Functor)
open import Categories.Functor.Properties using ([_]-resp-β‰…)
open import Function.Base using () renaming (id to idf)
-open import Category.Instance.Cospans π’ž using (Cospan; Cospans; Same)
+open import Category.Instance.Cospans π’ž using (Cospans)
open import Category.Instance.DecoratedCospans π’ž F using (DecoratedCospans)
-open import Functor.Instance.Cospan.Stack using (βŠ—)
-open import Functor.Instance.DecoratedCospan.Stack using () renaming (βŠ— to βŠ—β€²)
+open import Functor.Instance.Cospan.Stack π’ž using (module βŠ—)
+open import Functor.Instance.DecoratedCospan.Stack π’ž F using () renaming (module βŠ— to βŠ—β€²)
module π’ž = FinitelyCocompleteCategory π’ž
module π’Ÿ = SymmetricMonoidalCategory π’Ÿ
@@ -67,14 +69,14 @@ identity = record
open β‡’-Reasoning π’Ÿ.U
homomorphism : DecoratedCospans [ decorate (Cospans [ g ∘ f ]) β‰ˆ DecoratedCospans [ decorate g ∘ decorate f ] ]
-homomorphism {B} {C} {g} {A} {f} = record
+homomorphism {g} {f} = record
{ cospans-β‰ˆ = Cospans.Equiv.refl
; same-deco = same-deco
}
where
- open Cospan f using (N; fβ‚‚)
- open Cospan g using () renaming (N to M; f₁ to g₁)
+ open Cospan.Cospan f using (N; fβ‚‚)
+ open Cospan.Cospan g using () renaming (N to M; f₁ to g₁)
open π’Ÿ using (U; monoidal; _βŠ—β‚_; unitorΛ‘-commute-from) renaming (module unitorΛ‘ to Ξ»-)
open π’ž using (Β‘; βŠ₯; Β‘-unique; pushout) renaming ([_,_] to [_,_]β€²; _+₁_ to infixr 10 _+₁_ )
@@ -99,7 +101,7 @@ homomorphism {B} {C} {g} {A} {f} = record
F₁ ([ i₁ , iβ‚‚ ]β€² π’ž.∘ Β‘ +₁ Β‘ π’ž.∘ λ⇐′) ∘ Ξ΅ ∘ Ξ»β‡’ ∘ λ⇐ β‰ˆβŸ¨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ coherence-inv₃ monoidal ⟩
F₁ ([ i₁ , iβ‚‚ ]β€² π’ž.∘ Β‘ +₁ Β‘ π’ž.∘ λ⇐′) ∘ Ξ΅ ∘ Ξ»β‡’ ∘ ρ⇐ β‰ˆβŸ¨ refl⟩∘⟨ extendΚ³ unitorΛ‘-commute-from ⟨
F₁ ([ i₁ , iβ‚‚ ]β€² π’ž.∘ Β‘ +₁ Β‘ π’ž.∘ λ⇐′) ∘ Ξ»β‡’ ∘ id βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ pushΛ‘ F.homomorphism ⟩
- F₁ [ i₁ , iβ‚‚ ]β€² ∘ F₁ (Β‘ +₁ Β‘ π’ž.∘ λ⇐′) ∘ Ξ»β‡’ ∘ id βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ push-center (sym F.homomorphism) ⟩
+ F₁ [ i₁ , iβ‚‚ ]β€² ∘ F₁ (Β‘ +₁ Β‘ π’ž.∘ λ⇐′) ∘ Ξ»β‡’ ∘ id βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ push-center F.homomorphism ⟩
F₁ [ i₁ , iβ‚‚ ]β€² ∘ F₁ (Β‘ +₁ Β‘) ∘ F₁ λ⇐′ ∘ Ξ»β‡’ ∘ id βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ refl⟩∘⟨ refl⟩∘⟨ extendΚ³ (switch-fromtoΛ‘ ([ F.F ]-resp-β‰… unitorΛ‘) F.unitaryΛ‘) ⟨
F₁ [ i₁ , iβ‚‚ ]β€² ∘ F₁ (Β‘ +₁ Β‘) ∘ Ο† (βŠ₯ , βŠ₯) ∘ Ξ΅ βŠ—β‚ id ∘ id βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pullΛ‘ (sym serialize₁₂) ⟩
F₁ [ i₁ , iβ‚‚ ]β€² ∘ F₁ (Β‘ +₁ Β‘) ∘ Ο† (βŠ₯ , βŠ₯) ∘ Ξ΅ βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ refl⟩∘⟨ extendΚ³ (Ο†-commute (Β‘ , Β‘)) ⟨
@@ -126,19 +128,15 @@ Decorate = record
; F-resp-β‰ˆ = F-resp-β‰ˆ
}
-module βŠ— = Functor (βŠ— π’ž)
-module βŠ—β€² = Functor (βŠ—β€² π’ž F)
-open π’ž using (_+₁_)
-
Decorate-resp-βŠ— : DecoratedCospans [ decorate (βŠ—.₁ (f , g)) β‰ˆ βŠ—β€².₁ (decorate f , decorate g) ]
-Decorate-resp-βŠ— {f = f} {g = g} = record
- { cospans-β‰ˆ = Cospans.Equiv.refl
+Decorate-resp-βŠ— {f} {g} = record
+ { cospans-β‰ˆ = Cospan.β‰ˆ-refl
; same-deco = same-deco
}
where
- open Cospan f using (N)
- open Cospan g using () renaming (N to M)
+ open Cospan.Cospan f using (N)
+ open Cospan.Cospan g using () renaming (N to M)
open π’Ÿ using (U; monoidal; _βŠ—β‚_; unitorΛ‘-commute-from) renaming (module unitorΛ‘ to Ξ»-)
open π’ž using (Β‘; βŠ₯; Β‘-unique; pushout) renaming ([_,_] to [_,_]β€²; _+₁_ to infixr 10 _+₁_ )
@@ -165,4 +163,3 @@ Decorate-resp-βŠ— {f = f} {g = g} = record
F₁ (Β‘ +₁ Β‘) ∘ Ο† (βŠ₯ , βŠ₯) ∘ Ξ΅ βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ extendΚ³ (Ο†-commute (Β‘ , Β‘)) ⟨
Ο† (N , M) ∘ F₁ Β‘ βŠ—β‚ F₁ Β‘ ∘ Ξ΅ βŠ—β‚ Ξ΅ ∘ ρ⇐ β‰ˆβŸ¨ refl⟩∘⟨ pullΛ‘ (sym βŠ—-distrib-over-∘) ⟩
Ο† (N , M) ∘ (F₁ Β‘ ∘ Ξ΅) βŠ—β‚ (F₁ Β‘ ∘ Ξ΅) ∘ ρ⇐ ∎
-