aboutsummaryrefslogtreecommitdiff
path: root/Functor/Instance
diff options
context:
space:
mode:
Diffstat (limited to 'Functor/Instance')
-rw-r--r--Functor/Instance/Cospan/Embed.agda185
-rw-r--r--Functor/Instance/Cospan/Stack.agda144
2 files changed, 329 insertions, 0 deletions
diff --git a/Functor/Instance/Cospan/Embed.agda b/Functor/Instance/Cospan/Embed.agda
new file mode 100644
index 0000000..77f0361
--- /dev/null
+++ b/Functor/Instance/Cospan/Embed.agda
@@ -0,0 +1,185 @@
+{-# OPTIONS --without-K --safe #-}
+
+open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory)
+
+module Functor.Instance.Cospan.Embed {o ℓ e} (𝒞 : FinitelyCocompleteCategory o ℓ e) where
+
+import Categories.Diagram.Pushout as DiagramPushout
+import Categories.Diagram.Pushout.Properties as PushoutProperties
+import Categories.Morphism as Morphism
+import Categories.Morphism.Reasoning as ⇒-Reasoning
+import Category.Diagram.Pushout as Pushout′
+
+open import Categories.Category using (_[_,_]; _[_∘_]; _[_≈_])
+open import Categories.Category.Core using (Category)
+open import Categories.Functor.Core using (Functor)
+open import Category.Instance.Cospans 𝒞 using (Cospans)
+open import Data.Product.Base using (_,_)
+open import Function.Base using (id)
+open import Functor.Instance.Cospan.Stack using (⊗)
+
+module 𝒞 = FinitelyCocompleteCategory 𝒞
+module Cospans = Category Cospans
+
+open 𝒞 using (U; pushout; _+₁_)
+open Cospans using (_≈_)
+open DiagramPushout U using (Pushout)
+open Morphism U using (module ≅; _≅_)
+open PushoutProperties U using (up-to-iso)
+open Pushout′ U using (pushout-id-g; pushout-f-id)
+
+L₁ : {A B : 𝒞.Obj} → U [ A , B ] → Cospans [ A , B ]
+L₁ f = record
+ { f₁ = f
+ ; f₂ = 𝒞.id
+ }
+
+L-identity : {A : 𝒞.Obj} → L₁ 𝒞.id ≈ Cospans.id {A}
+L-identity = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = 𝒞.identity²
+ ; from∘f₂≈f₂′ = 𝒞.identity²
+ }
+
+L-homomorphism : {X Y Z : 𝒞.Obj} {f : U [ X , Y ]} {g : U [ Y , Z ]} → L₁ (U [ g ∘ f ]) ≈ Cospans [ L₁ g ∘ L₁ f ]
+L-homomorphism {X} {Y} {Z} {f} {g} = record
+ { ≅N = up-to-iso P′ P
+ ; from∘f₁≈f₁′ = pullˡ (P′.universal∘i₁≈h₁ {eq = P.commute})
+ ; from∘f₂≈f₂′ = P′.universal∘i₂≈h₂ {eq = P.commute} ○ sym 𝒞.identityʳ
+ }
+ where
+ open ⇒-Reasoning U
+ open 𝒞.HomReasoning
+ open 𝒞.Equiv
+ P P′ : Pushout 𝒞.id g
+ P = pushout 𝒞.id g
+ P′ = pushout-id-g
+ module P = Pushout P
+ module P′ = Pushout P′
+
+L-resp-≈ : {A B : 𝒞.Obj} {f g : U [ A , B ]} → U [ f ≈ g ] → Cospans [ L₁ f ≈ L₁ g ]
+L-resp-≈ {A} {B} {f} {g} f≈g = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = 𝒞.identityˡ ○ f≈g
+ ; from∘f₂≈f₂′ = 𝒞.identity²
+ }
+ where
+ open 𝒞.HomReasoning
+
+L : Functor U Cospans
+L = record
+ { F₀ = id
+ ; F₁ = L₁
+ ; identity = L-identity
+ ; homomorphism = L-homomorphism
+ ; F-resp-≈ = L-resp-≈
+ }
+
+R₁ : {A B : 𝒞.Obj} → U [ B , A ] → Cospans [ A , B ]
+R₁ g = record
+ { f₁ = 𝒞.id
+ ; f₂ = g
+ }
+
+R-identity : {A : 𝒞.Obj} → R₁ 𝒞.id ≈ Cospans.id {A}
+R-identity = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = 𝒞.identity²
+ ; from∘f₂≈f₂′ = 𝒞.identity²
+ }
+
+R-homomorphism : {X Y Z : 𝒞.Obj} {f : U [ Y , X ]} {g : U [ Z , Y ]} → R₁ (U [ f ∘ g ]) ≈ Cospans [ R₁ g ∘ R₁ f ]
+R-homomorphism {X} {Y} {Z} {f} {g} = record
+ { ≅N = up-to-iso P′ P
+ ; from∘f₁≈f₁′ = P′.universal∘i₁≈h₁ {eq = P.commute} ○ sym 𝒞.identityʳ
+ ; from∘f₂≈f₂′ = pullˡ (P′.universal∘i₂≈h₂ {eq = P.commute})
+ }
+ where
+ open ⇒-Reasoning U
+ open 𝒞.HomReasoning
+ open 𝒞.Equiv
+ P P′ : Pushout f 𝒞.id
+ P = pushout f 𝒞.id
+ P′ = pushout-f-id
+ module P = Pushout P
+ module P′ = Pushout P′
+
+R-resp-≈ : {A B : 𝒞.Obj} {f g : U [ A , B ]} → U [ f ≈ g ] → Cospans [ R₁ f ≈ R₁ g ]
+R-resp-≈ {A} {B} {f} {g} f≈g = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = 𝒞.identity²
+ ; from∘f₂≈f₂′ = 𝒞.identityˡ ○ f≈g
+ }
+ where
+ open 𝒞.HomReasoning
+
+R : Functor 𝒞.op Cospans
+R = record
+ { F₀ = id
+ ; F₁ = R₁
+ ; identity = R-identity
+ ; homomorphism = R-homomorphism
+ ; F-resp-≈ = R-resp-≈
+ }
+
+B₁ : {A B C : 𝒞.Obj} → U [ A , C ] → U [ B , C ] → Cospans [ A , B ]
+B₁ f g = record
+ { f₁ = f
+ ; f₂ = g
+ }
+
+B∘L : {W X Y Z : 𝒞.Obj} {f : U [ W , X ]} {g : U [ X , Y ]} {h : U [ Z , Y ]} → Cospans [ B₁ g h ∘ L₁ f ] ≈ B₁ (U [ g ∘ f ]) h
+B∘L {W} {X} {Y} {Z} {f} {g} {h} = record
+ { ≅N = up-to-iso P P′
+ ; from∘f₁≈f₁′ = pullˡ (P.universal∘i₁≈h₁ {eq = P′.commute})
+ ; from∘f₂≈f₂′ = pullˡ (P.universal∘i₂≈h₂ {eq = P′.commute}) ○ 𝒞.identityˡ
+ }
+ where
+ open ⇒-Reasoning U
+ open 𝒞.HomReasoning
+ open 𝒞.Equiv
+ P P′ : Pushout 𝒞.id g
+ P = pushout 𝒞.id g
+ P′ = pushout-id-g
+ module P = Pushout P
+ module P′ = Pushout P′
+
+R∘B : {W X Y Z : 𝒞.Obj} {f : U [ W , X ]} {g : U [ Y , X ]} {h : U [ Z , Y ]} → Cospans [ R₁ h ∘ B₁ f g ] ≈ B₁ f (U [ g ∘ h ])
+R∘B {W} {X} {Y} {Z} {f} {g} {h} = record
+ { ≅N = up-to-iso P P′
+ ; from∘f₁≈f₁′ = pullˡ (P.universal∘i₁≈h₁ {eq = P′.commute}) ○ 𝒞.identityˡ
+ ; from∘f₂≈f₂′ = pullˡ (P.universal∘i₂≈h₂ {eq = P′.commute})
+ }
+ where
+ open ⇒-Reasoning U
+ open 𝒞.HomReasoning
+ open 𝒞.Equiv
+ P P′ : Pushout g 𝒞.id
+ P = pushout g 𝒞.id
+ P′ = pushout-f-id
+ module P = Pushout P
+ module P′ = Pushout P′
+
+module _ where
+
+ open _≅_
+
+ ≅-L-R : ∀ {X Y : 𝒞.Obj} (X≅Y : X ≅ Y) → L₁ (to X≅Y) ≈ R₁ (from X≅Y)
+ ≅-L-R {X} {Y} X≅Y = record
+ { ≅N = X≅Y
+ ; from∘f₁≈f₁′ = isoʳ X≅Y
+ ; from∘f₂≈f₂′ = 𝒞.identityʳ
+ }
+
+module ⊗ = Functor (⊗ 𝒞)
+
+L-resp-⊗ : {X Y X′ Y′ : 𝒞.Obj} {a : U [ X , X′ ]} {b : U [ Y , Y′ ]} → L₁ (a +₁ b) ≈ ⊗.₁ (L₁ a , L₁ b)
+L-resp-⊗ {X} {Y} {X′} {Y′} {a} {b} = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = 𝒞.identityˡ
+ ; from∘f₂≈f₂′ = 𝒞.identityˡ ○ sym +-η ○ sym ([]-cong₂ identityʳ identityʳ)
+ }
+ where
+ open 𝒞.HomReasoning
+ open 𝒞.Equiv
+ open 𝒞 using (+-η; []-cong₂; identityˡ; identityʳ)
diff --git a/Functor/Instance/Cospan/Stack.agda b/Functor/Instance/Cospan/Stack.agda
new file mode 100644
index 0000000..b7664dc
--- /dev/null
+++ b/Functor/Instance/Cospan/Stack.agda
@@ -0,0 +1,144 @@
+{-# OPTIONS --without-K --safe #-}
+
+open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory)
+
+module Functor.Instance.Cospan.Stack {o ℓ e} (𝒞 : FinitelyCocompleteCategory o ℓ e) where
+
+import Categories.Diagram.Pushout as DiagramPushout
+import Categories.Diagram.Pushout.Properties as PushoutProperties
+import Categories.Morphism as Morphism
+import Categories.Morphism.Reasoning as ⇒-Reasoning
+
+open import Categories.Category.Core using (Category)
+open import Categories.Functor.Bifunctor using (Bifunctor)
+open import Category.Instance.Cospans 𝒞 using (Cospan; Cospans; Same; id-Cospan; compose)
+open import Category.Instance.FinitelyCocompletes {o} {ℓ} {e} using () renaming (_×_ to _×′_)
+open import Category.Instance.Properties.FinitelyCocompletes {o} {ℓ} {e} using (-+-; FinitelyCocompletes-CC)
+open import Data.Product.Base using (Σ; _,_; _×_; proj₁; proj₂)
+open import Functor.Exact using (RightExactFunctor; IsPushout⇒Pushout)
+open import Level using (Level; _⊔_; suc)
+
+module 𝒞 = FinitelyCocompleteCategory 𝒞
+module Cospans = Category Cospans
+
+open 𝒞 using (U; _+_; _+₁_; pushout; coproduct; [_,_]; ⊥; cocartesianCategory; monoidal)
+open Category U
+open DiagramPushout U using (Pushout)
+open PushoutProperties U using (up-to-iso)
+
+module 𝒞×𝒞 = FinitelyCocompleteCategory (𝒞 ×′ 𝒞)
+open 𝒞×𝒞 using () renaming (pushout to pushout′; U to U×U)
+open DiagramPushout U×U using () renaming (Pushout to Pushout′)
+
+open import Categories.Category.Monoidal.Utilities monoidal using (_⊗ᵢ_)
+
+together : {A A′ B B′ : Obj} → Cospan A B → Cospan A′ B′ → Cospan (A + A′) (B + B′)
+together A⇒B A⇒B′ = record
+ { f₁ = f₁ A⇒B +₁ f₁ A⇒B′
+ ; f₂ = f₂ A⇒B +₁ f₂ A⇒B′
+ }
+ where
+ open Cospan
+
+id⊗id≈id : {A B : Obj} → Same (together (id-Cospan {A}) (id-Cospan {B})) (id-Cospan {A + B})
+id⊗id≈id {A} {B} = record
+ { ≅N = ≅.refl
+ ; from∘f₁≈f₁′ = from∘f≈f′
+ ; from∘f₂≈f₂′ = from∘f≈f′
+ }
+ where
+ open Morphism U using (module ≅)
+ open HomReasoning
+ open 𝒞 using (+-η; []-cong₂)
+ open coproduct {A} {B} using (i₁; i₂)
+ from∘f≈f′ : id ∘ [ i₁ ∘ id , i₂ ∘ id ] 𝒞.≈ id
+ from∘f≈f′ = begin
+ id ∘ [ i₁ ∘ id , i₂ ∘ id ] ≈⟨ identityˡ ⟩
+ [ i₁ ∘ id , i₂ ∘ id ] ≈⟨ []-cong₂ identityʳ identityʳ ⟩
+ [ i₁ , i₂ ] ≈⟨ +-η ⟩
+ id ∎
+
+homomorphism
+ : {A A′ B B′ C C′ : Obj}
+ → (A⇒B : Cospan A B)
+ → (B⇒C : Cospan B C)
+ → (A⇒B′ : Cospan A′ B′)
+ → (B⇒C′ : Cospan B′ C′)
+ → Same (together (compose A⇒B B⇒C) (compose A⇒B′ B⇒C′)) (compose (together A⇒B A⇒B′) (together B⇒C B⇒C′) )
+homomorphism A⇒B B⇒C A⇒B′ B⇒C′ = record
+ { ≅N = ≅N
+ ; from∘f₁≈f₁′ = from∘f₁≈f₁′
+ ; from∘f₂≈f₂′ = from∘f₂≈f₂′
+ }
+ where
+ open Cospan
+ open Pushout
+ open HomReasoning
+ open ⇒-Reasoning U
+ open Morphism U using (_≅_)
+ open _≅_
+ open 𝒞 using (+₁∘+₁)
+ module -+- = RightExactFunctor (-+- {𝒞})
+ P₁ = pushout (f₂ A⇒B) (f₁ B⇒C)
+ P₂ = pushout (f₂ A⇒B′) (f₁ B⇒C′)
+ module P₁ = Pushout P₁
+ module P₂ = Pushout P₂
+ P₁×P₂ = pushout′ (f₂ A⇒B , f₂ A⇒B′) (f₁ B⇒C , f₁ B⇒C′)
+ module P₁×P₂ = Pushout′ P₁×P₂
+ P₃ = pushout (f₂ A⇒B +₁ f₂ A⇒B′) (f₁ B⇒C +₁ f₁ B⇒C′)
+ P₃′ = IsPushout⇒Pushout (-+-.F-resp-pushout P₁×P₂.isPushout)
+ ≅N : Q P₃′ ≅ Q P₃
+ ≅N = up-to-iso P₃′ P₃
+ from∘f₁≈f₁′ : from ≅N ∘ (f₁ (compose A⇒B B⇒C) +₁ f₁ (compose A⇒B′ B⇒C′)) ≈ f₁ (compose (together A⇒B A⇒B′) (together B⇒C B⇒C′))
+ from∘f₁≈f₁′ = begin
+ from ≅N ∘ (f₁ (compose A⇒B B⇒C) +₁ f₁ (compose A⇒B′ B⇒C′)) ≈⟨ Equiv.refl ⟩
+ from ≅N ∘ ((i₁ P₁ ∘ f₁ A⇒B) +₁ (i₁ P₂ ∘ f₁ A⇒B′)) ≈⟨ refl⟩∘⟨ +₁∘+₁ ⟨
+ from ≅N ∘ (i₁ P₁ +₁ i₁ P₂) ∘ (f₁ A⇒B +₁ f₁ A⇒B′) ≈⟨ Equiv.refl ⟩
+ from ≅N ∘ i₁ P₃′ ∘ f₁ (together A⇒B A⇒B′) ≈⟨ pullˡ (universal∘i₁≈h₁ P₃′) ⟩
+ i₁ P₃ ∘ f₁ (together A⇒B A⇒B′) ∎
+ from∘f₂≈f₂′ : from ≅N ∘ (f₂ (compose A⇒B B⇒C) +₁ f₂ (compose A⇒B′ B⇒C′)) ≈ f₂ (compose (together A⇒B A⇒B′) (together B⇒C B⇒C′))
+ from∘f₂≈f₂′ = begin
+ from ≅N ∘ (f₂ (compose A⇒B B⇒C) +₁ f₂ (compose A⇒B′ B⇒C′)) ≈⟨ Equiv.refl ⟩
+ from ≅N ∘ ((i₂ P₁ ∘ f₂ B⇒C) +₁ (i₂ P₂ ∘ f₂ B⇒C′)) ≈⟨ refl⟩∘⟨ +₁∘+₁ ⟨
+ from ≅N ∘ (i₂ P₁ +₁ i₂ P₂) ∘ (f₂ B⇒C +₁ f₂ B⇒C′) ≈⟨ Equiv.refl ⟩
+ from ≅N ∘ i₂ P₃′ ∘ f₂ (together B⇒C B⇒C′) ≈⟨ pullˡ (universal∘i₂≈h₂ P₃′) ⟩
+ i₂ P₃ ∘ f₂ (together B⇒C B⇒C′) ∎
+
+⊗-resp-≈
+ : {A A′ B B′ : Obj}
+ {f f′ : Cospan A B}
+ {g g′ : Cospan A′ B′}
+ → Same f f′
+ → Same g g′
+ → Same (together f g) (together f′ g′)
+⊗-resp-≈ {_} {_} {_} {_} {f} {f′} {g} {g′} ≈f ≈g = record
+ { ≅N = ≈f.≅N ⊗ᵢ ≈g.≅N
+ ; from∘f₁≈f₁′ = from∘f₁≈f₁′
+ ; from∘f₂≈f₂′ = from∘f₂≈f₂′
+ }
+ where
+ open 𝒞 using (-+-)
+ module ≈f = Same ≈f
+ module ≈g = Same ≈g
+ open HomReasoning
+ open Cospan
+ open 𝒞 using (+₁-cong₂; +₁∘+₁)
+ from∘f₁≈f₁′ : (≈f.from +₁ ≈g.from) ∘ (f₁ f +₁ f₁ g) ≈ f₁ f′ +₁ f₁ g′
+ from∘f₁≈f₁′ = begin 
+ (≈f.from +₁ ≈g.from) ∘ (f₁ f +₁ f₁ g) ≈⟨ +₁∘+₁ ⟩
+ (≈f.from ∘ f₁ f) +₁ (≈g.from ∘ f₁ g) ≈⟨ +₁-cong₂ (≈f.from∘f₁≈f₁′) (≈g.from∘f₁≈f₁′) ⟩
+ f₁ f′ +₁ f₁ g′ ∎
+ from∘f₂≈f₂′ : (≈f.from +₁ ≈g.from) ∘ (f₂ f +₁ f₂ g) ≈ f₂ f′ +₁ f₂ g′
+ from∘f₂≈f₂′ = begin 
+ (≈f.from +₁ ≈g.from) ∘ (f₂ f +₁ f₂ g) ≈⟨ +₁∘+₁ ⟩
+ (≈f.from ∘ f₂ f) +₁ (≈g.from ∘ f₂ g) ≈⟨ +₁-cong₂ (≈f.from∘f₂≈f₂′) (≈g.from∘f₂≈f₂′) ⟩
+ f₂ f′ +₁ f₂ g′ ∎
+
+⊗ : Bifunctor Cospans Cospans Cospans
+⊗ = record
+ { F₀ = λ { (A , A′) → A + A′ }
+ ; F₁ = λ { (f , g) → together f g }
+ ; identity = λ { {x , y} → id⊗id≈id {x} {y} }
+ ; homomorphism = λ { {_} {_} {_} {A⇒B , A⇒B′} {B⇒C , B⇒C′} → homomorphism A⇒B B⇒C A⇒B′ B⇒C′ }
+ ; F-resp-≈ = λ { (≈f , ≈g) → ⊗-resp-≈ ≈f ≈g }
+ }