diff options
102 files changed, 11239 insertions, 1006 deletions
diff --git a/Adjoint/Instance/List.agda b/Adjoint/Instance/List.agda new file mode 100644 index 0000000..1b65985 --- /dev/null +++ b/Adjoint/Instance/List.agda @@ -0,0 +1,107 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_; suc) + +module Adjoint.Instance.List {ℓ : Level} where + +import Data.List as L +import Data.List.Relation.Binary.Pointwise as PW + +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory) +open import Category.Instance.Setoids.SymmetricMonoidal {ℓ} {ℓ} using (Setoids-×) + +module S = SymmetricMonoidalCategory Setoids-× + +open import Categories.Adjoint using (_⊣_) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Functor using (Functor; id; _∘F_) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.Object.Monoid S.monoidal using (Monoid; IsMonoid; Monoid⇒) +open import Data.Monoid using (toMonoid; toMonoid⇒) +open import Data.Opaque.List using ([-]ₛ; Listₛ; mapₛ; foldₛ; ++ₛ-homo; []ₛ-homo; fold-mapₛ; fold) +open import Data.Product using (_,_; uncurry) +open import Data.Setoid using (∣_∣) +open import Function using (_⟶ₛ_; _⟨$⟩_) +open import Functor.Forgetful.Instance.Monoid {suc ℓ} {ℓ} {ℓ} using () renaming (Forget to Forget′) +open import Functor.Free.Instance.Monoid {ℓ} {ℓ} using (Listₘ; mapₘ; ListMonoid) renaming (Free to Free′) +open import Relation.Binary using (Setoid) + +open Monoid +open Monoid⇒ + +open import Categories.Category using (Category) + +Mon[S] : Category (suc ℓ) ℓ ℓ +Mon[S] = Monoids S.monoidal + +Free : Functor S.U Mon[S] +Free = Free′ + +Forget : Functor Mon[S] S.U +Forget = Forget′ S.monoidal + +opaque + unfolding [-]ₛ mapₘ + map-[-]ₛ + : {X Y : Setoid ℓ ℓ} + (f : X ⟶ₛ Y) + {x : ∣ X ∣} + → (open Setoid (Listₛ Y)) + → [-]ₛ ⟨$⟩ (f ⟨$⟩ x) + ≈ arr (mapₘ f) ⟨$⟩ ([-]ₛ ⟨$⟩ x) + map-[-]ₛ {X} {Y} f {x} = Setoid.refl (Listₛ Y) + +unit : NaturalTransformation id (Forget ∘F Free) +unit = ntHelper record + { η = λ X → [-]ₛ {ℓ} {ℓ} {X} + ; commute = map-[-]ₛ + } + +opaque + unfolding toMonoid ListMonoid + foldₘ : (X : Monoid) → Monoid⇒ (Listₘ (Carrier X)) X + foldₘ X .arr = foldₛ (toMonoid X) + foldₘ X .preserves-μ {xs , ys} = ++ₛ-homo (toMonoid X) xs ys + foldₘ X .preserves-η {_} = []ₛ-homo (toMonoid X) + +opaque + unfolding foldₘ toMonoid⇒ mapₘ + fold-mapₘ + : {X Y : Monoid} + (f : Monoid⇒ X Y) + {x : ∣ Listₛ (Carrier X) ∣} + → (open Setoid (Carrier Y)) + → arr (foldₘ Y) ⟨$⟩ (arr (mapₘ (arr f)) ⟨$⟩ x) + ≈ arr f ⟨$⟩ (arr (foldₘ X) ⟨$⟩ x) + fold-mapₘ {X} {Y} f = uncurry (fold-mapₛ (toMonoid X) (toMonoid Y)) (toMonoid⇒ X Y f) + +counit : NaturalTransformation (Free ∘F Forget) id +counit = ntHelper record + { η = foldₘ + ; commute = fold-mapₘ + } + +opaque + unfolding mapₘ foldₘ fold + zig : (Aₛ : Setoid ℓ ℓ) + {xs : ∣ Listₛ Aₛ ∣} + → (open Setoid (Listₛ Aₛ)) + → arr (foldₘ (Listₘ Aₛ)) ⟨$⟩ (arr (mapₘ [-]ₛ) ⟨$⟩ xs) ≈ xs + zig Aₛ {xs = L.[]} = Setoid.refl (Listₛ Aₛ) + zig Aₛ {xs = x L.∷ xs} = Setoid.refl Aₛ PW.∷ zig Aₛ {xs = xs} + +opaque + unfolding foldₘ fold + zag : (M : Monoid) + {x : ∣ Carrier M ∣} + → (open Setoid (Carrier M)) + → arr (foldₘ M) ⟨$⟩ ([-]ₛ ⟨$⟩ x) ≈ x + zag M {x} = Setoid.sym (Carrier M) (identityʳ M {x , _}) + +List⊣ : Free ⊣ Forget +List⊣ = record + { unit = unit + ; counit = counit + ; zig = λ {X} → zig X + ; zag = λ {M} → zag M + } diff --git a/Adjoint/Instance/Multiset.agda b/Adjoint/Instance/Multiset.agda new file mode 100644 index 0000000..c51baa9 --- /dev/null +++ b/Adjoint/Instance/Multiset.agda @@ -0,0 +1,112 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_; suc; 0ℓ) + +module Adjoint.Instance.Multiset {ℓ : Level} where + +open import Category.Instance.Setoids.SymmetricMonoidal {ℓ} {ℓ} using (Setoids-×) + +private + module S = Setoids-× + +import Categories.Object.Monoid S.monoidal as MonoidObject +import Data.List as L +import Data.List.Relation.Binary.Permutation.Setoid as ↭ +import Functor.Forgetful.Instance.CMonoid S.symmetric as CMonoid +import Functor.Forgetful.Instance.Monoid S.monoidal as Monoid +import Object.Monoid.Commutative S.symmetric as CMonoidObject + +open import Categories.Adjoint using (_⊣_) +open import Categories.Category using (Category) +open import Categories.Functor using (Functor; id; _∘F_) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Category.Construction.CMonoids using (CMonoids) +open import Data.CMonoid using (toCMonoid; toCMonoid⇒) +open import Data.Monoid using (toMonoid; toMonoid⇒) +open import Data.Opaque.Multiset using ([-]ₛ; Multisetₛ; foldₛ; ++ₛ-homo; []ₛ-homo; fold-mapₛ; fold) +open import Data.Product using (_,_; uncurry) +open import Data.Setoid using (∣_∣) +open import Function using (_⟶ₛ_; _⟨$⟩_) +open import Functor.Free.Instance.CMonoid {ℓ} {ℓ} using (Multisetₘ; mapₘ; MultisetCMonoid) renaming (Free to Free′) +open import Relation.Binary using (Setoid) + +open CMonoidObject using (CommutativeMonoid; CommutativeMonoid⇒) +open CommutativeMonoid using (Carrier; monoid; identityʳ) +open CommutativeMonoid⇒ using (arr; monoid⇒) +open MonoidObject using (Monoid; Monoid⇒) +open Monoid⇒ using (preserves-μ; preserves-η) + +CMon[S] : Category (suc ℓ) ℓ ℓ +CMon[S] = CMonoids S.symmetric + +Free : Functor S.U CMon[S] +Free = Free′ + +Forget : Functor CMon[S] S.U +Forget = Monoid.Forget ∘F CMonoid.Forget + +opaque + unfolding [-]ₛ + map-[-]ₛ + : {X Y : Setoid ℓ ℓ} + (f : X ⟶ₛ Y) + {x : ∣ X ∣} + → (open Setoid (Multisetₛ Y)) + → [-]ₛ ⟨$⟩ (f ⟨$⟩ x) + ≈ arr (mapₘ f) ⟨$⟩ ([-]ₛ ⟨$⟩ x) + map-[-]ₛ {X} {Y} f {x} = Setoid.refl (Multisetₛ Y) + +unit : NaturalTransformation id (Forget ∘F Free) +unit = ntHelper record + { η = λ X → [-]ₛ {ℓ} {ℓ} {X} + ; commute = map-[-]ₛ + } + +opaque + unfolding toMonoid MultisetCMonoid + foldₘ : (X : CommutativeMonoid) → CommutativeMonoid⇒ (Multisetₘ (Carrier X)) X + foldₘ X .monoid⇒ .Monoid⇒.arr = foldₛ (toCMonoid X) + foldₘ X .monoid⇒ .preserves-μ {xs , ys} = ++ₛ-homo (toCMonoid X) xs ys + foldₘ X .monoid⇒ .preserves-η {_} = []ₛ-homo (toCMonoid X) + +opaque + unfolding foldₘ toMonoid⇒ + fold-mapₘ + : {X Y : CommutativeMonoid} + (f : CommutativeMonoid⇒ X Y) + {x : ∣ Multisetₛ (Carrier X) ∣} + → (open Setoid (Carrier Y)) + → arr (foldₘ Y) ⟨$⟩ (arr (mapₘ (arr f)) ⟨$⟩ x) + ≈ arr f ⟨$⟩ (arr (foldₘ X) ⟨$⟩ x) + fold-mapₘ {X} {Y} f = uncurry (fold-mapₛ (toCMonoid X) (toCMonoid Y)) (toCMonoid⇒ X Y f) + +counit : NaturalTransformation (Free ∘F Forget) id +counit = ntHelper record + { η = foldₘ + ; commute = fold-mapₘ + } + +opaque + unfolding foldₘ fold Multisetₛ + zig : (Aₛ : Setoid ℓ ℓ) + {xs : ∣ Multisetₛ Aₛ ∣} + → (open Setoid (Multisetₛ Aₛ)) + → arr (foldₘ (Multisetₘ Aₛ)) ⟨$⟩ (arr (mapₘ [-]ₛ) ⟨$⟩ xs) ≈ xs + zig Aₛ {L.[]} = ↭.↭-refl Aₛ + zig Aₛ {x L.∷ xs} = ↭.prep (Setoid.refl Aₛ) (zig Aₛ) + +opaque + unfolding foldₘ fold + zag : (M : CommutativeMonoid) + {x : ∣ Carrier M ∣} + → (open Setoid (Carrier M)) + → arr (foldₘ M) ⟨$⟩ ([-]ₛ ⟨$⟩ x) ≈ x + zag M {x} = Setoid.sym (Carrier M) (identityʳ M {x , _}) + +Multiset⊣ : Free ⊣ Forget +Multiset⊣ = record + { unit = unit + ; counit = counit + ; zig = λ {X} → zig X + ; zag = λ {M} → zag M + } diff --git a/Category/Instance/Properties/FinitelyCocompletes.agda b/Category/Cartesian/Instance/FinitelyCocompletes.agda index dedfa16..5425233 100644 --- a/Category/Instance/Properties/FinitelyCocompletes.agda +++ b/Category/Cartesian/Instance/FinitelyCocompletes.agda @@ -1,23 +1,27 @@ {-# OPTIONS --without-K --safe #-} -open import Level using (Level) -module Category.Instance.Properties.FinitelyCocompletes {o ℓ e : Level} where +open import Level using (Level; suc; _⊔_) +module Category.Cartesian.Instance.FinitelyCocompletes {o ℓ e : Level} where + +import Categories.Morphism as Morphism import Categories.Morphism.Reasoning as ⇒-Reasoning open import Categories.Category.BinaryProducts using (BinaryProducts) open import Categories.Category.Cartesian.Bundle using (CartesianCategory) -open import Categories.Category.Product using (Product) renaming (_⁂_ to _⁂′_) open import Categories.Diagram.Coequalizer using (IsCoequalizer) +open import Categories.Functor.Bifunctor using (flip-bifunctor) open import Categories.Functor.Core using (Functor) -open import Categories.Functor using (_∘F_) renaming (id to idF) +open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_; niHelper) +open import Categories.NaturalTransformation.NaturalIsomorphism.Properties using (pointwise-iso) open import Categories.Object.Coproduct using (IsCoproduct; IsCoproduct⇒Coproduct; Coproduct) open import Categories.Object.Initial using (IsInitial) +open import Data.Product.Base using (_,_) renaming (_×_ to _×′_) + open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) open import Category.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes; FinitelyCocompletes-Cartesian; _×₁_) -open import Data.Product.Base using (_,_) renaming (_×_ to _×′_) open import Functor.Exact using (IsRightExact; RightExactFunctor) -open import Level using (_⊔_; suc) +open import Functor.Exact.Instance.Swap using (Swap) FinitelyCocompletes-CC : CartesianCategory (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) FinitelyCocompletes-CC = record @@ -202,9 +206,37 @@ module _ {𝒞 : FinitelyCocompleteCategory o ℓ e} where ; F-resp-coeq = +-resp-coeq 𝒞 } } + module x+y = RightExactFunctor -+- + + ↔-+- : 𝒞 × 𝒞 ⇒ 𝒞 + ↔-+- = -+- ∘ Swap 𝒞 𝒞 + module y+x = RightExactFunctor ↔-+- [x+y]+z : (𝒞 × 𝒞) × 𝒞 ⇒ 𝒞 [x+y]+z = -+- ∘ (-+- ×₁ id) + module [x+y]+z = RightExactFunctor [x+y]+z x+[y+z] : (𝒞 × 𝒞) × 𝒞 ⇒ 𝒞 x+[y+z] = -+- ∘ (id ×₁ -+-) ∘ assocˡ + module x+[y+z] = RightExactFunctor x+[y+z] + + assoc-≃ : [x+y]+z.F ≃ x+[y+z].F + assoc-≃ = pointwise-iso (λ { ((X , Y) , Z) → ≅.sym (𝒞.+-assoc {X} {Y} {Z})}) commute + where + open 𝒞 + module 𝒞×𝒞×𝒞 = FinitelyCocompleteCategory ((𝒞 × 𝒞) × 𝒞) + open Morphism U using (_≅_; module ≅) + module +-assoc {X} {Y} {Z} = _≅_ (≅.sym (+-assoc {X} {Y} {Z})) + open import Categories.Category.BinaryProducts using (BinaryProducts) + open import Categories.Object.Duality 𝒞.U using (Coproduct⇒coProduct) + op-binaryProducts : BinaryProducts op + op-binaryProducts = record { product = Coproduct⇒coProduct coproduct } + open BinaryProducts op-binaryProducts using () renaming (assocʳ∘⁂ to +₁∘assocˡ) + open Equiv + commute + : {((X , Y) , Z) : 𝒞×𝒞×𝒞.Obj} + {((X′ , Y′) , Z′) : 𝒞×𝒞×𝒞.Obj} + → (F : ((X , Y) , Z) 𝒞×𝒞×𝒞.⇒ ((X′ , Y′) , Z′)) + → (+-assoc.from 𝒞.∘ [x+y]+z.₁ F) + ≈ (x+[y+z].₁ F 𝒞.∘ +-assoc.from) + commute {(X , Y) , Z} {(X′ , Y′) , Z′} ((F , G) , H) = sym +₁∘assocˡ diff --git a/Category/Cartesian/Instance/SymMonCat.agda b/Category/Cartesian/Instance/SymMonCat.agda new file mode 100644 index 0000000..7d91d52 --- /dev/null +++ b/Category/Cartesian/Instance/SymMonCat.agda @@ -0,0 +1,16 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; suc; _⊔_) + +module Category.Cartesian.Instance.SymMonCat {o ℓ e : Level} where + +open import Category.Instance.SymMonCat {o} {ℓ} {e} using (SymMonCat) +open import Category.Instance.Properties.SymMonCat {o} {ℓ} {e} using (SymMonCat-Cartesian) +open import Categories.Category.Cartesian.Bundle using (CartesianCategory) + +SymMonCat-CC : CartesianCategory (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) +SymMonCat-CC = record + { U = SymMonCat + ; cartesian = SymMonCat-Cartesian + } + diff --git a/Category/Cocomplete/Finitely/Bundle.agda b/Category/Cocomplete/Finitely/Bundle.agda index 74f434f..8af8633 100644 --- a/Category/Cocomplete/Finitely/Bundle.agda +++ b/Category/Cocomplete/Finitely/Bundle.agda @@ -28,6 +28,7 @@ record FinitelyCocompleteCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where ; monoidal = monoidal ; symmetric = symmetric } + {-# INJECTIVE_FOR_INFERENCE symmetricMonoidalCategory #-} cocartesianCategory : CocartesianCategory o ℓ e cocartesianCategory = record diff --git a/Category/Cocomplete/Finitely/SymmetricMonoidal.agda b/Category/Cocomplete/Finitely/SymmetricMonoidal.agda index 26813eb..2b66d19 100644 --- a/Category/Cocomplete/Finitely/SymmetricMonoidal.agda +++ b/Category/Cocomplete/Finitely/SymmetricMonoidal.agda @@ -4,11 +4,11 @@ open import Categories.Category.Core using (Category) module Category.Cocomplete.Finitely.SymmetricMonoidal {o ℓ e} {𝒞 : Category o ℓ e} where -open import Categories.Category.Cocomplete.Finitely using (FinitelyCocomplete) +open import Categories.Category.Cocomplete.Finitely 𝒞 using (FinitelyCocomplete) open import Categories.Category.Cocartesian 𝒞 using (module CocartesianMonoidal; module CocartesianSymmetricMonoidal) -module FinitelyCocompleteSymmetricMonoidal (finCo : FinitelyCocomplete 𝒞) where +module FinitelyCocompleteSymmetricMonoidal (finCo : FinitelyCocomplete) where open FinitelyCocomplete finCo using (cocartesian) open CocartesianMonoidal cocartesian using (+-monoidal) public diff --git a/Category/Construction/CMonoids.agda b/Category/Construction/CMonoids.agda new file mode 100644 index 0000000..c2793cf --- /dev/null +++ b/Category/Construction/CMonoids.agda @@ -0,0 +1,57 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Core using (Category) +open import Categories.Category.Monoidal.Core using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) +open import Level using (Level; _⊔_) + +-- The category of commutative monoids internal to a symmetric monoidal category + +module Category.Construction.CMonoids + {o ℓ e : Level} + {𝒞 : Category o ℓ e} + {M : Monoidal 𝒞} + (symmetric : Symmetric M) + where + +open import Categories.Functor using (Functor) +open import Categories.Morphism.Reasoning 𝒞 +open import Object.Monoid.Commutative symmetric + +open Category 𝒞 +open Monoidal M +open HomReasoning +open CommutativeMonoid⇒ + +CMonoids : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e +CMonoids = record + { Obj = CommutativeMonoid + ; _⇒_ = CommutativeMonoid⇒ + ; _≈_ = λ f g → arr f ≈ arr g + ; id = record + { monoid⇒ = record + { arr = id + ; preserves-μ = identityˡ ○ introʳ ⊗.identity + ; preserves-η = identityˡ + } + } + ; _∘_ = λ f g → record + { monoid⇒ = record + { arr = arr f ∘ arr g + ; preserves-μ = glue (preserves-μ f) (preserves-μ g) ○ ∘-resp-≈ʳ (⟺ ⊗.homomorphism) + ; preserves-η = glueTrianglesˡ (preserves-η f) (preserves-η g) + } + } + ; assoc = assoc + ; sym-assoc = sym-assoc + ; identityˡ = identityˡ + ; identityʳ = identityʳ + ; identity² = identity² + ; equiv = record + { refl = refl + ; sym = sym + ; trans = trans + } + ; ∘-resp-≈ = ∘-resp-≈ + } + where open Equiv diff --git a/Category/Construction/CMonoids/Properties.agda b/Category/Construction/CMonoids/Properties.agda new file mode 100644 index 0000000..fab8b0b --- /dev/null +++ b/Category/Construction/CMonoids/Properties.agda @@ -0,0 +1,74 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Core using (Category) +open import Categories.Category.Monoidal.Core using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) +open import Level using (Level; _⊔_) + +module Category.Construction.CMonoids.Properties + {o ℓ e : Level} + {𝒞 : Category o ℓ e} + {monoidal : Monoidal 𝒞} + (symmetric : Symmetric monoidal) + where + +import Categories.Category.Monoidal.Reasoning monoidal as ⊗-Reasoning +import Categories.Morphism.Reasoning 𝒞 as ⇒-Reasoning +import Category.Construction.Monoids.Properties {o} {ℓ} {e} {𝒞} monoidal as Monoids + +open import Categories.Category using (_[_,_]) +open import Categories.Category.Monoidal.Symmetric.Properties symmetric using (module Shorthands) +open import Categories.Morphism using (_≅_) +open import Categories.Morphism.Notation using (_[_≅_]) +open import Categories.Object.Monoid monoidal using (Monoid) +open import Category.Construction.CMonoids symmetric using (CMonoids) +open import Data.Product using (Σ-syntax; _,_) +open import Object.Monoid.Commutative symmetric using (CommutativeMonoid) + +module 𝒞 = Category 𝒞 + +open CommutativeMonoid using (monoid) renaming (Carrier to ∣_∣) + +transport-by-iso + : {X : 𝒞.Obj} + → (M : CommutativeMonoid) + → 𝒞 [ ∣ M ∣ ≅ X ] + → Σ[ Xₘ ∈ CommutativeMonoid ] CMonoids [ M ≅ Xₘ ] +transport-by-iso {X} M ∣M∣≅X + using (Xₘ′ , M≅Xₘ′) ← Monoids.transport-by-iso {X} (monoid M) ∣M∣≅X = Xₘ , M≅Xₘ + where + module M≅Xₘ′ = _≅_ M≅Xₘ′ + open _≅_ ∣M∣≅X + module Xₘ′ = Monoid Xₘ′ + open CommutativeMonoid M + open 𝒞 using (_≈_; _∘_) + open Shorthands + open ⊗-Reasoning + open ⇒-Reasoning + open Symmetric symmetric using (_⊗₁_; module braiding) + comm : from ∘ μ ∘ to ⊗₁ to ≈ (from ∘ μ ∘ to ⊗₁ to) ∘ σ⇒ + comm = begin + from ∘ μ ∘ to ⊗₁ to ≈⟨ push-center (commutative) ⟩ + from ∘ μ ∘ σ⇒ ∘ to ⊗₁ to ≈⟨ pushʳ (pushʳ (braiding.⇒.commute (to , to))) ⟩ + (from ∘ μ ∘ to ⊗₁ to) ∘ σ⇒ ∎ + Xₘ : CommutativeMonoid + Xₘ = record + { Carrier = X + ; isCommutativeMonoid = record + { isMonoid = Xₘ′.isMonoid + ; commutative = comm + } + } + M⇒Xₘ : CMonoids [ M , Xₘ ] + M⇒Xₘ = record { monoid⇒ = M≅Xₘ′.from } + Xₘ⇒M : CMonoids [ Xₘ , M ] + Xₘ⇒M = record { monoid⇒ = M≅Xₘ′.to } + M≅Xₘ : CMonoids [ M ≅ Xₘ ] + M≅Xₘ = record + { from = M⇒Xₘ + ; to = Xₘ⇒M + ; iso = record + { isoˡ = isoˡ + ; isoʳ = isoʳ + } + } diff --git a/Category/Construction/Monoids/Properties.agda b/Category/Construction/Monoids/Properties.agda new file mode 100644 index 0000000..df48063 --- /dev/null +++ b/Category/Construction/Monoids/Properties.agda @@ -0,0 +1,108 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Core using (Category) +open import Categories.Category.Monoidal.Core using (Monoidal) +open import Level using (Level; _⊔_) + +module Category.Construction.Monoids.Properties + {o ℓ e : Level} + {𝒞 : Category o ℓ e} + (monoidal : Monoidal 𝒞) + where + +import Categories.Category.Monoidal.Reasoning monoidal as ⊗-Reasoning +import Categories.Morphism.Reasoning 𝒞 as ⇒-Reasoning + +open import Categories.Category using (_[_,_]) +open import Categories.Category.Construction.Monoids monoidal using (Monoids) +open import Categories.Category.Monoidal.Utilities monoidal using (module Shorthands; _⊗ᵢ_) +open import Categories.Morphism using (_≅_) +open import Categories.Morphism.Notation using (_[_≅_]) +open import Categories.Object.Monoid monoidal using (Monoid) +open import Data.Product using (Σ-syntax; _,_) + +module 𝒞 = Category 𝒞 + +open Monoid using () renaming (Carrier to ∣_∣) + +transport-by-iso + : {X : 𝒞.Obj} + → (M : Monoid) + → 𝒞 [ ∣ M ∣ ≅ X ] + → Σ[ Xₘ ∈ Monoid ] Monoids [ M ≅ Xₘ ] +transport-by-iso {X} M ∣M∣≅X = Xₘ , M≅Xₘ + where + open _≅_ ∣M∣≅X + open Monoid M + open 𝒞 using (_∘_; id; _≈_; module Equiv) + open Monoidal monoidal + using (_⊗₀_; _⊗₁_; assoc-commute-from; unitorˡ-commute-from; unitorʳ-commute-from) + open Shorthands + open ⊗-Reasoning + open ⇒-Reasoning + as + : (from ∘ μ ∘ to ⊗₁ to) + ∘ (from ∘ μ ∘ to ⊗₁ to) ⊗₁ id + ≈ (from ∘ μ ∘ to ⊗₁ to) + ∘ id ⊗₁ (from ∘ μ ∘ to ⊗₁ to) + ∘ α⇒ + as = extendˡ (begin + (μ ∘ to ⊗₁ to) ∘ (from ∘ μ ∘ to ⊗₁ to) ⊗₁ id ≈⟨ pullʳ merge₁ʳ ⟩ + μ ∘ (to ∘ from ∘ μ ∘ to ⊗₁ to) ⊗₁ to ≈⟨ refl⟩∘⟨ cancelˡ isoˡ ⟩⊗⟨refl ⟩ + μ ∘ (μ ∘ to ⊗₁ to) ⊗₁ to ≈⟨ refl⟩∘⟨ split₁ˡ ⟩ + μ ∘ μ ⊗₁ id ∘ (to ⊗₁ to) ⊗₁ to ≈⟨ extendʳ assoc ⟩ + μ ∘ (id ⊗₁ μ ∘ α⇒) ∘ (to ⊗₁ to) ⊗₁ to ≈⟨ pushʳ (pullʳ assoc-commute-from) ⟩ + (μ ∘ id ⊗₁ μ) ∘ to ⊗₁ (to ⊗₁ to) ∘ α⇒ ≈⟨ extendˡ (pullˡ merge₂ˡ) ⟩ + (μ ∘ to ⊗₁ (μ ∘ to ⊗₁ to)) ∘ α⇒ ≈⟨ (refl⟩∘⟨ refl⟩⊗⟨ insertˡ isoˡ) ⟩∘⟨refl ⟩ + (μ ∘ to ⊗₁ (to ∘ from ∘ μ ∘ to ⊗₁ to)) ∘ α⇒ ≈⟨ pushˡ (pushʳ split₂ʳ) ⟩ + (μ ∘ to ⊗₁ to) ∘ id ⊗₁ (from ∘ μ ∘ to ⊗₁ to) ∘ α⇒ ∎) + idˡ : λ⇒ ≈ (from ∘ μ ∘ to ⊗₁ to) ∘ (from ∘ η) ⊗₁ id + idˡ = begin + λ⇒ ≈⟨ insertˡ isoʳ ⟩ + from ∘ to ∘ λ⇒ ≈⟨ refl⟩∘⟨ unitorˡ-commute-from ⟨ + from ∘ λ⇒ ∘ id ⊗₁ to ≈⟨ refl⟩∘⟨ pushˡ identityˡ ⟩ + from ∘ μ ∘ η ⊗₁ id ∘ id ⊗₁ to ≈⟨ refl⟩∘⟨ refl⟩∘⟨ serialize₁₂ ⟨ + from ∘ μ ∘ η ⊗₁ to ≈⟨ refl⟩∘⟨ refl⟩∘⟨ insertˡ isoˡ ⟩⊗⟨refl ⟩ + from ∘ μ ∘ (to ∘ from ∘ η) ⊗₁ to ≈⟨ pushʳ (pushʳ split₁ʳ) ⟩ + (from ∘ μ ∘ to ⊗₁ to) ∘ (from ∘ η) ⊗₁ id ∎ + idʳ : ρ⇒ ≈ (from ∘ μ ∘ to ⊗₁ to) ∘ id ⊗₁ (from ∘ η) + idʳ = begin + ρ⇒ ≈⟨ insertˡ isoʳ ⟩ + from ∘ to ∘ ρ⇒ ≈⟨ refl⟩∘⟨ unitorʳ-commute-from ⟨ + from ∘ ρ⇒ ∘ to ⊗₁ id ≈⟨ refl⟩∘⟨ pushˡ identityʳ ⟩ + from ∘ μ ∘ id ⊗₁ η ∘ to ⊗₁ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ serialize₂₁ ⟨ + from ∘ μ ∘ to ⊗₁ η ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ insertˡ isoˡ ⟩ + from ∘ μ ∘ to ⊗₁ (to ∘ from ∘ η) ≈⟨ pushʳ (pushʳ split₂ʳ) ⟩ + (from ∘ μ ∘ to ⊗₁ to) ∘ id ⊗₁ (from ∘ η) ∎ + Xₘ : Monoid + Xₘ = record + { Carrier = X + ; isMonoid = record + { μ = from ∘ μ ∘ to ⊗₁ to + ; η = from ∘ η + ; assoc = as + ; identityˡ = idˡ + ; identityʳ = idʳ + } + } + M⇒Xₘ : Monoids [ M , Xₘ ] + M⇒Xₘ = record + { arr = from + ; preserves-μ = pushʳ (insertʳ (_≅_.isoˡ (∣M∣≅X ⊗ᵢ ∣M∣≅X))) + ; preserves-η = Equiv.refl + } + Xₘ⇒M : Monoids [ Xₘ , M ] + Xₘ⇒M = record + { arr = to + ; preserves-μ = cancelˡ isoˡ + ; preserves-η = cancelˡ isoˡ + } + M≅Xₘ : Monoids [ M ≅ Xₘ ] + M≅Xₘ = record + { from = M⇒Xₘ + ; to = Xₘ⇒M + ; iso = record + { isoˡ = isoˡ + ; isoʳ = isoʳ + } + } diff --git a/Category/Diagram/Cospan.agda b/Category/Diagram/Cospan.agda new file mode 100644 index 0000000..b988651 --- /dev/null +++ b/Category/Diagram/Cospan.agda @@ -0,0 +1,117 @@ +{-# OPTIONS --without-K --safe #-} + +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Level using (Level; _⊔_) + +module Category.Diagram.Cospan {o ℓ e : Level} (𝒞 : FinitelyCocompleteCategory o ℓ e) where + +import Categories.Morphism as Morphism +import Categories.Morphism.Reasoning as ⇒-Reasoning + +open import Relation.Binary using (IsEquivalence) + +module 𝒞 = FinitelyCocompleteCategory 𝒞 + +open 𝒞 hiding (i₁; i₂; _≈_) +open ⇒-Reasoning U using (switch-fromtoˡ; glueTrianglesˡ) +open Morphism U using (_≅_; module ≅) + +record Cospan (A B : Obj) : Set (o ⊔ ℓ) where + + constructor cospan + + field + {N} : Obj + f₁ : A ⇒ N + f₂ : B ⇒ N + +private + variable + A B C D : Obj + +record _≈_ (C D : Cospan A B) : Set (ℓ ⊔ e) where + + module C = Cospan C + module D = Cospan D + + field + ≅N : C.N ≅ D.N + + open _≅_ ≅N public + module ≅N = _≅_ ≅N + + field + from∘f₁≈f₁ : from ∘ C.f₁ 𝒞.≈ D.f₁ + from∘f₂≈f₂ : from ∘ C.f₂ 𝒞.≈ D.f₂ + +infix 4 _≈_ + +private + variable + f g h : Cospan A B + +≈-refl : f ≈ f +≈-refl {f = cospan f₁ f₂} = record + { ≅N = ≅.refl + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ + } + where abstract + from∘f₁≈f₁ : id ∘ f₁ 𝒞.≈ f₁ + from∘f₁≈f₁ = identityˡ + from∘f₂≈f₂ : id ∘ f₂ 𝒞.≈ f₂ + from∘f₂≈f₂ = identityˡ + +≈-sym : f ≈ g → g ≈ f +≈-sym f≈g = record + { ≅N = ≅.sym ≅N + ; from∘f₁≈f₁ = a + ; from∘f₂≈f₂ = b + } + where abstract + open _≈_ f≈g + a : ≅N.to ∘ D.f₁ 𝒞.≈ C.f₁ + a = Equiv.sym (switch-fromtoˡ ≅N from∘f₁≈f₁) + b : ≅N.to ∘ D.f₂ 𝒞.≈ C.f₂ + b = Equiv.sym (switch-fromtoˡ ≅N from∘f₂≈f₂) + +≈-trans : f ≈ g → g ≈ h → f ≈ h +≈-trans f≈g g≈h = record + { ≅N = ≅.trans f≈g.≅N g≈h.≅N + ; from∘f₁≈f₁ = a + ; from∘f₂≈f₂ = b + } + where abstract + module f≈g = _≈_ f≈g + module g≈h = _≈_ g≈h + a : (g≈h.≅N.from ∘ f≈g.≅N.from) ∘ f≈g.C.f₁ 𝒞.≈ g≈h.D.f₁ + a = glueTrianglesˡ g≈h.from∘f₁≈f₁ f≈g.from∘f₁≈f₁ + b : (g≈h.≅N.from ∘ f≈g.≅N.from) ∘ f≈g.C.f₂ 𝒞.≈ g≈h.D.f₂ + b = glueTrianglesˡ g≈h.from∘f₂≈f₂ f≈g.from∘f₂≈f₂ + +≈-equiv : {A B : 𝒞.Obj} → IsEquivalence (_≈_ {A} {B}) +≈-equiv = record + { refl = ≈-refl + ; sym = ≈-sym + ; trans = ≈-trans + } + +compose : Cospan A B → Cospan B C → Cospan A C +compose (cospan f g) (cospan h i) = cospan (i₁ ∘ f) (i₂ ∘ i) + where + open pushout g h + +identity : Cospan A A +identity = cospan id id + +compose-3 : Cospan A B → Cospan B C → Cospan C D → Cospan A D +compose-3 (cospan f₁ f₂) (cospan g₁ g₂) (cospan h₁ h₂) = cospan (P₃.i₁ ∘ P₁.i₁ ∘ f₁) (P₃.i₂ ∘ P₂.i₂ ∘ h₂) + where + module P₁ = pushout f₂ g₁ + module P₂ = pushout g₂ h₁ + module P₃ = pushout P₁.i₂ P₂.i₁ + +_⊗_ : Cospan A B → Cospan C D → Cospan (A + C) (B + D) +_⊗_ (cospan f₁ f₂) (cospan g₁ g₂) = cospan (f₁ +₁ g₁) (f₂ +₁ g₂) + +infixr 10 _⊗_ diff --git a/Category/Instance/Cospans.agda b/Category/Instance/Cospans.agda index d54499d..d17a58d 100644 --- a/Category/Instance/Cospans.agda +++ b/Category/Instance/Cospans.agda @@ -7,7 +7,9 @@ open import Level using (_⊔_) module Category.Instance.Cospans {o ℓ e} (𝒞 : FinitelyCocompleteCategory o ℓ e) where -open FinitelyCocompleteCategory 𝒞 +module 𝒞 = FinitelyCocompleteCategory 𝒞 +open 𝒞 using (U; pushout) +open Category U hiding (_≈_) open import Categories.Diagram.Pushout U using (Pushout) open import Categories.Diagram.Pushout.Properties U using (pushout-resp-≈; up-to-iso) @@ -15,8 +17,8 @@ open import Categories.Morphism U using (_≅_; module ≅) open import Categories.Morphism.Reasoning U using ( switch-fromtoˡ - ; glueTrianglesˡ ; pullˡ ; pullʳ + ; cancelˡ ) open import Category.Diagram.Pushout U @@ -26,254 +28,213 @@ open import Category.Diagram.Pushout U ; extend-i₁-iso ; extend-i₂-iso ) -private +open import Category.Diagram.Cospan 𝒞 using (Cospan; compose; compose-3; identity; _≈_; cospan; ≈-trans; ≈-sym; ≈-equiv) +private variable - A B C D X Y Z : Obj - f g h : A ⇒ B - -record Cospan (A B : Obj) : Set (o ⊔ ℓ) where - - field - {N} : Obj - f₁ : A ⇒ N - f₂ : B ⇒ N - -compose : Cospan A B → Cospan B C → Cospan A C -compose c₁ c₂ = record { f₁ = p.i₁ ∘ C₁.f₁ ; f₂ = p.i₂ ∘ C₂.f₂ } - where - module C₁ = Cospan c₁ - module C₂ = Cospan c₂ - module p = pushout C₁.f₂ C₂.f₁ - -id-Cospan : Cospan A A -id-Cospan = record { f₁ = id ; f₂ = id } - -compose-3 : Cospan A B → Cospan B C → Cospan C D → Cospan A D -compose-3 c₁ c₂ c₃ = record { f₁ = P₃.i₁ ∘ P₁.i₁ ∘ C₁.f₁ ; f₂ = P₃.i₂ ∘ P₂.i₂ ∘ C₃.f₂ } - where - module C₁ = Cospan c₁ - module C₂ = Cospan c₂ - module C₃ = Cospan c₃ - module P₁ = pushout C₁.f₂ C₂.f₁ - module P₂ = pushout C₂.f₂ C₃.f₁ - module P₃ = pushout P₁.i₂ P₂.i₁ + A B C D : Obj -record Same (C C′ : Cospan A B) : Set (ℓ ⊔ e) where - - module C = Cospan C - module C′ = Cospan C′ - - field - ≅N : C.N ≅ C′.N - - open _≅_ ≅N public - module ≅N = _≅_ ≅N - - field - from∘f₁≈f₁′ : from ∘ C.f₁ ≈ C′.f₁ - from∘f₂≈f₂′ : from ∘ C.f₂ ≈ C′.f₂ - -same-refl : {C : Cospan A B} → Same C C -same-refl = record - { ≅N = ≅.refl - ; from∘f₁≈f₁′ = identityˡ - ; from∘f₂≈f₂′ = identityˡ - } - -same-sym : {C C′ : Cospan A B} → Same C C′ → Same C′ C -same-sym C≅C′ = record - { ≅N = ≅.sym ≅N - ; from∘f₁≈f₁′ = Equiv.sym (switch-fromtoˡ ≅N from∘f₁≈f₁′) - ; from∘f₂≈f₂′ = Equiv.sym (switch-fromtoˡ ≅N from∘f₂≈f₂′) - } - where - open Same C≅C′ - -same-trans : {C C′ C″ : Cospan A B} → Same C C′ → Same C′ C″ → Same C C″ -same-trans C≈C′ C′≈C″ = record - { ≅N = ≅.trans C≈C′.≅N C′≈C″.≅N - ; from∘f₁≈f₁′ = glueTrianglesˡ C′≈C″.from∘f₁≈f₁′ C≈C′.from∘f₁≈f₁′ - ; from∘f₂≈f₂′ = glueTrianglesˡ C′≈C″.from∘f₂≈f₂′ C≈C′.from∘f₂≈f₂′ - } - where - module C≈C′ = Same C≈C′ - module C′≈C″ = Same C′≈C″ +private + variable + f g h : Cospan A B -compose-idˡ : {C : Cospan A B} → Same (compose C id-Cospan) C -compose-idˡ {_} {_} {C} = record +compose-idˡ : compose f identity ≈ f +compose-idˡ {f = cospan {N} f₁ f₂} = record { ≅N = ≅P - ; from∘f₁≈f₁′ = begin - ≅P.from ∘ P.i₁ ∘ C.f₁ ≈⟨ assoc ⟨ - (≅P.from ∘ P.i₁) ∘ C.f₁ ≈⟨ P.universal∘i₁≈h₁ ⟩∘⟨refl ⟩ - id ∘ C.f₁ ≈⟨ identityˡ ⟩ - C.f₁ ∎ - ; from∘f₂≈f₂′ = begin - ≅P.from ∘ P.i₂ ∘ id ≈⟨ refl⟩∘⟨ identityʳ ⟩ - ≅P.from ∘ P.i₂ ≈⟨ P.universal∘i₂≈h₂ ⟩ - C.f₂ ∎ + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where open HomReasoning - module C = Cospan C - P = pushout C.f₂ id + P P′ : Pushout f₂ id + P = pushout f₂ id + P′ = pushout-f-id {f = f₂} module P = Pushout P - P′ = pushout-f-id {f = C.f₂} + ≅P : P.Q ≅ N ≅P = up-to-iso P P′ module ≅P = _≅_ ≅P - -compose-idʳ : {C : Cospan A B} → Same (compose id-Cospan C) C -compose-idʳ {_} {_} {C} = record + abstract + from∘f₁≈f₁ : ≅P.from ∘ P.i₁ ∘ f₁ 𝒞.≈ f₁ + from∘f₁≈f₁ = begin + ≅P.from ∘ P.i₁ ∘ f₁ ≈⟨ cancelˡ P.universal∘i₁≈h₁ ⟩ + f₁ ∎ + from∘f₂≈f₂ : ≅P.from ∘ P.i₂ ∘ id 𝒞.≈ f₂ + from∘f₂≈f₂ = begin + ≅P.from ∘ P.i₂ ∘ id ≈⟨ refl⟩∘⟨ identityʳ ⟩ + ≅P.from ∘ P.i₂ ≈⟨ P.universal∘i₂≈h₂ ⟩ + f₂ ∎ + +compose-idʳ : compose identity f ≈ f +compose-idʳ {f = cospan {N} f₁ f₂} = record { ≅N = ≅P - ; from∘f₁≈f₁′ = begin - ≅P.from ∘ P.i₁ ∘ id ≈⟨ refl⟩∘⟨ identityʳ ⟩ - ≅P.from ∘ P.i₁ ≈⟨ P.universal∘i₁≈h₁ ⟩ - C.f₁ ∎ - ; from∘f₂≈f₂′ = begin - ≅P.from ∘ P.i₂ ∘ C.f₂ ≈⟨ assoc ⟨ - (≅P.from ∘ P.i₂) ∘ C.f₂ ≈⟨ P.universal∘i₂≈h₂ ⟩∘⟨refl ⟩ - id ∘ C.f₂ ≈⟨ identityˡ ⟩ - C.f₂ ∎ + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where open HomReasoning - module C = Cospan C - P = pushout id C.f₁ + P P′ : Pushout id f₁ + P = pushout id f₁ module P = Pushout P - P′ = pushout-id-g {g = C.f₁} + P′ = pushout-id-g {g = f₁} + ≅P : P.Q ≅ N ≅P = up-to-iso P P′ module ≅P = _≅_ ≅P - -compose-id² : Same {A} (compose id-Cospan id-Cospan) id-Cospan + abstract + from∘f₁≈f₁ : ≅P.from ∘ P.i₁ ∘ id 𝒞.≈ f₁ + from∘f₁≈f₁ = begin + ≅P.from ∘ P.i₁ ∘ id ≈⟨ refl⟩∘⟨ identityʳ ⟩ + ≅P.from ∘ P.i₁ ≈⟨ P.universal∘i₁≈h₁ ⟩ + f₁ ∎ + from∘f₂≈f₂ : ≅P.from ∘ P.i₂ ∘ f₂ 𝒞.≈ f₂ + from∘f₂≈f₂ = begin + ≅P.from ∘ P.i₂ ∘ f₂ ≈⟨ cancelˡ P.universal∘i₂≈h₂ ⟩ + f₂ ∎ + +compose-id² : compose identity identity ≈ identity {A} compose-id² = compose-idˡ -compose-3-right - : {c₁ : Cospan A B} - {c₂ : Cospan B C} - {c₃ : Cospan C D} - → Same (compose c₁ (compose c₂ c₃)) (compose-3 c₁ c₂ c₃) -compose-3-right {_} {_} {_} {_} {c₁} {c₂} {c₃} = record - { ≅N = up-to-iso P₄′ P₄ - ; from∘f₁≈f₁′ = sym-assoc ○ P₄′.universal∘i₁≈h₁ ⟩∘⟨refl ○ assoc - ; from∘f₂≈f₂′ = sym-assoc ○ P₄′.universal∘i₂≈h₂ ⟩∘⟨refl +compose-3-right : compose f (compose g h) ≈ compose-3 f g h +compose-3-right {f = f} {g = g} {h = h} = record + { ≅N = ≅N + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where open HomReasoning - module C₁ = Cospan c₁ - module C₂ = Cospan c₂ - module C₃ = Cospan c₃ + module C₁ = Cospan f + module C₂ = Cospan g + module C₃ = Cospan h + P₁ : Pushout C₁.f₂ C₂.f₁ P₁ = pushout C₁.f₂ C₂.f₁ + P₂ : Pushout C₂.f₂ C₃.f₁ P₂ = pushout C₂.f₂ C₃.f₁ module P₁ = Pushout P₁ module P₂ = Pushout P₂ + P₃ : Pushout P₁.i₂ P₂.i₁ P₃ = pushout P₁.i₂ P₂.i₁ module P₃ = Pushout P₃ + P₄ P₄′ : Pushout C₁.f₂ (P₂.i₁ ∘ C₂.f₁) P₄ = glue-i₂ P₁ P₃ - module P₄ = Pushout P₄ P₄′ = pushout C₁.f₂ (P₂.i₁ ∘ C₂.f₁) + module P₄ = Pushout P₄ module P₄′ = Pushout P₄′ - -compose-3-left - : {c₁ : Cospan A B} - {c₂ : Cospan B C} - {c₃ : Cospan C D} - → Same (compose (compose c₁ c₂) c₃) (compose-3 c₁ c₂ c₃) -compose-3-left {_} {_} {_} {_} {c₁} {c₂} {c₃} = record + ≅N : P₄′.Q ≅ P₄.Q + ≅N = up-to-iso P₄′ P₄ + module ≅N = _≅_ ≅N + abstract + from∘f₁≈f₁ : ≅N.from ∘ P₄′.i₁ ∘ C₁.f₁ 𝒞.≈ P₃.i₁ ∘ P₁.i₁ ∘ C₁.f₁ + from∘f₁≈f₁ = sym-assoc ○ P₄′.universal∘i₁≈h₁ ⟩∘⟨refl ○ assoc + from∘f₂≈f₂ : ≅N.from ∘ P₄′.i₂ ∘ P₂.i₂ ∘ C₃.f₂ 𝒞.≈ P₄.i₂ ∘ P₂.i₂ ∘ C₃.f₂ + from∘f₂≈f₂ = sym-assoc ○ P₄′.universal∘i₂≈h₂ ⟩∘⟨refl + +compose-3-left : compose (compose f g) h ≈ compose-3 f g h +compose-3-left {f = f} {g = g} {h = h} = record { ≅N = up-to-iso P₄′ P₄ - ; from∘f₁≈f₁′ = sym-assoc ○ P₄′.universal∘i₁≈h₁ ⟩∘⟨refl - ; from∘f₂≈f₂′ = sym-assoc ○ P₄′.universal∘i₂≈h₂ ⟩∘⟨refl ○ assoc + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where open HomReasoning - module C₁ = Cospan c₁ - module C₂ = Cospan c₂ - module C₃ = Cospan c₃ + module C₁ = Cospan f + module C₂ = Cospan g + module C₃ = Cospan h + P₁ : Pushout C₁.f₂ C₂.f₁ P₁ = pushout C₁.f₂ C₂.f₁ + P₂ : Pushout C₂.f₂ C₃.f₁ P₂ = pushout C₂.f₂ C₃.f₁ module P₁ = Pushout P₁ module P₂ = Pushout P₂ + P₃ : Pushout P₁.i₂ P₂.i₁ P₃ = pushout P₁.i₂ P₂.i₁ module P₃ = Pushout P₃ + P₄ P₄′ : Pushout (P₁.i₂ ∘ C₂.f₂) C₃.f₁ P₄ = glue-i₁ P₂ P₃ - module P₄ = Pushout P₄ P₄′ = pushout (P₁.i₂ ∘ C₂.f₂) C₃.f₁ + module P₄ = Pushout P₄ module P₄′ = Pushout P₄′ + ≅N : P₄′.Q ≅ P₄.Q + ≅N = up-to-iso P₄′ P₄ + module ≅N = _≅_ ≅N + abstract + from∘f₁≈f₁ : ≅N.from ∘ P₄′.i₁ ∘ P₁.i₁ ∘ C₁.f₁ 𝒞.≈ P₄.i₁ ∘ P₁.i₁ ∘ C₁.f₁ + from∘f₁≈f₁ = sym-assoc ○ P₄′.universal∘i₁≈h₁ ⟩∘⟨refl + from∘f₂≈f₂ : ≅N.from ∘ P₄′.i₂ ∘ C₃.f₂ 𝒞.≈ P₃.i₂ ∘ P₂.i₂ ∘ C₃.f₂ + from∘f₂≈f₂ = sym-assoc ○ P₄′.universal∘i₂≈h₂ ⟩∘⟨refl ○ assoc compose-assoc : {c₁ : Cospan A B} {c₂ : Cospan B C} {c₃ : Cospan C D} - → Same (compose c₁ (compose c₂ c₃)) (compose (compose c₁ c₂) c₃) -compose-assoc = same-trans compose-3-right (same-sym compose-3-left) + → compose c₁ (compose c₂ c₃) ≈ (compose (compose c₁ c₂) c₃) +compose-assoc = ≈-trans compose-3-right (≈-sym compose-3-left) compose-sym-assoc : {c₁ : Cospan A B} {c₂ : Cospan B C} {c₃ : Cospan C D} - → Same (compose (compose c₁ c₂) c₃) (compose c₁ (compose c₂ c₃)) -compose-sym-assoc = same-trans compose-3-left (same-sym compose-3-right) + → compose (compose c₁ c₂) c₃ ≈ compose c₁ (compose c₂ c₃) +compose-sym-assoc = ≈-trans compose-3-left (≈-sym compose-3-right) compose-equiv : {c₂ c₂′ : Cospan B C} {c₁ c₁′ : Cospan A B} - → Same c₂ c₂′ - → Same c₁ c₁′ - → Same (compose c₁ c₂) (compose c₁′ c₂′) + → c₂ ≈ c₂′ + → c₁ ≈ c₁′ + → compose c₁ c₂ ≈ compose c₁′ c₂′ compose-equiv {_} {_} {_} {c₂} {c₂′} {c₁} {c₁′} ≈C₂ ≈C₁ = record - { ≅N = up-to-iso P P″ - ; from∘f₁≈f₁′ = begin - ≅P.from ∘ P.i₁ ∘ C₁.f₁ ≈⟨ assoc ⟨ - (≅P.from ∘ P.i₁) ∘ C₁.f₁ ≈⟨ P.universal∘i₁≈h₁ ⟩∘⟨refl ⟩ - (P′.i₁ ∘ ≈C₁.from) ∘ C₁.f₁ ≈⟨ assoc ⟩ - P′.i₁ ∘ ≈C₁.from ∘ C₁.f₁ ≈⟨ refl⟩∘⟨ ≈C₁.from∘f₁≈f₁′ ⟩ - P′.i₁ ∘ C₁′.f₁ ∎ - ; from∘f₂≈f₂′ = begin - ≅P.from ∘ P.i₂ ∘ C₂.f₂ ≈⟨ assoc ⟨ - (≅P.from ∘ P.i₂) ∘ C₂.f₂ ≈⟨ P.universal∘i₂≈h₂ ⟩∘⟨refl ⟩ - (P′.i₂ ∘ ≈C₂.from) ∘ C₂.f₂ ≈⟨ assoc ⟩ - P′.i₂ ∘ ≈C₂.from ∘ C₂.f₂ ≈⟨ refl⟩∘⟨ ≈C₂.from∘f₂≈f₂′ ⟩ - P′.i₂ ∘ C₂′.f₂ ∎ + { ≅N = ≅P + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where module C₁ = Cospan c₁ module C₁′ = Cospan c₁′ module C₂ = Cospan c₂ module C₂′ = Cospan c₂′ + P : Pushout C₁.f₂ C₂.f₁ P = pushout C₁.f₂ C₂.f₁ + P′ : Pushout C₁′.f₂ C₂′.f₁ P′ = pushout C₁′.f₂ C₂′.f₁ - module ≈C₁ = Same ≈C₁ - module ≈C₂ = Same ≈C₂ + module ≈C₁ = _≈_ ≈C₁ + module ≈C₂ = _≈_ ≈C₂ P′′ : Pushout (≈C₁.to ∘ C₁′.f₂) (≈C₂.to ∘ C₂′.f₁) P′′ = extend-i₂-iso (extend-i₁-iso P′ (≅.sym ≈C₁.≅N)) (≅.sym ≈C₂.≅N) P″ : Pushout C₁.f₂ C₂.f₁ P″ = pushout-resp-≈ P′′ - (Equiv.sym (switch-fromtoˡ ≈C₁.≅N ≈C₁.from∘f₂≈f₂′)) - (Equiv.sym (switch-fromtoˡ ≈C₂.≅N ≈C₂.from∘f₁≈f₁′)) + (Equiv.sym (switch-fromtoˡ ≈C₁.≅N ≈C₁.from∘f₂≈f₂)) + (Equiv.sym (switch-fromtoˡ ≈C₂.≅N ≈C₂.from∘f₁≈f₁)) module P = Pushout P module P′ = Pushout P′ ≅P : P.Q ≅ P′.Q ≅P = up-to-iso P P″ module ≅P = _≅_ ≅P open HomReasoning + abstract + from∘f₁≈f₁ : ≅P.from ∘ P.i₁ ∘ C₁.f₁ 𝒞.≈ P′.i₁ ∘ C₁′.f₁ + from∘f₁≈f₁ = begin + ≅P.from ∘ P.i₁ ∘ C₁.f₁ ≈⟨ pullˡ P.universal∘i₁≈h₁ ⟩ + (P′.i₁ ∘ ≈C₁.from) ∘ C₁.f₁ ≈⟨ pullʳ ≈C₁.from∘f₁≈f₁ ⟩ + P′.i₁ ∘ C₁′.f₁ ∎ + from∘f₂≈f₂ : ≅P.from ∘ P.i₂ ∘ C₂.f₂ 𝒞.≈ P′.i₂ ∘ C₂′.f₂ + from∘f₂≈f₂ = begin + ≅P.from ∘ P.i₂ ∘ C₂.f₂ ≈⟨ pullˡ P.universal∘i₂≈h₂ ⟩ + (P′.i₂ ∘ ≈C₂.from) ∘ C₂.f₂ ≈⟨ pullʳ ≈C₂.from∘f₂≈f₂ ⟩ + P′.i₂ ∘ C₂′.f₂ ∎ Cospans : Category o (o ⊔ ℓ) (ℓ ⊔ e) Cospans = record { Obj = Obj ; _⇒_ = Cospan - ; _≈_ = Same - ; id = id-Cospan + ; _≈_ = _≈_ + ; id = identity ; _∘_ = flip compose ; assoc = compose-assoc ; sym-assoc = compose-sym-assoc ; identityˡ = compose-idˡ ; identityʳ = compose-idʳ ; identity² = compose-id² - ; equiv = record - { refl = same-refl - ; sym = same-sym - ; trans = same-trans - } + ; equiv = ≈-equiv ; ∘-resp-≈ = compose-equiv } diff --git a/Category/Instance/DecoratedCospans.agda b/Category/Instance/DecoratedCospans.agda index c0c62c7..b527265 100644 --- a/Category/Instance/DecoratedCospans.agda +++ b/Category/Instance/DecoratedCospans.agda @@ -18,15 +18,18 @@ module Category.Instance.DecoratedCospans module 𝒞 = FinitelyCocompleteCategory 𝒞 module 𝒟 = SymmetricMonoidalCategory 𝒟 +import Categories.Category.Monoidal.Utilities as ⊗-Util import Category.Instance.Cospans 𝒞 as Cospans +import Category.Diagram.Cospan 𝒞 as Cospan open import Categories.Category using (Category; _[_∘_]) open import Categories.Category.Cocartesian using (module CocartesianMonoidal) open import Categories.Diagram.Pushout using (Pushout) open import Categories.Diagram.Pushout.Properties 𝒞.U using (up-to-iso) -open import Categories.Functor.Properties using ([_]-resp-≅) +open import Categories.Functor.Properties using ([_]-resp-≅; [_]-resp-square) open import Categories.Morphism.Reasoning using (switch-fromtoˡ; glueTrianglesˡ) open import Cospan.Decorated 𝒞 F using (DecoratedCospan) +open import Relation.Binary using (IsEquivalence) open import Data.Product using (_,_) open import Function using (flip) open import Level using (_⊔_) @@ -50,7 +53,7 @@ private compose : DecoratedCospan A B → DecoratedCospan B C → DecoratedCospan A C compose c₁ c₂ = record - { cospan = Cospans.compose C₁.cospan C₂.cospan + { cospan = Cospan.compose C₁.cospan C₂.cospan ; decoration = F₁ [ i₁ , i₂ ] ∘ φ ∘ s⊗t } where @@ -67,64 +70,80 @@ compose c₁ c₂ = record identity : DecoratedCospan A A identity = record - { cospan = Cospans.id-Cospan + { cospan = Cospan.identity ; decoration = 𝒟.U [ F₁ 𝒞.initial.! ∘ ε ] } -record Same (C₁ C₂ : DecoratedCospan A B) : Set (ℓ ⊔ e ⊔ e′) where +record _≈_ (C₁ C₂ : DecoratedCospan A B) : Set (ℓ ⊔ e ⊔ e′) where - module C₁ = DecoratedCospan C₁ - module C₂ = DecoratedCospan C₂ + private + module C₁ = DecoratedCospan C₁ + module C₂ = DecoratedCospan C₂ field - cospans-≈ : Cospans.Same C₁.cospan C₂.cospan + cospans-≈ : C₁.cospan Cospan.≈ C₂.cospan - open Cospans.Same cospans-≈ public - open 𝒟 - open Morphism U using (_≅_) + open Cospan._≈_ cospans-≈ public + open Morphism 𝒟.U using (_≅_) field - same-deco : F₁ ≅N.from ∘ C₁.decoration ≈ C₂.decoration + same-deco : F₁ ≅N.from 𝒟.∘ C₁.decoration 𝒟.≈ C₂.decoration ≅F[N] : F₀ C₁.N ≅ F₀ C₂.N ≅F[N] = [ F′ ]-resp-≅ ≅N -same-refl : {C : DecoratedCospan A B} → Same C C -same-refl = record - { cospans-≈ = Cospans.same-refl - ; same-deco = F-identity ⟩∘⟨refl ○ identityˡ - } - where - open 𝒟 - open HomReasoning +infix 4 _≈_ -same-sym : {C C′ : DecoratedCospan A B} → Same C C′ → Same C′ C -same-sym C≅C′ = record - { cospans-≈ = Cospans.same-sym cospans-≈ - ; same-deco = sym (switch-fromtoˡ 𝒟.U ≅F[N] same-deco) - } - where - open Same C≅C′ - open 𝒟.Equiv - -same-trans : {C C′ C″ : DecoratedCospan A B} → Same C C′ → Same C′ C″ → Same C C″ -same-trans C≈C′ C′≈C″ = record - { cospans-≈ = Cospans.same-trans C≈C′.cospans-≈ C′≈C″.cospans-≈ - ; same-deco = - homomorphism ⟩∘⟨refl ○ - glueTrianglesˡ 𝒟.U C′≈C″.same-deco C≈C′.same-deco - } - where - module C≈C′ = Same C≈C′ - module C′≈C″ = Same C′≈C″ - open 𝒟.HomReasoning +module _ where + + open 𝒟.HomReasoning + open 𝒟.Equiv + open 𝒟 using (identityˡ) + + private + variable + f g h : DecoratedCospan A B + + abstract + + ≈-refl : f ≈ f + ≈-refl = record + { cospans-≈ = Cospan.≈-refl + ; same-deco = F-identity ⟩∘⟨refl ○ identityˡ + } + + ≈-sym : f ≈ g → g ≈ f + ≈-sym f≈g = record + { cospans-≈ = Cospan.≈-sym cospans-≈ + ; same-deco = sym (switch-fromtoˡ 𝒟.U ≅F[N] same-deco) + } + where + open _≈_ f≈g + + ≈-trans : f ≈ g → g ≈ h → f ≈ h + ≈-trans f≈g g≈h = record + { cospans-≈ = Cospan.≈-trans f≈g.cospans-≈ g≈h.cospans-≈ + ; same-deco = + homomorphism ⟩∘⟨refl ○ + glueTrianglesˡ 𝒟.U g≈h.same-deco f≈g.same-deco + } + where + module f≈g = _≈_ f≈g + module g≈h = _≈_ g≈h + + ≈-equiv : {A B : 𝒞.Obj} → IsEquivalence (_≈_ {A} {B}) + ≈-equiv = record + { refl = ≈-refl + ; sym = ≈-sym + ; trans = ≈-trans + } compose-assoc : {c₁ : DecoratedCospan A B} {c₂ : DecoratedCospan B C} {c₃ : DecoratedCospan C D} - → Same (compose c₁ (compose c₂ c₃)) (compose (compose c₁ c₂) c₃) -compose-assoc {_} {_} {_} {_} {c₁} {c₂} {c₃} = record + → compose c₁ (compose c₂ c₃) ≈ compose (compose c₁ c₂) c₃ +compose-assoc {A} {B} {C} {D} {c₁} {c₂} {c₃} = record { cospans-≈ = Cospans.compose-assoc ; same-deco = deco-assoc } @@ -133,14 +152,19 @@ compose-assoc {_} {_} {_} {_} {c₁} {c₂} {c₃} = record module C₂ = DecoratedCospan c₂ module C₃ = DecoratedCospan c₃ open 𝒞 using (+-assoc; pushout; [_,_]; _+₁_; _+_) renaming (_∘_ to _∘′_; id to id′) + p₁ : Pushout 𝒞.U C₁.f₂ C₂.f₁ p₁ = pushout C₁.f₂ C₂.f₁ + p₂ : Pushout 𝒞.U C₂.f₂ C₃.f₁ p₂ = pushout C₂.f₂ C₃.f₁ module P₁ = Pushout p₁ module P₂ = Pushout p₂ + p₃ : Pushout 𝒞.U P₁.i₂ P₂.i₁ p₃ = pushout P₁.i₂ P₂.i₁ + p₁₃ p₄ : Pushout 𝒞.U C₁.f₂ (P₂.i₁ ∘′ C₂.f₁) p₁₃ = glue-i₂ p₁ p₃ - p₂₃ = glue-i₁ p₂ p₃ p₄ = pushout C₁.f₂ (P₂.i₁ ∘′ C₂.f₁) + p₂₃ p₅ : Pushout 𝒞.U (P₁.i₂ ∘′ C₂.f₂) C₃.f₁ + p₂₃ = glue-i₁ p₂ p₃ p₅ = pushout (P₁.i₂ ∘′ C₂.f₂) C₃.f₁ module P₃ = Pushout p₃ module P₄ = Pushout p₄ @@ -151,36 +175,50 @@ compose-assoc {_} {_} {_} {_} {c₁} {c₂} {c₃} = record module P₄≅P₁₃ = _≅_ (up-to-iso p₄ p₁₃) module P₅≅P₂₃ = _≅_ (up-to-iso p₅ p₂₃) + N M P Q R : 𝒞.Obj N = C₁.N M = C₂.N P = C₃.N Q = P₁.Q R = P₂.Q - φ = ⊗-homo.η - φ-commute = ⊗-homo.commute - - a = C₁.f₂ - b = C₂.f₁ - c = C₂.f₂ - d = C₂.f₁ + f : N 𝒞.⇒ Q f = P₁.i₁ + + g : M 𝒞.⇒ Q g = P₁.i₂ + + h : M 𝒞.⇒ R h = P₂.i₁ + + i : P 𝒞.⇒ R i = P₂.i₂ + j : Q 𝒞.⇒ P₃.Q j = P₃.i₁ + + k : R 𝒞.⇒ P₃.Q k = P₃.i₂ + w : N 𝒞.⇒ P₄.Q w = P₄.i₁ + + x : R 𝒞.⇒ P₄.Q x = P₄.i₂ + + y : Q 𝒞.⇒ P₅.Q y = P₅.i₁ + + z : P 𝒞.⇒ P₅.Q z = P₅.i₂ + l : P₂₃.Q 𝒞.⇒ P₅.Q l = P₅≅P₂₃.to + + m : P₄.Q 𝒞.⇒ P₁₃.Q m = P₄≅P₁₃.from - module +-assoc = _≅_ +-assoc + module +-assoc {m} {n} {o} = _≅_ (+-assoc {m} {n} {o}) module _ where @@ -194,23 +232,19 @@ compose-assoc {_} {_} {_} {_} {c₁} {c₂} {c₃} = record open 𝒞.HomReasoning open 𝒞.Equiv - copairings : ((l ∘ m) ∘ [ w , x ]) ∘ (id +₁ [ h , i ]) ≈ [ y , z ] ∘ ([ f , g ] +₁ id) ∘ +-assoc.from + copairings : ((l ∘ m) ∘ [ w , x ]) ∘ (id +₁ [ h , i ]) 𝒞.≈ [ y , z ] ∘ ([ f , g ] +₁ id) ∘ +-assoc.from copairings = begin - ((l ∘ m) ∘ [ w , x ]) ∘ (id +₁ [ h , i ]) ≈⟨ pushˡ assoc ⟩ - l ∘ (m ∘ [ w , x ]) ∘ (id +₁ [ h , i ]) ≈⟨ refl⟩∘⟨ ∘[] ⟩∘⟨refl ⟩ - l ∘ [ m ∘ w , m ∘ x ] ∘ (id +₁ [ h , i ]) ≈⟨ refl⟩∘⟨ []-cong₂ (P₄.universal∘i₁≈h₁) (P₄.universal∘i₂≈h₂) ⟩∘⟨refl ⟩ - l ∘ [ j ∘ f , k ] ∘ (id +₁ [ h , i ]) ≈⟨ pullˡ ∘[] ⟩ - [ l ∘ j ∘ f , l ∘ k ] ∘ (id +₁ [ h , i ]) ≈⟨ []-congʳ (pullˡ P₂₃.universal∘i₁≈h₁) ⟩∘⟨refl ⟩ - [ y ∘ f , l ∘ k ] ∘ (id +₁ [ h , i ]) ≈⟨ []∘+₁ ⟩ - [ (y ∘ f) ∘ id , (l ∘ k) ∘ [ h , i ] ] ≈⟨ []-cong₂ identityʳ (pullʳ ∘[]) ⟩ - [ y ∘ f , l ∘ [ k ∘ h , k ∘ i ] ] ≈⟨ []-congˡ (refl⟩∘⟨ []-congʳ P₃.commute) ⟨ - [ y ∘ f , l ∘ [ j ∘ g , k ∘ i ] ] ≈⟨ []-congˡ ∘[] ⟩ - [ y ∘ f , [ l ∘ j ∘ g , l ∘ k ∘ i ] ] ≈⟨ []-congˡ ([]-congˡ P₂₃.universal∘i₂≈h₂) ⟩ - [ y ∘ f , [ l ∘ j ∘ g , z ] ] ≈⟨ []-congˡ ([]-congʳ (pullˡ P₂₃.universal∘i₁≈h₁)) ⟩ - [ y ∘ f , [ y ∘ g , z ] ] ≈⟨ []∘assocˡ ⟨ - [ [ y ∘ f , y ∘ g ] , z ] ∘ +-assoc.from ≈⟨ []-cong₂ ∘[] identityʳ ⟩∘⟨refl ⟨ - [ y ∘ [ f , g ] , z ∘ id ] ∘ +-assoc.from ≈⟨ pullˡ []∘+₁ ⟨ - [ y , z ] ∘ ([ f , g ] +₁ id) ∘ +-assoc.from ∎ + ((l ∘ m) ∘ [ w , x ]) ∘ (id +₁ [ h , i ]) ≈⟨ ∘[] ⟩∘⟨refl ⟩ + [(l ∘ m) ∘ w , (l ∘ m) ∘ x ] ∘ (id +₁ [ h , i ]) ≈⟨ []-cong₂ (pullʳ P₄.universal∘i₁≈h₁) (pullʳ P₄.universal∘i₂≈h₂) ⟩∘⟨refl ⟩ + [ l ∘ j ∘ f , l ∘ k ] ∘ (id +₁ [ h , i ]) ≈⟨ []-congʳ (pullˡ P₂₃.universal∘i₁≈h₁) ⟩∘⟨refl ⟩ + [ y ∘ f , l ∘ k ] ∘ (id +₁ [ h , i ]) ≈⟨ []∘+₁ ⟩ + [ (y ∘ f) ∘ id , (l ∘ k) ∘ [ h , i ] ] ≈⟨ []-cong₂ identityʳ ∘[] ⟩ + [ y ∘ f , [ (l ∘ k) ∘ h , (l ∘ k) ∘ i ] ] ≈⟨ []-congˡ ([]-cong₂ (pullʳ (sym P₃.commute)) (assoc ○ P₂₃.universal∘i₂≈h₂)) ⟩ + [ y ∘ f , [ l ∘ j ∘ g , z ] ] ≈⟨ []-congˡ ([]-congʳ (pullˡ P₂₃.universal∘i₁≈h₁)) ⟩ + [ y ∘ f , [ y ∘ g , z ] ] ≈⟨ []∘assocˡ ⟨ + [ [ y ∘ f , y ∘ g ] , z ] ∘ +-assoc.from ≈⟨ []-cong₂ ∘[] identityʳ ⟩∘⟨refl ⟨ + [ y ∘ [ f , g ] , z ∘ id ] ∘ +-assoc.from ≈⟨ pullˡ []∘+₁ ⟨ + [ y , z ] ∘ ([ f , g ] +₁ id) ∘ +-assoc.from ∎ module _ where @@ -219,153 +253,160 @@ compose-assoc {_} {_} {_} {_} {c₁} {c₂} {c₃} = record open 𝒟 using (_⊗₀_; _⊗₁_; id; _∘_; _≈_; assoc; sym-assoc; identityʳ; ⊗; identityˡ; triangle; assoc-commute-to; assoc-commute-from) open 𝒟 using (_⇒_; unit) - α⇒ = 𝒟.associator.from - α⇐ = 𝒟.associator.to - - λ⇒ = 𝒟.unitorˡ.from - λ⇐ = 𝒟.unitorˡ.to - - ρ⇒ = 𝒟.unitorʳ.from - ρ⇐ = 𝒟.unitorʳ.to - - module α≅ = 𝒟.associator - module λ≅ = 𝒟.unitorˡ - module ρ≅ = 𝒟.unitorʳ + open ⊗-Util 𝒟.monoidal using (module Shorthands) + open Shorthands using (α⇒; α⇐; λ⇒; λ⇐; ρ⇒; ρ⇐) open Coherence 𝒟.monoidal using (λ₁≅ρ₁⇐) open 𝒟.Equiv + +-α⇒ : {m n o : 𝒞.Obj} → m + (n + o) 𝒞.⇒ (m + n) + o +-α⇒ = +-assoc.from + +-α⇐ : {m n o : 𝒞.Obj} → (m + n) + o 𝒞.⇒ m + (n + o) +-α⇐ = +-assoc.to - s : unit ⇒ F₀ C₁.N + s : unit ⇒ F₀ N s = C₁.decoration - t : unit ⇒ F₀ C₂.N + t : unit ⇒ F₀ M t = C₂.decoration - u : unit ⇒ F₀ C₃.N + u : unit ⇒ F₀ P u = C₃.decoration - F-copairings : F₁ (l ∘′ m) ∘ F₁ [ w , x ] ∘ F₁ (id′ +₁ [ h , i ]) ≈ F₁ [ y , z ] ∘ F₁ ([ f , g ] +₁ id′) ∘ F₁ (+-assoc.from) - F-copairings = begin - F₁ (l ∘′ m) ∘ F₁ [ w , x ] ∘ F₁ (id′ +₁ [ h , i ]) ≈⟨ pushˡ homomorphism ⟨ - F₁ ((l ∘′ m) ∘′ [ w , x ]) ∘ F₁ (id′ +₁ [ h , i ]) ≈⟨ homomorphism ⟨ - F₁ (((l ∘′ m) ∘′ [ w , x ]) ∘′ (id′ +₁ [ h , i ])) ≈⟨ F-resp-≈ copairings ⟩ - F₁ ([ y , z ] ∘′ ([ f , g ] +₁ id′) ∘′ +-assoc.from) ≈⟨ homomorphism ⟩ - F₁ [ y , z ] ∘ F₁ (([ f , g ] +₁ id′) ∘′ +-assoc.from) ≈⟨ refl⟩∘⟨ homomorphism ⟩ - F₁ [ y , z ] ∘ F₁ ([ f , g ] +₁ id′) ∘ F₁ +-assoc.from ∎ - - coherences : φ (N , M + P) ∘ id ⊗₁ φ (M , P) ≈ F₁ +-assoc.to ∘ φ (N + M , P) ∘ φ (N , M) ⊗₁ id ∘ α⇐ - coherences = begin - φ (N , M + P) ∘ id ⊗₁ φ (M , P) ≈⟨ insertʳ α≅.isoʳ ⟩ - ((φ (N , M + P) ∘ id ⊗₁ φ (M , P)) ∘ α⇒) ∘ α⇐ ≈⟨ assoc ⟩∘⟨refl ⟩ - (φ (N , M + P) ∘ id ⊗₁ φ (M , P) ∘ α⇒) ∘ α⇐ ≈⟨ assoc ⟩ - φ (N , M + P) ∘ (id ⊗₁ φ (M , P) ∘ α⇒) ∘ α⇐ ≈⟨ extendʳ associativity ⟨ - F₁ +-assoc.to ∘ (φ (N + M , P) ∘ φ (N , M) ⊗₁ id) ∘ α⇐ ≈⟨ refl⟩∘⟨ assoc ⟩ - F₁ +-assoc.to ∘ φ (N + M , P) ∘ φ (N , M) ⊗₁ id ∘ α⇐ ∎ - - triangle-to : α⇒ ∘ ρ⇐ ⊗₁ id ≈ id ⊗₁ λ⇐ - triangle-to = begin - α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ pullˡ identityˡ ⟨ - id ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ ⊗.identity ⟩∘⟨refl ⟨ - id ⊗₁ id ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩⊗⟨ λ≅.isoˡ ⟩∘⟨refl ⟨ - id ⊗₁ (λ⇐ ∘ λ⇒) ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ identityʳ ⟩⊗⟨refl ⟩∘⟨refl ⟨ - (id ∘ id) ⊗₁ (λ⇐ ∘ λ⇒) ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ pushˡ ⊗-distrib-over-∘ ⟩ - id ⊗₁ λ⇐ ∘ id ⊗₁ λ⇒ ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩∘⟨ pullˡ triangle ⟩ - id ⊗₁ λ⇐ ∘ ρ⇒ ⊗₁ id ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟨ - id ⊗₁ λ⇐ ∘ (ρ⇒ ∘ ρ⇐) ⊗₁ (id ∘ id) ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ identityˡ ⟩ - id ⊗₁ λ⇐ ∘ (ρ⇒ ∘ ρ⇐) ⊗₁ id ≈⟨ refl⟩∘⟨ ρ≅.isoʳ ⟩⊗⟨refl ⟩ - id ⊗₁ λ⇐ ∘ id ⊗₁ id ≈⟨ refl⟩∘⟨ ⊗.identity ⟩ - id ⊗₁ λ⇐ ∘ id ≈⟨ identityʳ ⟩ - id ⊗₁ λ⇐ ∎ - - unitors : s ⊗₁ (t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈ α⇒ ∘ (s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ - unitors = begin - s ⊗₁ (t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ pushˡ split₂ʳ ⟩ - s ⊗₁ t ⊗₁ u ∘ id ⊗₁ ρ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ λ₁≅ρ₁⇐ ⟩∘⟨refl ⟨ - s ⊗₁ t ⊗₁ u ∘ id ⊗₁ λ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ triangle-to ⟨ - s ⊗₁ t ⊗₁ u ∘ α⇒ ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ extendʳ assoc-commute-from ⟨ - α⇒ ∘ (s ⊗₁ t) ⊗₁ u ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pushˡ split₁ʳ ⟨ - α⇒ ∘ (s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ∎ - - F-l∘m = F₁ (l ∘′ m) + F[l∘m] : F₀ P₄.Q ⇒ F₀ P₅.Q + F[l∘m] = F₁ (l ∘′ m) + + F[w,x] : F₀ (N + R) ⇒ F₀ P₄.Q F[w,x] = F₁ [ w , x ] + + F[h,i] : F₀ (M + P) ⇒ F₀ R F[h,i] = F₁ [ h , i ] + + F[y,z] : F₀ (Q + P) ⇒ F₀ P₅.Q F[y,z] = F₁ [ y , z ] + + F[f,g] : F₀ (N + M) ⇒ F₀ Q F[f,g] = F₁ [ f , g ] - F-[f,g]+id = F₁ ([ f , g ] +₁ id′) - F-id+[h,i] = F₁ (id′ +₁ [ h , i ]) - φ-N,R = φ (N , R) - φ-M,P = φ (M , P) - φ-N+M,P = φ (N + M , P) - φ-N+M = φ (N , M) - φ-N,M+P = φ (N , M + P) - φ-N,M = φ (N , M) - φ-Q,P = φ (Q , P) + + F[[f,g]+id] : F₀ ((N + M) + P) ⇒ F₀ (Q + P) + F[[f,g]+id] = F₁ ([ f , g ] +₁ id′) + + F[id+[h,i]] : F₀ (N + (M + P)) ⇒ F₀ (N + R) + F[id+[h,i]] = F₁ (id′ +₁ [ h , i ]) + + φ[N,R] : F₀ N ⊗₀ F₀ R 𝒟.⇒ F₀ (N + R) + φ[N,R] = ⊗-homo.η (N , R) + + φ[M,P] : F₀ M ⊗₀ F₀ P 𝒟.⇒ F₀ (M + P) + φ[M,P] = ⊗-homo.η (M , P) + + φ[N+M,P] : F₀ (N + M) ⊗₀ F₀ P 𝒟.⇒ F₀ ((N + M) + P) + φ[N+M,P] = ⊗-homo.η (N + M , P) + + φ[N,M] : F₀ N ⊗₀ F₀ M 𝒟.⇒ F₀ (N + M) + φ[N,M] = ⊗-homo.η (N , M) + + φ[N,M+P] : F₀ N ⊗₀ F₀ (M + P) 𝒟.⇒ F₀ (N + (M + P)) + φ[N,M+P] = ⊗-homo.η (N , M + P) + + φ[Q,P] : F₀ Q ⊗₀ F₀ P 𝒟.⇒ F₀ (Q + P) + φ[Q,P] = ⊗-homo.η (Q , P) + + s⊗[t⊗u] : unit 𝒟.⇒ F₀ N ⊗₀ (F₀ M ⊗₀ F₀ P) s⊗[t⊗u] = s ⊗₁ (t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ + + [s⊗t]⊗u : unit 𝒟.⇒ (F₀ N ⊗₀ F₀ M) ⊗₀ F₀ P [s⊗t]⊗u = (s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ - deco-assoc - : F-l∘m ∘ F[w,x] ∘ φ-N,R ∘ s ⊗₁ (F[h,i] ∘ φ-M,P ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ - ≈ F[y,z] ∘ φ-Q,P ∘ (F[f,g] ∘ φ-N,M ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ - deco-assoc = begin - F-l∘m ∘ F[w,x] ∘ φ-N,R ∘ s ⊗₁ (F[h,i] ∘ φ-M,P ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ pullˡ refl ⟩ - (F-l∘m ∘ F[w,x]) ∘ φ-N,R ∘ s ⊗₁ (F[h,i] ∘ φ-M,P ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ split₂ˡ ⟩∘⟨refl ⟩ - (F-l∘m ∘ F[w,x]) ∘ φ-N,R ∘ (id ⊗₁ F[h,i] ∘ s ⊗₁ (φ-M,P ∘ t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ split₂ˡ) ⟩∘⟨refl ⟩ - (F-l∘m ∘ F[w,x]) ∘ φ-N,R ∘ (id ⊗₁ F[h,i] ∘ id ⊗₁ φ-M,P ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc ⟩ - (F-l∘m ∘ F[w,x]) ∘ φ-N,R ∘ id ⊗₁ F[h,i] ∘ (id ⊗₁ φ-M,P ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F-identity ⟩⊗⟨refl ⟩∘⟨ refl ⟨ - (F-l∘m ∘ F[w,x]) ∘ φ-N,R ∘ F₁ id′ ⊗₁ F[h,i] ∘ (id ⊗₁ φ-M,P ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (id′ , [ h , i ])) ⟩ - (F-l∘m ∘ F[w,x]) ∘ F-id+[h,i] ∘ φ-N,M+P ∘ (id ⊗₁ φ-M,P ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ pullˡ assoc ⟩ - (F-l∘m ∘ F[w,x] ∘ F-id+[h,i]) ∘ φ-N,M+P ∘ (id ⊗₁ φ-M,P ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc ⟩ - (F-l∘m ∘ F[w,x] ∘ F-id+[h,i]) ∘ φ-N,M+P ∘ id ⊗₁ φ-M,P ∘ s⊗[t⊗u] ≈⟨ refl⟩∘⟨ sym-assoc ⟩ - (F-l∘m ∘ F[w,x] ∘ F-id+[h,i]) ∘ (φ-N,M+P ∘ id ⊗₁ φ-M,P) ∘ s⊗[t⊗u] ≈⟨ F-copairings ⟩∘⟨ coherences ⟩∘⟨ unitors ⟩ - (F[y,z] ∘ F-[f,g]+id ∘ F₁ +-α⇒) ∘ (F₁ +-α⇐ ∘ φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ sym-assoc ⟩∘⟨ assoc ⟩ - ((F[y,z] ∘ F-[f,g]+id) ∘ F₁ +-α⇒) ∘ F₁ +-α⇐ ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ assoc ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ F₁ +-α⇒ ∘ F₁ +-α⇐ ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ pushˡ homomorphism ⟨ - (F[y,z] ∘ F-[f,g]+id) ∘ F₁ (+-α⇒ ∘′ +-α⇐) ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ F-resp-≈ +-assoc.isoʳ ⟩∘⟨refl ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ F₁ id′ ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ F-identity ⟩∘⟨refl ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ id ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ identityˡ ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ sym-assoc ⟩∘⟨refl ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ ((φ-N+M,P ∘ φ-N,M ⊗₁ id) ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ cancelInner α≅.isoˡ ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ (φ-N+M,P ∘ φ-N,M ⊗₁ id) ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ assoc ⟩ - (F[y,z] ∘ F-[f,g]+id) ∘ φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ [s⊗t]⊗u ≈⟨ assoc ⟩ - F[y,z] ∘ F-[f,g]+id ∘ φ-N+M,P ∘ φ-N,M ⊗₁ id ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟨ - F[y,z] ∘ F-[f,g]+id ∘ φ-N+M,P ∘ (φ-N,M ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute ([ f , g ] , id′)) ⟨ - F[y,z] ∘ φ-Q,P ∘ F[f,g] ⊗₁ F₁ id′ ∘ (φ-N,M ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ F-identity ⟩∘⟨ refl ⟩ - F[y,z] ∘ φ-Q,P ∘ F[f,g] ⊗₁ id ∘ (φ-N,M ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟨ - F[y,z] ∘ φ-Q,P ∘ (F[f,g] ∘ φ-N,M ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ∎ - -compose-idʳ : {C : DecoratedCospan A B} → Same (compose identity C) C + abstract + F-copairings : F[l∘m] ∘ F[w,x] ∘ F[id+[h,i]] 𝒟.≈ F[y,z] ∘ F[[f,g]+id] ∘ F₁ +-assoc.from + F-copairings = begin + F[l∘m] ∘ F[w,x] ∘ F[id+[h,i]] ≈⟨ pushˡ homomorphism ⟨ + F₁ ((l ∘′ m) ∘′ [ w , x ]) ∘ F[id+[h,i]] ≈⟨ [ F′ ]-resp-square copairings ⟩ + F[y,z] ∘ F₁ (([ f , g ] +₁ id′) ∘′ +-assoc.from) ≈⟨ refl⟩∘⟨ homomorphism ⟩ + F[y,z] ∘ F[[f,g]+id] ∘ F₁ +-assoc.from ∎ + + coherences : φ[N,M+P] ∘ id ⊗₁ φ[M,P] 𝒟.≈ F₁ +-assoc.to ∘ φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐ + coherences = begin + φ[N,M+P] ∘ id ⊗₁ φ[M,P] ≈⟨ refl⟩∘⟨ insertʳ 𝒟.associator.isoʳ ⟩ + φ[N,M+P] ∘ (id ⊗₁ φ[M,P] ∘ α⇒) ∘ α⇐ ≈⟨ extendʳ associativity ⟨ + F₁ +-assoc.to ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id) ∘ α⇐ ≈⟨ refl⟩∘⟨ assoc ⟩ + F₁ +-assoc.to ∘ φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐ ∎ + + triangle-to : α⇒ {𝒟.unit} {𝒟.unit} {𝒟.unit} ∘ ρ⇐ ⊗₁ id 𝒟.≈ id ⊗₁ λ⇐ + triangle-to = begin + α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ pullˡ identityˡ ⟨ + id ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ ⊗.identity ⟩∘⟨refl ⟨ + id ⊗₁ id ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩⊗⟨ 𝒟.unitorˡ.isoˡ ⟩∘⟨refl ⟨ + id ⊗₁ (λ⇐ ∘ λ⇒) ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ identityʳ ⟩⊗⟨refl ⟩∘⟨refl ⟨ + (id ∘ id) ⊗₁ (λ⇐ ∘ λ⇒) ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ pushˡ ⊗-distrib-over-∘ ⟩ + id ⊗₁ λ⇐ ∘ id ⊗₁ λ⇒ ∘ α⇒ ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩∘⟨ pullˡ triangle ⟩ + id ⊗₁ λ⇐ ∘ ρ⇒ ⊗₁ id ∘ ρ⇐ ⊗₁ id ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟨ + id ⊗₁ λ⇐ ∘ (ρ⇒ ∘ ρ⇐) ⊗₁ (id ∘ id) ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ identityˡ ⟩ + id ⊗₁ λ⇐ ∘ (ρ⇒ ∘ ρ⇐) ⊗₁ id ≈⟨ refl⟩∘⟨ 𝒟.unitorʳ.isoʳ ⟩⊗⟨refl ⟩ + id ⊗₁ λ⇐ ∘ id ⊗₁ id ≈⟨ refl⟩∘⟨ ⊗.identity ⟩ + id ⊗₁ λ⇐ ∘ id ≈⟨ identityʳ ⟩ + id ⊗₁ λ⇐ ∎ + + unitors : s ⊗₁ (t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ 𝒟.≈ α⇒ ∘ (s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ + unitors = begin + s ⊗₁ (t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ pushˡ split₂ʳ ⟩ + s ⊗₁ t ⊗₁ u ∘ id ⊗₁ ρ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ λ₁≅ρ₁⇐ ⟩∘⟨refl ⟨ + s ⊗₁ t ⊗₁ u ∘ id ⊗₁ λ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ triangle-to ⟨ + s ⊗₁ t ⊗₁ u ∘ α⇒ ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ extendʳ assoc-commute-from ⟨ + α⇒ ∘ (s ⊗₁ t) ⊗₁ u ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pushˡ split₁ʳ ⟨ + α⇒ ∘ (s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ∎ + + deco-assoc + : F[l∘m] ∘ (F[w,x] ∘ φ[N,R] ∘ s ⊗₁ (F[h,i] ∘ φ[M,P] ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐) + 𝒟.≈ F[y,z] ∘ φ[Q,P] ∘ (F[f,g] ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ + deco-assoc = begin + F[l∘m] ∘ F[w,x] ∘ φ[N,R] ∘ s ⊗₁ (F[h,i] ∘ φ[M,P] ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ pullˡ refl ⟩ + (F[l∘m] ∘ F[w,x]) ∘ φ[N,R] ∘ s ⊗₁ (F[h,i] ∘ φ[M,P] ∘ t ⊗₁ u ∘ ρ⇐) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ split₂ˡ ⟩∘⟨refl ⟩ + (F[l∘m] ∘ F[w,x]) ∘ φ[N,R] ∘ (id ⊗₁ F[h,i] ∘ s ⊗₁ (φ[M,P] ∘ t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ split₂ˡ) ⟩∘⟨refl ⟩ + (F[l∘m] ∘ F[w,x]) ∘ φ[N,R] ∘ (id ⊗₁ F[h,i] ∘ id ⊗₁ φ[M,P] ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc ⟩ + (F[l∘m] ∘ F[w,x]) ∘ φ[N,R] ∘ id ⊗₁ F[h,i] ∘ (id ⊗₁ φ[M,P] ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F-identity ⟩⊗⟨refl ⟩∘⟨refl ⟨ + (F[l∘m] ∘ F[w,x]) ∘ φ[N,R] ∘ F₁ id′ ⊗₁ F[h,i] ∘ (id ⊗₁ φ[M,P] ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (⊗-homo.commute (id′ , [ h , i ])) ⟩ + (F[l∘m] ∘ F[w,x]) ∘ F[id+[h,i]] ∘ φ[N,M+P] ∘ (id ⊗₁ φ[M,P] ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ pullˡ assoc ⟩ + (F[l∘m] ∘ F[w,x] ∘ F[id+[h,i]]) ∘ φ[N,M+P] ∘ (id ⊗₁ φ[M,P] ∘ s ⊗₁ (t ⊗₁ u ∘ ρ⇐)) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc ⟩ + (F[l∘m] ∘ F[w,x] ∘ F[id+[h,i]]) ∘ φ[N,M+P] ∘ id ⊗₁ φ[M,P] ∘ s⊗[t⊗u] ≈⟨ refl⟩∘⟨ sym-assoc ⟩ + (F[l∘m] ∘ F[w,x] ∘ F[id+[h,i]]) ∘ (φ[N,M+P] ∘ id ⊗₁ φ[M,P]) ∘ s⊗[t⊗u] ≈⟨ F-copairings ⟩∘⟨ coherences ⟩∘⟨ unitors ⟩ + (F[y,z] ∘ F[[f,g]+id] ∘ F₁ +-α⇒) ∘ (F₁ +-α⇐ ∘ φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ sym-assoc ⟩∘⟨ assoc ⟩ + ((F[y,z] ∘ F[[f,g]+id]) ∘ F₁ +-α⇒) ∘ F₁ +-α⇐ ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ assoc ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ F₁ +-α⇒ ∘ F₁ +-α⇐ ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ pushˡ homomorphism ⟨ + (F[y,z] ∘ F[[f,g]+id]) ∘ F₁ (+-α⇒ ∘′ +-α⇐) ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ F-resp-≈ +-assoc.isoʳ ⟩∘⟨refl ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ F₁ id′ ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ F-identity ⟩∘⟨refl ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ id ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ identityˡ ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ sym-assoc ⟩∘⟨refl ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ ((φ[N+M,P] ∘ φ[N,M] ⊗₁ id) ∘ α⇐) ∘ α⇒ ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ cancelInner 𝒟.associator.isoˡ ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ (φ[N+M,P] ∘ φ[N,M] ⊗₁ id) ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ assoc ⟩ + (F[y,z] ∘ F[[f,g]+id]) ∘ φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ [s⊗t]⊗u ≈⟨ assoc ⟩ + F[y,z] ∘ F[[f,g]+id] ∘ φ[N+M,P] ∘ φ[N,M] ⊗₁ id ∘ [s⊗t]⊗u ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟨ + F[y,z] ∘ F[[f,g]+id] ∘ φ[N+M,P] ∘ (φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (⊗-homo.commute ([ f , g ] , id′)) ⟨ + F[y,z] ∘ φ[Q,P] ∘ F[f,g] ⊗₁ F₁ id′ ∘ (φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ F-identity ⟩∘⟨refl ⟩ + F[y,z] ∘ φ[Q,P] ∘ F[f,g] ⊗₁ id ∘ (φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟨ + F[y,z] ∘ φ[Q,P] ∘ (F[f,g] ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ u ∘ ρ⇐ ∎ + +compose-idʳ : {C : DecoratedCospan A B} → compose identity C ≈ C compose-idʳ {A} {_} {C} = record { cospans-≈ = Cospans.compose-idʳ ; same-deco = deco-id } where - open DecoratedCospan C - - open 𝒞 using (pushout; [_,_]; ⊥; _+₁_; ¡) - + open 𝒞 using (pushout; [_,_]; ⊥; _+₁_; ¡; _+_) P = pushout 𝒞.id f₁ P′ = pushout-id-g {g = f₁} ≅P = up-to-iso P P′ - open Morphism 𝒞.U using (_≅_) module ≅P = _≅_ ≅P - open Pushout P - open 𝒞 using (cocartesian) renaming (id to id′; _∘_ to _∘′_) - open CocartesianMonoidal 𝒞.U cocartesian using (⊥+A≅A) - module ⊥+A≅A {a} = _≅_ (⊥+A≅A {a}) - module _ where - open 𝒞 using ( _⇒_ ; _∘_ ; _≈_ ; id ; U @@ -374,17 +415,13 @@ compose-idʳ {A} {_} {C} = record ; ∘[] ; []∘+₁ ; inject₂ ; assoc ; module HomReasoning ; module Dual ; module Equiv ) - open Equiv - open Dual.op-binaryProducts cocartesian using () renaming (⟨⟩-cong₂ to []-cong₂) - open ⇒-Reasoning U open HomReasoning - - copairing-id : ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (¡ +₁ id)) ∘ ⊥+A≅A.to ≈ id + copairing-id : ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (¡ +₁ id)) ∘ ⊥+A≅A.to 𝒞.≈ id copairing-id = begin ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (¡ +₁ id)) ∘ ⊥+A≅A.to ≈⟨ assoc ⟩ (≅P.from ∘ [ i₁ , i₂ ]) ∘ (¡ +₁ id) ∘ ⊥+A≅A.to ≈⟨ assoc ⟩ @@ -394,62 +431,54 @@ compose-idʳ {A} {_} {C} = record [ f₁ ∘ ¡ , id ∘ id ] ∘ ⊥+A≅A.to ≈⟨ []-cong₂ (sym (¡-unique (f₁ ∘ ¡))) identity² ⟩∘⟨refl ⟩ [ ¡ , id ] ∘ ⊥+A≅A.to ≈⟨ inject₂ ⟩ id ∎ - module _ where - open 𝒟 using ( id ; _∘_ ; _≈_ ; _⇒_ ; U ; assoc ; sym-assoc; identityˡ - ; monoidal ; _⊗₁_ ; unit ; unitorˡ ; unitorʳ + ; monoidal ; _⊗₀_; _⊗₁_ ; unit ; unitorˡ ; unitorʳ ) - open ⊗-Reasoning monoidal open ⇒-Reasoning U - - φ = ⊗-homo.η - φ-commute = ⊗-homo.commute - - module λ≅ = unitorˡ - λ⇒ = λ≅.from - λ⇐ = unitorˡ.to - ρ⇐ = unitorʳ.to - - open Coherence monoidal using (λ₁≅ρ₁⇐) + open ⊗-Util 𝒟.monoidal using (module Shorthands) + open Shorthands using (λ⇒; λ⇐; ρ⇐) + open Coherence 𝒟.monoidal using (λ₁≅ρ₁⇐) open 𝒟.Equiv - s : unit ⇒ F₀ N s = decoration - - cohere-s : φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈ F₁ ⊥+A≅A.to ∘ s - cohere-s = begin - φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ identityˡ ⟨ - id ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ F-identity ⟩∘⟨refl ⟨ - F₁ id′ ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ F-resp-≈ ⊥+A≅A.isoˡ ⟩∘⟨refl ⟨ - F₁ (⊥+A≅A.to ∘′ ⊥+A≅A.from) ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ - F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₁₂ ⟩ - F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ φ (⊥ , N) ∘ (ε ⊗₁ id) ∘ (id ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩ - F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ (φ (⊥ , N) ∘ (ε ⊗₁ id)) ∘ (id ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ unitaryˡ ⟩ - F₁ ⊥+A≅A.to ∘ λ⇒ ∘ (id ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ λ₁≅ρ₁⇐ ⟨ - F₁ ⊥+A≅A.to ∘ λ⇒ ∘ (id ⊗₁ s) ∘ λ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ 𝒟.unitorˡ-commute-to ⟨ - F₁ ⊥+A≅A.to ∘ λ⇒ ∘ λ⇐ ∘ s ≈⟨ refl⟩∘⟨ cancelˡ λ≅.isoʳ ⟩ - F₁ ⊥+A≅A.to ∘ s ∎ - - deco-id : F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ (A , N) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈ s - deco-id = begin - F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ (A , N) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (A , N) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟩ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (A , N) ∘ (F₁ ¡ ⊗₁ id) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ F-identity ⟩∘⟨refl ⟨ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (A , N) ∘ (F₁ ¡ ⊗₁ F₁ id′) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (¡ , id′)) ⟩ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ F₁ (¡ +₁ id′) ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ - F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘ φ (⊥ , N) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ cohere-s ⟩ - F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘ F₁ ⊥+A≅A.to ∘ s ≈⟨ pushˡ homomorphism ⟨ - F₁ (((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘′ ⊥+A≅A.to) ∘ s ≈⟨ F-resp-≈ copairing-id ⟩∘⟨refl ⟩ - F₁ id′ ∘ s ≈⟨ F-identity ⟩∘⟨refl ⟩ - id ∘ s ≈⟨ identityˡ ⟩ - s ∎ - -compose-idˡ : {C : DecoratedCospan A B} → Same (compose C identity) C + φ[⊥,N] : F₀ ⊥ ⊗₀ F₀ N ⇒ F₀ (⊥ + N) + φ[⊥,N] = ⊗-homo.η (⊥ , N) + φ[A,N] : F₀ A ⊗₀ F₀ N ⇒ F₀ (A + N) + φ[A,N] = ⊗-homo.η (A , N) + abstract + cohere-s : φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ 𝒟.≈ F₁ ⊥+A≅A.to ∘ s + cohere-s = begin + φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ identityˡ ⟨ + id ∘ φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ F-identity ⟩∘⟨refl ⟨ + F₁ id′ ∘ φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ F-resp-≈ ⊥+A≅A.isoˡ ⟩∘⟨refl ⟨ + F₁ (⊥+A≅A.to ∘′ ⊥+A≅A.from) ∘ φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ + F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ φ[⊥,N] ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₁₂ ⟩ + F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ φ[⊥,N] ∘ ε ⊗₁ id ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩ + F₁ ⊥+A≅A.to ∘ F₁ ⊥+A≅A.from ∘ (φ[⊥,N] ∘ ε ⊗₁ id) ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ unitaryˡ ⟩ + F₁ ⊥+A≅A.to ∘ λ⇒ ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ λ₁≅ρ₁⇐ ⟨ + F₁ ⊥+A≅A.to ∘ λ⇒ ∘ id ⊗₁ s ∘ λ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ 𝒟.unitorˡ-commute-to ⟨ + F₁ ⊥+A≅A.to ∘ λ⇒ ∘ λ⇐ ∘ s ≈⟨ refl⟩∘⟨ cancelˡ 𝒟.unitorˡ.isoʳ ⟩ + F₁ ⊥+A≅A.to ∘ s ∎ + deco-id : F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ[A,N] ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ 𝒟.≈ s + deco-id = begin + F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ[A,N] ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[A,N] ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟩ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[A,N] ∘ (F₁ ¡ ⊗₁ id) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ F-identity ⟩∘⟨refl ⟨ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[A,N] ∘ (F₁ ¡ ⊗₁ F₁ id′) ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (⊗-homo.commute (¡ , id′)) ⟩ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ F₁ (¡ +₁ id′) ∘ φ[⊥,N] ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ + F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘ φ[⊥,N] ∘ (ε ⊗₁ s) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ cohere-s ⟩ + F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘ F₁ ⊥+A≅A.to ∘ s ≈⟨ pushˡ homomorphism ⟨ + F₁ (((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (¡ +₁ id′)) ∘′ ⊥+A≅A.to) ∘ s ≈⟨ F-resp-≈ copairing-id ⟩∘⟨refl ⟩ + F₁ id′ ∘ s ≈⟨ F-identity ⟩∘⟨refl ⟩ + id ∘ s ≈⟨ identityˡ ⟩ + s ∎ + +compose-idˡ : {C : DecoratedCospan A B} → compose C identity ≈ C compose-idˡ {_} {B} {C} = record { cospans-≈ = Cospans.compose-idˡ ; same-deco = deco-id @@ -497,7 +526,7 @@ compose-idˡ {_} {B} {C} = record open ⇒-Reasoning U open HomReasoning - copairing-id : ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (id +₁ ¡)) ∘ A+⊥≅A.to ≈ id + copairing-id : ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (id +₁ ¡)) ∘ A+⊥≅A.to 𝒞.≈ id copairing-id = begin ((≅P.from ∘ [ i₁ , i₂ ]) ∘ (id +₁ ¡)) ∘ A+⊥≅A.to ≈⟨ assoc ⟩ (≅P.from ∘ [ i₁ , i₂ ]) ∘ (id +₁ ¡) ∘ A+⊥≅A.to ≈⟨ assoc ⟩ @@ -510,11 +539,12 @@ compose-idˡ {_} {B} {C} = record module _ where + open 𝒞 using (_+_) open 𝒟 using ( id ; _∘_ ; _≈_ ; _⇒_ ; U ; assoc ; sym-assoc; identityˡ - ; monoidal ; _⊗₁_ ; unit ; unitorˡ ; unitorʳ + ; monoidal ; _⊗₀_; _⊗₁_ ; unit ; unitorˡ ; unitorʳ ; unitorʳ-commute-to ; module Equiv ) @@ -523,129 +553,146 @@ compose-idˡ {_} {B} {C} = record open ⊗-Reasoning monoidal open ⇒-Reasoning U - φ = ⊗-homo.η - φ-commute = ⊗-homo.commute + φ[N,⊥] : F₀ N ⊗₀ F₀ ⊥ 𝒟.⇒ F₀ (N + ⊥) + φ[N,⊥] = ⊗-homo.η (N , ⊥) - module ρ≅ = unitorʳ - ρ⇒ = ρ≅.from - ρ⇐ = ρ≅.to + φ[N,B] : F₀ N ⊗₀ F₀ B 𝒟.⇒ F₀ (N + B) + φ[N,B] = ⊗-homo.η (N , B) s : unit ⇒ F₀ N s = decoration - cohere-s : φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈ F₁ A+⊥≅A.to ∘ s - cohere-s = begin - φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ identityˡ ⟨ - id ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ F-identity ⟩∘⟨refl ⟨ - F₁ id′ ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ F-resp-≈ A+⊥≅A.isoˡ ⟩∘⟨refl ⟨ - F₁ (A+⊥≅A.to ∘′ A+⊥≅A.from) ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ - F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₂₁ ⟩ - F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ φ (N , ⊥) ∘ (id ⊗₁ ε) ∘ (s ⊗₁ id) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩ - F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ (φ (N , ⊥) ∘ (id ⊗₁ ε)) ∘ (s ⊗₁ id) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ unitaryʳ ⟩ - F₁ A+⊥≅A.to ∘ ρ⇒ ∘ (s ⊗₁ id) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ unitorʳ-commute-to ⟨ - F₁ A+⊥≅A.to ∘ ρ⇒ ∘ ρ⇐ ∘ s ≈⟨ refl⟩∘⟨ cancelˡ ρ≅.isoʳ ⟩ - F₁ A+⊥≅A.to ∘ s ∎ - - deco-id : F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈ s - deco-id = begin - F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₂ˡ ⟩ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (N , B) ∘ (id ⊗₁ F₁ ¡) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F-identity ⟩⊗⟨refl ⟩∘⟨refl ⟨ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ (N , B) ∘ (F₁ id′ ⊗₁ F₁ ¡) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (id′ , ¡)) ⟩ - F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ F₁ (id′ +₁ ¡) ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ - F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘ φ (N , ⊥) ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ cohere-s ⟩ - F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘ F₁ A+⊥≅A.to ∘ s ≈⟨ pushˡ homomorphism ⟨ - F₁ (((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘′ A+⊥≅A.to) ∘ s ≈⟨ F-resp-≈ copairing-id ⟩∘⟨refl ⟩ - F₁ id′ ∘ s ≈⟨ F-identity ⟩∘⟨refl ⟩ - id ∘ s ≈⟨ identityˡ ⟩ - s ∎ - -compose-id² : Same {A} (compose identity identity) identity + open ⊗-Util 𝒟.monoidal using (module Shorthands) + open Shorthands using (α⇒; α⇐; λ⇒; λ⇐; ρ⇒; ρ⇐) + + abstract + cohere-s : φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ 𝒟.≈ F₁ A+⊥≅A.to ∘ s + cohere-s = begin + φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ identityˡ ⟨ + id ∘ φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ F-identity ⟩∘⟨refl ⟨ + F₁ id′ ∘ φ[N,⊥] ∘ (s ⊗₁ ε) ∘ ρ⇐ ≈⟨ F-resp-≈ A+⊥≅A.isoˡ ⟩∘⟨refl ⟨ + F₁ (A+⊥≅A.to ∘′ A+⊥≅A.from) ∘ φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ + F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₂₁ ⟩ + F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ φ[N,⊥] ∘ id ⊗₁ ε ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩ + F₁ A+⊥≅A.to ∘ F₁ A+⊥≅A.from ∘ (φ[N,⊥] ∘ id ⊗₁ ε) ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ unitaryʳ ⟩ + F₁ A+⊥≅A.to ∘ ρ⇒ ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ unitorʳ-commute-to ⟨ + F₁ A+⊥≅A.to ∘ ρ⇒ ∘ ρ⇐ ∘ s ≈⟨ refl⟩∘⟨ cancelˡ unitorʳ.isoʳ ⟩ + F₁ A+⊥≅A.to ∘ s ∎ + + deco-id : F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ[N,B] ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ 𝒟.≈ s + deco-id = begin + F₁ ≅P.from ∘ F₁ [ i₁ , i₂ ] ∘ φ[N,B] ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[N,B] ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₂ˡ ⟩ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[N,B] ∘ id ⊗₁ F₁ ¡ ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ F-identity ⟩⊗⟨refl ⟩∘⟨refl ⟨ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ φ[N,B] ∘ F₁ id′ ⊗₁ F₁ ¡ ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (⊗-homo.commute (id′ , ¡)) ⟩ + F₁ (≅P.from ∘′ [ i₁ , i₂ ]) ∘ F₁ (id′ +₁ ¡) ∘ φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ + F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘ φ[N,⊥] ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ cohere-s ⟩ + F₁ ((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘ F₁ A+⊥≅A.to ∘ s ≈⟨ pushˡ homomorphism ⟨ + F₁ (((≅P.from ∘′ [ i₁ , i₂ ]) ∘′ (id′ +₁ ¡)) ∘′ A+⊥≅A.to) ∘ s ≈⟨ F-resp-≈ copairing-id ⟩∘⟨refl ⟩ + F₁ id′ ∘ s ≈⟨ F-identity ⟩∘⟨refl ⟩ + id ∘ s ≈⟨ identityˡ ⟩ + s ∎ + +compose-id² : compose identity identity ≈ identity {A} compose-id² = compose-idˡ compose-equiv : {c₂ c₂′ : DecoratedCospan B C} {c₁ c₁′ : DecoratedCospan A B} - → Same c₂ c₂′ - → Same c₁ c₁′ - → Same (compose c₁ c₂) (compose c₁′ c₂′) + → c₂ ≈ c₂′ + → c₁ ≈ c₁′ + → compose c₁ c₂ ≈ (compose c₁′ c₂′) compose-equiv {_} {_} {_} {c₂} {c₂′} {c₁} {c₁′} ≅C₂ ≅C₁ = record { cospans-≈ = ≅C₂∘C₁ ; same-deco = F≅N∘C₂∘C₁≈C₂′∘C₁′ } where - module ≅C₁ = Same ≅C₁ - module ≅C₂ = Same ≅C₂ + module ≅C₁ = _≈_ ≅C₁ + module ≅C₂ = _≈_ ≅C₂ module C₁ = DecoratedCospan c₁ module C₁′ = DecoratedCospan c₁′ module C₂ = DecoratedCospan c₂ module C₂′ = DecoratedCospan c₂′ ≅C₂∘C₁ = Cospans.compose-equiv ≅C₂.cospans-≈ ≅C₁.cospans-≈ - module ≅C₂∘C₁ = Cospans.Same ≅C₂∘C₁ + module ≅C₂∘C₁ = Cospan._≈_ ≅C₂∘C₁ P = 𝒞.pushout C₁.f₂ C₂.f₁ P′ = 𝒞.pushout C₁′.f₂ C₂′.f₁ module P = Pushout P module P′ = Pushout P′ - s = C₁.decoration - t = C₂.decoration - s′ = C₁′.decoration - t′ = C₂′.decoration + N M N′ M′ : 𝒞.Obj N = C₁.N M = C₂.N N′ = C₁′.N M′ = C₂′.N - φ = ⊗-homo.η - φ-commute = ⊗-homo.commute + s : 𝒟.unit 𝒟.⇒ F₀ N + s = C₁.decoration + t : 𝒟.unit 𝒟.⇒ F₀ M + t = C₂.decoration + + s′ : 𝒟.unit 𝒟.⇒ F₀ N′ + s′ = C₁′.decoration + + t′ : 𝒟.unit 𝒟.⇒ F₀ M′ + t′ = C₂′.decoration + + Q⇒ : ≅C₂∘C₁.C.N 𝒞.⇒ ≅C₂∘C₁.D.N Q⇒ = ≅C₂∘C₁.≅N.from + + N⇒ : ≅C₁.C.N 𝒞.⇒ ≅C₁.D.N N⇒ = ≅C₁.≅N.from + + M⇒ : ≅C₂.C.N 𝒞.⇒ ≅C₂.D.N M⇒ = ≅C₂.≅N.from module _ where - ρ⇒ = 𝒟.unitorʳ.from - ρ⇐ = 𝒟.unitorʳ.to + open ⊗-Util 𝒟.monoidal using (module Shorthands) + open Shorthands using (ρ⇒; ρ⇐) - open 𝒞 using ([_,_]; ∘[]; _+₁_; []∘+₁) renaming (_∘_ to _∘′_) + open 𝒞 using ([_,_]; ∘[]; _+_; _+₁_; []∘+₁) renaming (_∘_ to _∘′_) open 𝒞.Dual.op-binaryProducts 𝒞.cocartesian using () renaming (⟨⟩-cong₂ to []-cong₂) open 𝒟 + φ[N,M] : F₀ N ⊗₀ F₀ M 𝒟.⇒ F₀ (N + M) + φ[N,M] = ⊗-homo.η (N , M) + + φ[N′,M′] : F₀ N′ ⊗₀ F₀ M′ 𝒟.⇒ F₀ (N′ + M′) + φ[N′,M′] = ⊗-homo.η (N′ , M′) + open ⊗-Reasoning monoidal open ⇒-Reasoning U - F≅N∘C₂∘C₁≈C₂′∘C₁′ : F₁ Q⇒ ∘ F₁ [ P.i₁ , P.i₂ ] ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈ F₁ [ P′.i₁ , P′.i₂ ] ∘ φ (N′ , M′) ∘ s′ ⊗₁ t′ ∘ ρ⇐ - F≅N∘C₂∘C₁≈C₂′∘C₁′ = begin - F₁ Q⇒ ∘ F₁ [ P.i₁ , P.i₂ ] ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ - F₁ (Q⇒ ∘′ [ P.i₁ , P.i₂ ]) ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ ∘[] ⟩∘⟨refl ⟩ - F₁ ([ Q⇒ ∘′ P.i₁ , Q⇒ ∘′ P.i₂ ]) ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ ([]-cong₂ P.universal∘i₁≈h₁ P.universal∘i₂≈h₂) ⟩∘⟨refl ⟩ - F₁ ([ P′.i₁ ∘′ N⇒ , P′.i₂ ∘′ M⇒ ]) ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ []∘+₁ ⟩∘⟨refl ⟨ - F₁ ([ P′.i₁ , P′.i₂ ] ∘′ (N⇒ +₁ M⇒)) ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ - F₁ [ P′.i₁ , P′.i₂ ] ∘ F₁ (N⇒ +₁ M⇒) ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (N⇒ , M⇒)) ⟨ - F₁ [ P′.i₁ , P′.i₂ ] ∘ φ (N′ , M′) ∘ F₁ N⇒ ⊗₁ F₁ M⇒ ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ ⊗-distrib-over-∘ ⟨ - F₁ [ P′.i₁ , P′.i₂ ] ∘ φ (N′ , M′) ∘ (F₁ N⇒ ∘ s) ⊗₁ (F₁ M⇒ ∘ t) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ≅C₁.same-deco ⟩⊗⟨ ≅C₂.same-deco ⟩∘⟨refl ⟩ - F₁ [ P′.i₁ , P′.i₂ ] ∘ φ (N′ , M′) ∘ s′ ⊗₁ t′ ∘ ρ⇐ ∎ + abstract + F≅N∘C₂∘C₁≈C₂′∘C₁′ : F₁ Q⇒ ∘ F₁ [ P.i₁ , P.i₂ ] ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ 𝒟.≈ F₁ [ P′.i₁ , P′.i₂ ] ∘ φ[N′,M′] ∘ s′ ⊗₁ t′ ∘ ρ⇐ + F≅N∘C₂∘C₁≈C₂′∘C₁′ = begin + F₁ Q⇒ ∘ F₁ [ P.i₁ , P.i₂ ] ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟨ + F₁ (Q⇒ ∘′ [ P.i₁ , P.i₂ ]) ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ ∘[] ⟩∘⟨refl ⟩ + F₁ ([ Q⇒ ∘′ P.i₁ , Q⇒ ∘′ P.i₂ ]) ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ ([]-cong₂ P.universal∘i₁≈h₁ P.universal∘i₂≈h₂) ⟩∘⟨refl ⟩ + F₁ ([ P′.i₁ ∘′ N⇒ , P′.i₂ ∘′ M⇒ ]) ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ F-resp-≈ []∘+₁ ⟩∘⟨refl ⟨ + F₁ ([ P′.i₁ , P′.i₂ ] ∘′ (N⇒ +₁ M⇒)) ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ pushˡ homomorphism ⟩ + F₁ [ P′.i₁ , P′.i₂ ] ∘ F₁ (N⇒ +₁ M⇒) ∘ φ[N,M] ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (⊗-homo.commute (N⇒ , M⇒)) ⟨ + F₁ [ P′.i₁ , P′.i₂ ] ∘ φ[N′,M′] ∘ F₁ N⇒ ⊗₁ F₁ M⇒ ∘ s ⊗₁ t ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ ⊗-distrib-over-∘ ⟨ + F₁ [ P′.i₁ , P′.i₂ ] ∘ φ[N′,M′] ∘ (F₁ N⇒ ∘ s) ⊗₁ (F₁ M⇒ ∘ t) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ≅C₁.same-deco ⟩⊗⟨ ≅C₂.same-deco ⟩∘⟨refl ⟩ + F₁ [ P′.i₁ , P′.i₂ ] ∘ φ[N′,M′] ∘ s′ ⊗₁ t′ ∘ ρ⇐ ∎ DecoratedCospans : Category o (o ⊔ ℓ ⊔ ℓ′) (ℓ ⊔ e ⊔ e′) DecoratedCospans = record { Obj = 𝒞.Obj ; _⇒_ = DecoratedCospan - ; _≈_ = Same + ; _≈_ = _≈_ ; id = identity ; _∘_ = flip compose ; assoc = compose-assoc - ; sym-assoc = same-sym (compose-assoc) + ; sym-assoc = ≈-sym compose-assoc ; identityˡ = compose-idˡ ; identityʳ = compose-idʳ ; identity² = compose-id² - ; equiv = record - { refl = same-refl - ; sym = same-sym - ; trans = same-trans - } + ; equiv = ≈-equiv ; ∘-resp-≈ = compose-equiv } diff --git a/Category/Instance/FinitelyCocompletes.agda b/Category/Instance/FinitelyCocompletes.agda index 0847165..9bee58e 100644 --- a/Category/Instance/FinitelyCocompletes.agda +++ b/Category/Instance/FinitelyCocompletes.agda @@ -1,29 +1,31 @@ {-# OPTIONS --without-K --safe #-} -open import Level using (Level) + +open import Level using (Level; suc; _⊔_) + module Category.Instance.FinitelyCocompletes {o ℓ e : Level} where +import Category.Instance.One.Properties as OneProps + open import Categories.Category using (_[_,_]) open import Categories.Category.BinaryProducts using (BinaryProducts) open import Categories.Category.Cartesian using (Cartesian) -open import Categories.Category.Helper using (categoryHelper) -open import Categories.Category.Monoidal.Instance.Cats using () renaming (module Product to Products) open import Categories.Category.Core using (Category) +open import Categories.Category.Helper using (categoryHelper) open import Categories.Category.Instance.Cats using (Cats) open import Categories.Category.Instance.One using (One; One-⊤) +open import Categories.Category.Monoidal.Instance.Cats using () renaming (module Product to Products) open import Categories.Category.Product using (πˡ; πʳ; _※_; _⁂_) renaming (Product to ProductCat) open import Categories.Diagram.Coequalizer using (IsCoequalizer) open import Categories.Functor using (Functor) renaming (id to idF) +open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_; associator; unitorˡ; unitorʳ; module ≃; _ⓘₕ_) open import Categories.Object.Coproduct using (IsCoproduct) open import Categories.Object.Initial using (IsInitial) open import Categories.Object.Product.Core using (Product) -open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_; associator; unitorˡ; unitorʳ; module ≃; _ⓘₕ_) open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) open import Category.Cocomplete.Finitely.Product using (FinitelyCocomplete-×) -open import Category.Instance.One.Properties using (One-FinitelyCocomplete) open import Data.Product.Base using (_,_; proj₁; proj₂; map; dmap; zip′) -open import Functor.Exact using (∘-RightExactFunctor; RightExactFunctor; idREF; IsRightExact; rightexact) open import Function.Base using (id; flip) -open import Level using (Level; suc; _⊔_) +open import Functor.Exact using (∘-RightExactFunctor; RightExactFunctor; idREF; IsRightExact; rightexact) FinitelyCocompletes : Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) FinitelyCocompletes = categoryHelper record @@ -48,7 +50,7 @@ FinitelyCocompletes = categoryHelper record One-FCC : FinitelyCocompleteCategory o ℓ e One-FCC = record { U = One - ; finCo = One-FinitelyCocomplete + ; finCo = OneProps.finitelyCocomplete } _×_ diff --git a/Category/Instance/One/Properties.agda b/Category/Instance/One/Properties.agda index 1452669..c6261bb 100644 --- a/Category/Instance/One/Properties.agda +++ b/Category/Instance/One/Properties.agda @@ -8,7 +8,7 @@ open import Categories.Category.Core using (Category) open import Categories.Category.Instance.One using () renaming (One to One′) One : Category o ℓ e -One = One′ +One = One′ {o} {ℓ} {e} open import Categories.Category.Cocartesian One using (Cocartesian) open import Categories.Category.Cocomplete.Finitely One using (FinitelyCocomplete) @@ -16,20 +16,20 @@ open import Categories.Object.Initial One using (Initial) open import Categories.Category.Cocartesian One using (BinaryCoproducts) -One-Initial : Initial -One-Initial = _ +initial : Initial +initial = _ -One-BinaryCoproducts : BinaryCoproducts -One-BinaryCoproducts = _ +coproducts : BinaryCoproducts +coproducts = _ -One-Cocartesian : Cocartesian -One-Cocartesian = record - { initial = One-Initial - ; coproducts = One-BinaryCoproducts +cocartesian : Cocartesian +cocartesian = record + { initial = initial + ; coproducts = coproducts } -One-FinitelyCocomplete : FinitelyCocomplete -One-FinitelyCocomplete = record - { cocartesian = One-Cocartesian +finitelyCocomplete : FinitelyCocomplete +finitelyCocomplete = record + { cocartesian = cocartesian ; coequalizer = _ } diff --git a/Category/Instance/Properties/SymMonCat.agda b/Category/Instance/Properties/SymMonCat.agda new file mode 100644 index 0000000..fa15295 --- /dev/null +++ b/Category/Instance/Properties/SymMonCat.agda @@ -0,0 +1,166 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --lossy-unification #-} + +open import Level using (Level; suc; _⊔_) +module Category.Instance.Properties.SymMonCat {o ℓ e : Level} where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.NaturalTransformation.NaturalIsomorphism.Monoidal.Symmetric as SMNI +import Categories.Functor.Monoidal.Symmetric {o} {o} {ℓ} {ℓ} {e} {e} as SMF + +open import Category.Instance.SymMonCat {o} {ℓ} {e} using (SymMonCat) + +open import Categories.Category using (Category; _[_≈_]; _[_∘_]) +open import Categories.Object.Product.Core SymMonCat using (Product) +open import Categories.Object.Terminal SymMonCat using (Terminal) +open import Categories.Category.Instance.One using (One) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Category.Cartesian SymMonCat using (Cartesian) +open import Categories.Category.Cartesian.Bundle using (CartesianCategory) +open import Categories.Functor.Monoidal.Properties using (idF-SymmetricMonoidal; ∘-SymmetricMonoidal) +open import Categories.Category.BinaryProducts SymMonCat using (BinaryProducts) +open import Categories.Functor.Monoidal.Construction.Product + using () + renaming + ( πˡ-SymmetricMonoidalFunctor to πˡ-SMF + ; πʳ-SymmetricMonoidalFunctor to πʳ-SMF + ; ※-SymmetricMonoidalFunctor to ※-SMF + ) +open import Categories.Category.Monoidal.Construction.Product using (Product-SymmetricMonoidalCategory) +open import Categories.Category.Product.Properties using () renaming (project₁ to p₁; project₂ to p₂; unique to u) +open import Data.Product.Base using (_,_; proj₁; proj₂) + +open SMF.Lax using (SymmetricMonoidalFunctor) +open SMNI.Lax using (SymmetricMonoidalNaturalIsomorphism; id; isEquivalence) + +module Cone + {A B X : SymmetricMonoidalCategory o ℓ e} + {F : SymmetricMonoidalFunctor X A} + {G : SymmetricMonoidalFunctor X B} where + + module A = SymmetricMonoidalCategory A + module B = SymmetricMonoidalCategory B + module X = SymmetricMonoidalCategory X + module F = SymmetricMonoidalFunctor X A F + module G = SymmetricMonoidalFunctor X B G + + A×B : SymmetricMonoidalCategory o ℓ e + A×B = (Product-SymmetricMonoidalCategory A B) + + πˡ : SymmetricMonoidalFunctor A×B A + πˡ = πˡ-SMF {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} + + πʳ : SymmetricMonoidalFunctor A×B B + πʳ = πʳ-SMF {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} + + module _ where + open Category A.U + open Equiv + open ⇒-Reasoning A.U + open ⊗-Reasoning A.monoidal + project₁ : SymMonCat [ SymMonCat [ πˡ ∘ ※-SMF F G ] ≈ F ] + project₁ = record + { U = p₁ {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {X.U} {F.F} {G.F} + ; F⇒G-isMonoidal = record + { ε-compat = identityˡ ○ identityʳ + ; ⊗-homo-compat = λ { {C} {D} → identityˡ ○ refl⟩∘⟨ sym A.⊗.identity } + } + } + module _ (H : SymmetricMonoidalFunctor X A×B) (eq₁ : SymMonCat [ SymMonCat [ πˡ ∘ H ] ≈ F ]) where + private + module H = SymmetricMonoidalFunctor X A×B H + open SymmetricMonoidalNaturalIsomorphism eq₁ + ε-compat₁ : ⇐.η X.unit A.∘ F.ε A.≈ H.ε .proj₁ + ε-compat₁ = refl⟩∘⟨ sym ε-compat ○ cancelˡ (iso.isoˡ X.unit) ○ identityʳ + ⊗-homo-compat₁ + : ∀ {C D} + → ⇐.η (X.⊗.₀ (C , D)) ∘ F.⊗-homo.η (C , D) + ≈ H.⊗-homo.η (C , D) .proj₁ ∘ A.⊗.₁ (⇐.η C , ⇐.η D) + ⊗-homo-compat₁ {C} {D} = + insertʳ + (sym ⊗-distrib-over-∘ + ○ iso.isoʳ C ⟩⊗⟨ iso.isoʳ D + ○ A.⊗.identity) + ○ (pullʳ (sym ⊗-homo-compat) + ○ cancelˡ (iso.isoˡ (X.⊗.₀ (C , D))) + ○ identityʳ) ⟩∘⟨refl + + module _ where + open Category B.U + open Equiv + open ⇒-Reasoning B.U + open ⊗-Reasoning B.monoidal + project₂ : SymMonCat [ SymMonCat [ πʳ ∘ ※-SMF F G ] ≈ G ] + project₂ = record + { U = p₂ {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {X.U} {F.F} {G.F} + ; F⇒G-isMonoidal = record + { ε-compat = identityˡ ○ identityʳ + ; ⊗-homo-compat = λ { {C} {D} → identityˡ ○ refl⟩∘⟨ sym B.⊗.identity } + } + } + module _ (H : SymmetricMonoidalFunctor X A×B) (eq₂ : SymMonCat [ SymMonCat [ πʳ ∘ H ] ≈ G ]) where + private + module H = SymmetricMonoidalFunctor X A×B H + open SymmetricMonoidalNaturalIsomorphism eq₂ + ε-compat₂ : ⇐.η X.unit ∘ G.ε ≈ H.ε .proj₂ + ε-compat₂ = refl⟩∘⟨ sym ε-compat ○ cancelˡ (iso.isoˡ X.unit) ○ identityʳ + ⊗-homo-compat₂ + : ∀ {C D} + → ⇐.η (X.⊗.₀ (C , D)) ∘ G.⊗-homo.η (C , D) + ≈ H.⊗-homo.η (C , D) .proj₂ ∘ B.⊗.₁ (⇐.η C , ⇐.η D) + ⊗-homo-compat₂ {C} {D} = + insertʳ + (sym ⊗-distrib-over-∘ + ○ iso.isoʳ C ⟩⊗⟨ iso.isoʳ D + ○ B.⊗.identity) + ○ (pullʳ (sym ⊗-homo-compat) + ○ cancelˡ (iso.isoˡ (X.⊗.₀ (C , D))) + ○ identityʳ) ⟩∘⟨refl + + unique + : (H : SymmetricMonoidalFunctor X A×B) + → SymMonCat [ SymMonCat [ πˡ ∘ H ] ≈ F ] + → SymMonCat [ SymMonCat [ πʳ ∘ H ] ≈ G ] + → SymMonCat [ ※-SMF F G ≈ H ] + unique H eq₁ eq₂ = record + { U = u {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {X.U} {F.F} {G.F} {H.F} eq₁.U eq₂.U + ; F⇒G-isMonoidal = record + { ε-compat = ε-compat₁ H eq₁ , ε-compat₂ H eq₂ + ; ⊗-homo-compat = ⊗-homo-compat₁ H eq₁ , ⊗-homo-compat₂ H eq₂ + } + } + where + module H = SymmetricMonoidalFunctor X A×B H + module eq₁ = SymmetricMonoidalNaturalIsomorphism eq₁ + module eq₂ = SymmetricMonoidalNaturalIsomorphism eq₂ + +prod-SymMonCat : ∀ {A B} → Product A B +prod-SymMonCat {A} {B} = record + { A×B = Product-SymmetricMonoidalCategory A B + ; π₁ = πˡ-SMF {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} + ; π₂ = πʳ-SMF {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} + ; ⟨_,_⟩ = ※-SMF + ; project₁ = λ { {X} {f} {g} → Cone.project₁ {A} {B} {X} {f} {g} } + ; project₂ = λ { {X} {f} {g} → Cone.project₂ {A} {B} {X} {f} {g} } + ; unique = λ { {X} {h} {f} {g} eq₁ eq₂ → Cone.unique {A} {B} {X} {f} {g} h eq₁ eq₂ } + } + +SymMonCat-BinaryProducts : BinaryProducts +SymMonCat-BinaryProducts = record { product = prod-SymMonCat } + +SymMonCat-Terminal : Terminal +SymMonCat-Terminal = record + { ⊤ = record + { U = One + ; monoidal = _ + ; symmetric = _ + } + ; ⊤-is-terminal = _ + } + +SymMonCat-Cartesian : Cartesian +SymMonCat-Cartesian = record + { terminal = SymMonCat-Terminal + ; products = SymMonCat-BinaryProducts + } diff --git a/Category/Instance/Setoids/SymmetricMonoidal.agda b/Category/Instance/Setoids/SymmetricMonoidal.agda index fa4d903..995ddf3 100644 --- a/Category/Instance/Setoids/SymmetricMonoidal.agda +++ b/Category/Instance/Setoids/SymmetricMonoidal.agda @@ -1,33 +1,47 @@ {-# OPTIONS --without-K --safe #-} -module Category.Instance.Setoids.SymmetricMonoidal {ℓ} where +open import Level using (Level; _⊔_; suc) +module Category.Instance.Setoids.SymmetricMonoidal {c ℓ : Level} where -open import Categories.Category.Cartesian.Monoidal using (module CartesianMonoidal) open import Categories.Category.Instance.Setoids using (Setoids) open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian; Setoids-Cocartesian) renaming (Setoids-Monoidal to ×-monoidal) -open import Categories.Category.Cartesian.SymmetricMonoidal (Setoids ℓ ℓ) Setoids-Cartesian +open import Categories.Category.Cartesian.SymmetricMonoidal (Setoids c ℓ) Setoids-Cartesian using () renaming (symmetric to ×-symmetric) -open import Level using (suc) -open import Categories.Category.Cocartesian (Setoids ℓ ℓ) +open import Categories.Category.Cocartesian (Setoids c (c ⊔ ℓ)) using (module CocartesianMonoidal; module CocartesianSymmetricMonoidal) -open CocartesianMonoidal (Setoids-Cocartesian {ℓ} {ℓ}) using (+-monoidal) -open CocartesianSymmetricMonoidal (Setoids-Cocartesian {ℓ} {ℓ}) using (+-symmetric) +open CocartesianMonoidal (Setoids-Cocartesian {c} {ℓ}) using (+-monoidal) +open CocartesianSymmetricMonoidal (Setoids-Cocartesian {c} {ℓ}) using (+-symmetric) -Setoids-× : SymmetricMonoidalCategory (suc ℓ) ℓ ℓ +open import Categories.Category using (Category) +open import Categories.Category.Monoidal using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) + +opaque + + ×-monoidal′ : Monoidal (Setoids c ℓ) + ×-monoidal′ = ×-monoidal {c} {ℓ} + + ×-symmetric′ : Symmetric ×-monoidal′ + ×-symmetric′ = ×-symmetric + +Setoids-× : SymmetricMonoidalCategory (suc (c ⊔ ℓ)) (c ⊔ ℓ) (c ⊔ ℓ) Setoids-× = record - { U = Setoids ℓ ℓ - ; monoidal = ×-monoidal - ; symmetric = ×-symmetric + { U = Setoids c ℓ + ; monoidal = ×-monoidal′ + ; symmetric = ×-symmetric′ } -Setoids-+ : SymmetricMonoidalCategory (suc ℓ) ℓ ℓ +Setoids-+ : SymmetricMonoidalCategory (suc (c ⊔ ℓ)) (c ⊔ ℓ) (c ⊔ ℓ) Setoids-+ = record - { U = Setoids ℓ ℓ + { U = Setoids c (c ⊔ ℓ) ; monoidal = +-monoidal ; symmetric = +-symmetric } + +module Setoids-× = SymmetricMonoidalCategory Setoids-× +module Setoids-+ = SymmetricMonoidalCategory Setoids-+ diff --git a/Category/Instance/SymMonCat.agda b/Category/Instance/SymMonCat.agda new file mode 100644 index 0000000..e4b136c --- /dev/null +++ b/Category/Instance/SymMonCat.agda @@ -0,0 +1,74 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --lossy-unification #-} + +open import Level using (Level; suc; _⊔_) +module Category.Instance.SymMonCat {o ℓ e : Level} where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Functor.Monoidal.Symmetric as SMF +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Morphism as Morphism +import Categories.NaturalTransformation.NaturalIsomorphism.Monoidal.Symmetric as SMNI +import Categories.Category.Monoidal.Utilities as MonoidalUtil +import Categories.Category.Monoidal.Braided.Properties as BraidedProperties +open import Relation.Binary.Core using (Rel) + +open import Categories.Category using (Category; _[_,_]; _[_∘_]) +open import Categories.Category.Helper using (categoryHelper) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Properties using (idF-SymmetricMonoidal; ∘-SymmetricMonoidal) + +open SMF.Lax using (SymmetricMonoidalFunctor) +open SMNI.Lax using (SymmetricMonoidalNaturalIsomorphism; id; isEquivalence) + +assoc + : {A B C D : SymmetricMonoidalCategory o ℓ e} + {F : SymmetricMonoidalFunctor A B} + {G : SymmetricMonoidalFunctor B C} + {H : SymmetricMonoidalFunctor C D} + → SymmetricMonoidalNaturalIsomorphism + (∘-SymmetricMonoidal (∘-SymmetricMonoidal H G) F) + (∘-SymmetricMonoidal H (∘-SymmetricMonoidal G F)) +assoc {A} {B} {C} {D} {F} {G} {H} = SMNI.Lax.associator {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} {C} {D} {F} {G} {H} + +identityˡ + : {A B : SymmetricMonoidalCategory o ℓ e} + {F : SymmetricMonoidalFunctor A B} + → SymmetricMonoidalNaturalIsomorphism (∘-SymmetricMonoidal (idF-SymmetricMonoidal B) F) F +identityˡ {A} {B} {F} = SMNI.Lax.unitorˡ {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} {F} + +identityʳ + : {A B : SymmetricMonoidalCategory o ℓ e} + {F : SymmetricMonoidalFunctor A B} + → SymmetricMonoidalNaturalIsomorphism (∘-SymmetricMonoidal F (idF-SymmetricMonoidal A)) F +identityʳ {A} {B} {F} = SMNI.Lax.unitorʳ {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} {F} + +Compose + : {A B C : SymmetricMonoidalCategory o ℓ e} + → SymmetricMonoidalFunctor B C + → SymmetricMonoidalFunctor A B + → SymmetricMonoidalFunctor A C +Compose {A} {B} {C} F G = ∘-SymmetricMonoidal {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} {C} F G + +∘-resp-≈ + : {A B C : SymmetricMonoidalCategory o ℓ e} + {f h : SymmetricMonoidalFunctor B C} + {g i : SymmetricMonoidalFunctor A B} + → SymmetricMonoidalNaturalIsomorphism f h + → SymmetricMonoidalNaturalIsomorphism g i + → SymmetricMonoidalNaturalIsomorphism (∘-SymmetricMonoidal f g) (∘-SymmetricMonoidal h i) +∘-resp-≈ {A} {B} {C} {F} {F′} {G} {G′} F≃F′ G≃G′ = SMNI.Lax._ⓘₕ_ {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} {C} {G} {G′} {F} {F′} F≃F′ G≃G′ + +SymMonCat : Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) +SymMonCat = categoryHelper record + { Obj = SymmetricMonoidalCategory o ℓ e + ; _⇒_ = SymmetricMonoidalFunctor {o} {o} {ℓ} {ℓ} {e} {e} + ; _≈_ = λ { {A} {B} → SymmetricMonoidalNaturalIsomorphism {o} {ℓ} {e} {o} {ℓ} {e} {A} {B} } + ; id = λ { {X} → idF-SymmetricMonoidal X } + ; _∘_ = λ { {A} {B} {C} F G → Compose {A} {B} {C} F G } + ; assoc = λ { {A} {B} {C} {D} {F} {G} {H} → assoc {A} {B} {C} {D} {F} {G} {H} } + ; identityˡ = λ { {A} {B} {F} → identityˡ {A} {B} {F} } + ; identityʳ = λ { {A} {B} {F} → identityʳ {A} {B} {F} } + ; equiv = isEquivalence + ; ∘-resp-≈ = λ { {A} {B} {C} {f} {h} {g} {i} f≈h g≈i → ∘-resp-≈ {A} {B} {C} {f} {h} {g} {i} f≈h g≈i } + } diff --git a/Category/Monoidal/Instance/Cospans.agda b/Category/Monoidal/Instance/Cospans.agda index 228ddea..a1648db 100644 --- a/Category/Monoidal/Instance/Cospans.agda +++ b/Category/Monoidal/Instance/Cospans.agda @@ -18,8 +18,9 @@ open import Categories.Category.Monoidal.Core using (Monoidal) open import Categories.Functor using (Functor) open import Categories.Functor.Properties using ([_]-resp-≅) open import Categories.NaturalTransformation.NaturalIsomorphism using (niHelper) -open import Category.Instance.Cospans 𝒞 using (Cospans; Cospan) -open import Category.Instance.Properties.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes-CC) +open import Category.Instance.Cospans 𝒞 using (Cospans) +open import Category.Diagram.Cospan using (Cospan; cospan) +open import Category.Cartesian.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes-CC) open import Category.Monoidal.Instance.Cospans.Lift {o} {ℓ} {e} using (module Square) open import Data.Product.Base using (_,_) open import Functor.Instance.Cospan.Stack 𝒞 using (⊗) @@ -61,7 +62,6 @@ CospansMonoidal = record ; pentagon = pentagon } where - module ⊗ = Functor ⊗ module Cospans = Category Cospans module Unitorˡ = Square ⊥+--id module Unitorʳ = Square -+⊥-id @@ -83,16 +83,15 @@ CospansMonoidal = record CospansBraided : Braided CospansMonoidal CospansBraided = record { braiding = niHelper record - { η = λ { (X , Y) → Braiding.FX≅GX′.from {X , Y} } - ; η⁻¹ = λ { (Y , X) → Braiding.FX≅GX′.to {Y , X} } - ; commute = λ { {X , Y} {X′ , Y′} (f , g) → Braiding.from (record { f₁ = f₁ f , f₁ g ; f₂ = f₂ f , f₂ g }) } - ; iso = λ { (X , Y) → Braiding.FX≅GX′.iso {X , Y} } + { η = λ (X , Y) → Braiding.FX≅GX′.from {X , Y} + ; η⁻¹ = λ (Y , X) → Braiding.FX≅GX′.to {Y , X} + ; commute = λ (cospan f₁ f₂ , cospan g₁ g₂) → Braiding.from (cospan (f₁ , g₁) (f₂ , g₂)) + ; iso = λ (X , Y) → Braiding.FX≅GX′.iso {X , Y} } ; hexagon₁ = sym L-resp-⊗ ⟩∘⟨ refl ⟩∘⟨ sym L-resp-⊗ ○ refl⟩∘⟨ sym homomorphism ○ sym homomorphism ○ L-resp-≈ hex₁ ○ homomorphism ○ refl⟩∘⟨ homomorphism ; hexagon₂ = sym L-resp-⊗ ⟩∘⟨refl ⟩∘⟨ sym L-resp-⊗ ○ sym homomorphism ⟩∘⟨refl ○ sym homomorphism ○ L-resp-≈ hex₂ ○ homomorphism ○ homomorphism ⟩∘⟨refl } where - open Cospan module Cospans = Category Cospans open Cospans.Equiv open Cospans.HomReasoning diff --git a/Category/Monoidal/Instance/Cospans/Lift.agda b/Category/Monoidal/Instance/Cospans/Lift.agda index fa31fcb..c7e7516 100644 --- a/Category/Monoidal/Instance/Cospans/Lift.agda +++ b/Category/Monoidal/Instance/Cospans/Lift.agda @@ -4,7 +4,7 @@ open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategor module Category.Monoidal.Instance.Cospans.Lift {o ℓ e} where -open import Category.Instance.Cospans using (Cospans; Cospan; Same) +open import Category.Instance.Cospans using (Cospans) open import Categories.Category.Core using (Category) @@ -16,6 +16,7 @@ import Category.Diagram.Pushout as PushoutDiagram′ import Functor.Instance.Cospan.Embed as CospanEmbed open import Categories.Category using (_[_,_]; _[_≈_]; _[_∘_]; module Definitions) +open import Category.Diagram.Cospan using (Cospan; cospan) open import Categories.Functor.Core using (Functor) open import Categories.Functor.Properties using ([_]-resp-≅) open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; _≃_) @@ -26,7 +27,7 @@ module _ {𝒞 : FinitelyCocompleteCategory o ℓ e} {𝒟 : FinitelyCocompleteC module 𝒞 = FinitelyCocompleteCategory 𝒞 module 𝒟 = FinitelyCocompleteCategory 𝒟 - open CospanEmbed 𝒟 using (L; B₁; B∘L; R∘B; ≅-L-R) + open CospanEmbed 𝒟 using (L; B∘L; R∘B; ≅-L-R) module Square {F G : Functor 𝒞.U 𝒟.U} (F≃G : F ≃ G) where @@ -51,21 +52,21 @@ module _ {𝒞 : FinitelyCocompleteCategory o ℓ e} {𝒟 : FinitelyCocompleteC open Cospan fg renaming (f₁ to f; f₂ to g) open 𝒟 using (_∘_) - squares⇒cospan : Cospans 𝒟 [ B₁ (G.₁ f ∘ FX≅GX.from) (G.₁ g ∘ FX≅GX.from) ≈ B₁ (F.₁ f) (F.₁ g) ] + squares⇒cospan : Cospans 𝒟 [ cospan (G.₁ f ∘ FX≅GX.from) (G.₁ g ∘ FX≅GX.from) ≈ cospan (F.₁ f) (F.₁ g) ] squares⇒cospan = record { ≅N = ≅.sym 𝒟.U FX≅GX - ; from∘f₁≈f₁′ = sym (switch-fromtoˡ FX≅GX (⇒.commute f)) - ; from∘f₂≈f₂′ = sym (switch-fromtoˡ FX≅GX (⇒.commute g)) + ; from∘f₁≈f₁ = sym (switch-fromtoˡ FX≅GX (⇒.commute f)) + ; from∘f₂≈f₂ = sym (switch-fromtoˡ FX≅GX (⇒.commute g)) } where open 𝒟.Equiv using (sym) - from : Cospans 𝒟 [ Cospans 𝒟 [ L.₁ (⇒.η Y) ∘ B₁ (F.₁ f) (F.₁ g) ] ≈ Cospans 𝒟 [ B₁ (G.₁ f) (G.₁ g) ∘ L.₁ (⇒.η X) ] ] + from : Cospans 𝒟 [ Cospans 𝒟 [ L.₁ (⇒.η Y) ∘ cospan (F.₁ f) (F.₁ g) ] ≈ Cospans 𝒟 [ cospan (G.₁ f) (G.₁ g) ∘ L.₁ (⇒.η X) ] ] from = sym (switch-tofromˡ FX≅GX′ (refl⟩∘⟨ B∘L ○ ≅-L-R FX≅GX ⟩∘⟨refl ○ R∘B ○ squares⇒cospan)) where open Cospans.Equiv using (sym) - to : Cospans 𝒟 [ Cospans 𝒟 [ L.₁ (⇐.η Y) ∘ B₁ (G.₁ f) (G.₁ g) ] ≈ Cospans 𝒟 [ B₁ (F.₁ f) (F.₁ g) ∘ L.₁ (⇐.η X) ] ] + to : Cospans 𝒟 [ Cospans 𝒟 [ L.₁ (⇐.η Y) ∘ cospan (G.₁ f) (G.₁ g) ] ≈ Cospans 𝒟 [ cospan (F.₁ f) (F.₁ g) ∘ L.₁ (⇐.η X) ] ] to = switch-fromtoʳ FX≅GX′ (pullʳ B∘L ○ ≅-L-R FX≅GX ⟩∘⟨refl ○ R∘B ○ squares⇒cospan) where open ⇒-Reasoning (Cospans 𝒟) using (pullʳ) diff --git a/Category/Monoidal/Instance/DecoratedCospans.agda b/Category/Monoidal/Instance/DecoratedCospans.agda new file mode 100644 index 0000000..3df57ee --- /dev/null +++ b/Category/Monoidal/Instance/DecoratedCospans.agda @@ -0,0 +1,410 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --lossy-unification #-} + +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module Category.Monoidal.Instance.DecoratedCospans + {o o′ ℓ ℓ′ e e′} + (𝒞 : FinitelyCocompleteCategory o ℓ e) + {𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′} + (F : SymmetricMonoidalFunctor (smc 𝒞) 𝒟) where + +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Morphism as Morphism +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Category.Monoidal.Properties as ⊗-Properties +import Categories.Category.Monoidal.Braided.Properties as σ-Properties + +open import Categories.Category using (_[_,_]; _[_≈_]; _[_∘_]; Category) +open import Categories.Category.BinaryProducts using (BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cartesian.Bundle using (CartesianCategory) +open import Categories.Category.Cocartesian using (module CocartesianMonoidal; module CocartesianSymmetricMonoidal) +open import Categories.Category.Monoidal.Braided using (Braided) +open import Categories.Category.Monoidal.Core using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) +open import Categories.Category.Monoidal.Utilities using (module Shorthands) +open import Categories.Functor using (Functor; _∘F_) renaming (id to idF) +open import Categories.Functor.Hom using (Hom[_][_,-]) +open import Categories.Functor.Properties using ([_]-resp-≅) +open import Categories.Functor.Monoidal.Construction.Product using (⁂-SymmetricMonoidalFunctor) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; _ⓘˡ_; niHelper) +open import Categories.NaturalTransformation.Core using (NaturalTransformation; _∘ᵥ_; ntHelper) +open import Categories.NaturalTransformation.Equivalence using () renaming (_≃_ to _≋_) +open import Category.Instance.DecoratedCospans 𝒞 F using (DecoratedCospans) +open import Category.Cartesian.Instance.FinitelyCocompletes {o} {ℓ} {e} using (module x+y; module y+x; [x+y]+z; x+[y+z]; assoc-≃) +open import Category.Monoidal.Instance.DecoratedCospans.Lift {o} {ℓ} {e} {o′} {ℓ′} {e′} using (module Square) +open import Cospan.Decorated 𝒞 F using (DecoratedCospan) +open import Data.Product.Base using (_,_) +open import Function.Base using () renaming (id to idf) +open import Functor.Instance.DecoratedCospan.Stack 𝒞 F using (⊗) +open import Functor.Instance.DecoratedCospan.Embed 𝒞 F using (L; L-resp-⊗; B₁) + +open import Category.Monoidal.Instance.DecoratedCospans.Products 𝒞 F +open CocartesianMonoidal 𝒞.U 𝒞.cocartesian using (⊥+--id; -+⊥-id; ⊥+A≅A; A+⊥≅A; +-monoidal) +open import Categories.Category.Monoidal.Utilities +-monoidal using (associator-naturalIsomorphism) + +module LiftUnitorˡ where + module F = SymmetricMonoidalFunctor F + open 𝒞 using (⊥; _+-; i₂; _+_; _+₁_; ¡; +₁-cong₂; ¡-unique; -+-) + open Shorthands 𝒟.monoidal using (ρ⇒; ρ⇐; λ⇒) + ≃₁ : NaturalTransformation (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F (⊥ +-)) + ≃₁ = ntHelper record + { η = λ { X → record + { to = λ { f → 𝒟.U [ F.⊗-homo.η (⊥ , X) ∘ 𝒟.U [ 𝒟.⊗.₁ (𝒟.U [ F.₁ 𝒞.initial.! ∘ F.ε ] , f) ∘ ρ⇐ ] ] } + ; cong = λ { x → refl⟩∘⟨ refl⟩⊗⟨ x ⟩∘⟨refl } } + } + ; commute = ned + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + ned : {X Y : 𝒞.Obj} (f : X 𝒞.⇒ Y) {x : 𝒟.unit 𝒟.⇒ F.₀ X} + → F.⊗-homo.η (⊥ , Y) ∘ (F.₁ ¡ 𝒟.∘ F.ε) ⊗₁ (F.F₁ f ∘ x ∘ id) ∘ ρ⇐ + 𝒟.≈ F.₁ (𝒞.id +₁ f) ∘ (F.⊗-homo.η (𝒞.⊥ , X) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ x ∘ ρ⇐) ∘ id + ned {X} {Y} f {x} = begin + F.⊗-homo.η (⊥ , Y) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ (F.₁ f ∘ x ∘ id) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ identityʳ) ⟩∘⟨refl ⟩ + F.⊗-homo.η (⊥ , Y) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ (F.₁ f ∘ x) ∘ ρ⇐ ≈⟨ push-center split₂ˡ ⟩ + F.⊗-homo.η (⊥ , Y) ∘ id ⊗₁ F.₁ f ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ x ∘ ρ⇐ ≈⟨ refl⟩∘⟨ F.identity ⟩⊗⟨refl ⟩∘⟨refl ⟨ + F.⊗-homo.η (⊥ , Y) ∘ F.₁ 𝒞.id ⊗₁ F.₁ f ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ x ∘ ρ⇐ ≈⟨ extendʳ (F.⊗-homo.commute (𝒞.id , f)) ⟩ + F.₁ (𝒞.id +₁ f) ∘ F.⊗-homo.η (⊥ , X) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ x ∘ ρ⇐ ≈⟨ refl⟩∘⟨ identityʳ ⟨ + F.₁ (𝒞.id +₁ f) ∘ (F.⊗-homo.η (⊥ , X) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ x ∘ ρ⇐) ∘ id ∎ + ≃₂ : NaturalTransformation (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F idF) + ≃₂ = ntHelper record + { η = λ { X → record { to = idf ; cong = idf } } + ; commute = λ { f → refl } + } + where + open 𝒟.Equiv + open NaturalIsomorphism using (F⇐G) + ≃₂≋≃₁ : (F⇐G (Hom[ 𝒟.U ][ 𝒟.unit ,-] ⓘˡ (F.F ⓘˡ ⊥+--id))) ∘ᵥ ≃₂ ≋ ≃₁ + ≃₂≋≃₁ {X} {f} = begin + F.₁ λ⇐ ∘ f ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ 𝒟.unitorʳ.isoʳ ⟨ + F.₁ λ⇐ ∘ f ∘ ρ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ coherence₃ ⟩∘⟨refl ⟨ + F.₁ λ⇐ ∘ f ∘ λ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ 𝒟.unitorˡ-commute-from ⟨ + F.₁ λ⇐ ∘ λ⇒ ∘ id ⊗₁ f ∘ ρ⇐ ≈⟨ pushˡ (introˡ F.identity) ⟩ + F.₁ 𝒞.id ∘ F.₁ λ⇐ ∘ λ⇒ ∘ id ⊗₁ f ∘ ρ⇐ ≈⟨ F.F-resp-≈ (-+-.identity) ⟩∘⟨refl ⟨ + F.₁ (𝒞.id +₁ 𝒞.id) ∘ F.₁ λ⇐ ∘ λ⇒ ∘ id ⊗₁ f ∘ ρ⇐ ≈⟨ F.F-resp-≈ (+₁-cong₂ (¡-unique 𝒞.id) 𝒞.Equiv.refl) ⟩∘⟨refl ⟨ + F.₁ (¡ +₁ 𝒞.id) ∘ F.₁ λ⇐ ∘ λ⇒ ∘ id ⊗₁ f ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ unitorˡ) F.unitaryˡ) ⟨ + F.₁ (¡ +₁ 𝒞.id) ∘ F.⊗-homo.η (⊥ , X) ∘ F.ε ⊗₁ id ∘ id ⊗₁ f ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ (sym serialize₁₂) ⟩ + F.₁ (¡ +₁ 𝒞.id) ∘ F.⊗-homo.η (⊥ , X) ∘ F.ε ⊗₁ f ∘ ρ⇐ ≈⟨ extendʳ (F.⊗-homo.commute (¡ , 𝒞.id)) ⟨ + F.⊗-homo.η (⊥ , X) ∘ (F.₁ ¡ ⊗₁ F.₁ 𝒞.id) ∘ F.ε ⊗₁ f ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ F.identity ⟩∘⟨refl ⟩ + F.⊗-homo.η (⊥ , X) ∘ (F.₁ ¡ ⊗₁ id) ∘ F.ε ⊗₁ f ∘ ρ⇐ ≈⟨ pull-center (sym split₁ˡ) ⟩ + F.⊗-homo.η (⊥ , X) ∘ (F.₁ ¡ ∘ F.ε) ⊗₁ f ∘ ρ⇐ ∎ + where + open Shorthands 𝒞.monoidal using (λ⇐) + open CocartesianMonoidal 𝒞.U 𝒞.cocartesian using (unitorˡ) + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ) + open ⊗-Reasoning 𝒟.monoidal + open ⊗-Properties 𝒟.monoidal using (coherence₃) + open ⇒-Reasoning 𝒟.U + module -+- = Functor -+- + module Unitorˡ = Square {𝒞} {𝒞} {𝒟} {𝒟} {F} {F} ⊥+--id ≃₁ ≃₂ ≃₂≋≃₁ +open LiftUnitorˡ using (module Unitorˡ) + +module LiftUnitorʳ where + module F = SymmetricMonoidalFunctor F + open 𝒞 using (⊥; -+_; i₁; _+_; _+₁_; ¡; +₁-cong₂; ¡-unique; -+-) + open Shorthands 𝒟.monoidal using (ρ⇒; ρ⇐) + ≃₁ : NaturalTransformation (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F (-+ ⊥)) + ≃₁ = ntHelper record + { η = λ { X → record + { to = λ { f → 𝒟.U [ F.⊗-homo.η (X , ⊥) ∘ 𝒟.U [ 𝒟.⊗.₁ (f , 𝒟.U [ F.₁ 𝒞.initial.! ∘ F.ε ]) ∘ ρ⇐ ] ] } + ; cong = λ { x → refl⟩∘⟨ x ⟩⊗⟨refl ⟩∘⟨refl } } + } + ; commute = ned + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + ned : {X Y : 𝒞.Obj} (f : X 𝒞.⇒ Y) {x : 𝒟.unit 𝒟.⇒ F.₀ X} + → F.⊗-homo.η (Y , ⊥) ∘ (F.F₁ f ∘ x ∘ id) ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ + 𝒟.≈ F.₁ (f +₁ 𝒞.id) ∘ (F.⊗-homo.η (X , ⊥) ∘ x ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐) ∘ id + ned {X} {Y} f {x} = begin + F.⊗-homo.η (Y , ⊥) ∘ (F.₁ f ∘ x ∘ id) ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F.⊗-homo.η (Y , ⊥) ∘ (F.₁ f ∘ x) ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ≈⟨ push-center split₁ˡ ⟩ + F.⊗-homo.η (Y , ⊥) ∘ F.₁ f ⊗₁ id ∘ x ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ F.identity ⟩∘⟨refl ⟨ + F.⊗-homo.η (Y , ⊥) ∘ F.₁ f ⊗₁ F.₁ 𝒞.id ∘ x ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ≈⟨ extendʳ (F.⊗-homo.commute (f , 𝒞.id)) ⟩ + F.₁ (f +₁ 𝒞.id) ∘ F.⊗-homo.η (X , ⊥) ∘ x ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ identityʳ ⟨ + F.₁ (f +₁ 𝒞.id) ∘ (F.⊗-homo.η (X , ⊥) ∘ x ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐) ∘ id ∎ + ≃₂ : NaturalTransformation (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F idF) + ≃₂ = ntHelper record + { η = λ { X → record { to = idf ; cong = idf } } + ; commute = λ { f → refl } + } + where + open 𝒟.Equiv + open NaturalIsomorphism using (F⇐G) + ≃₂≋≃₁ : (F⇐G (Hom[ 𝒟.U ][ 𝒟.unit ,-] ⓘˡ (F.F ⓘˡ -+⊥-id))) ∘ᵥ ≃₂ ≋ ≃₁ + ≃₂≋≃₁ {X} {f} = begin + F.₁ i₁ ∘ f ∘ id ≈⟨ refl⟩∘⟨ refl⟩∘⟨ 𝒟.unitorʳ.isoʳ ⟨ + F.₁ ρ⇐′ ∘ f ∘ ρ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ 𝒟.unitorʳ-commute-from ⟨ + F.₁ ρ⇐′ ∘ ρ⇒ ∘ f ⊗₁ id ∘ ρ⇐ ≈⟨ pushˡ (introˡ F.identity) ⟩ + F.₁ 𝒞.id ∘ F.₁ ρ⇐′ ∘ ρ⇒ ∘ f ⊗₁ id ∘ ρ⇐ ≈⟨ F.F-resp-≈ (-+-.identity) ⟩∘⟨refl ⟨ + F.₁ (𝒞.id +₁ 𝒞.id) ∘ F.₁ ρ⇐′ ∘ ρ⇒ ∘ f ⊗₁ id ∘ ρ⇐ ≈⟨ F.F-resp-≈ (+₁-cong₂ 𝒞.Equiv.refl (¡-unique 𝒞.id)) ⟩∘⟨refl ⟨ + F.₁ (𝒞.id +₁ ¡) ∘ F.₁ ρ⇐′ ∘ ρ⇒ ∘ f ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ unitorʳ) F.unitaryʳ) ⟨ + F.₁ (𝒞.id +₁ ¡) ∘ F.⊗-homo.η (X , ⊥) ∘ id ⊗₁ F.ε ∘ f ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ (sym serialize₂₁) ⟩ + F.₁ (𝒞.id +₁ ¡) ∘ F.⊗-homo.η (X , ⊥) ∘ f ⊗₁ F.ε ∘ ρ⇐ ≈⟨ extendʳ (F.⊗-homo.commute (𝒞.id , ¡)) ⟨ + F.⊗-homo.η (X , ⊥) ∘ (F.₁ 𝒞.id ⊗₁ F.₁ ¡) ∘ f ⊗₁ F.ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ F.identity ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F.⊗-homo.η (X , ⊥) ∘ (id ⊗₁ F.₁ ¡) ∘ f ⊗₁ F.ε ∘ ρ⇐ ≈⟨ pull-center (sym split₂ˡ) ⟩ + F.⊗-homo.η (X , ⊥) ∘ f ⊗₁ (F.₁ ¡ ∘ F.ε) ∘ ρ⇐ ∎ + where + open Shorthands 𝒞.monoidal using () renaming (ρ⇐ to ρ⇐′) + open CocartesianMonoidal 𝒞.U 𝒞.cocartesian using (unitorʳ) + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + module -+- = Functor -+- + module Unitorʳ = Square {𝒞} {𝒞} {𝒟} {𝒟} {F} {F} -+⊥-id ≃₁ ≃₂ ≃₂≋≃₁ +open LiftUnitorʳ using (module Unitorʳ) + +module LiftAssociator where + module F = SymmetricMonoidalFunctor F + open 𝒞 using (⊥; -+_; i₁; _+_; _+₁_; ¡; +₁-cong₂; ¡-unique; -+-) + open Shorthands 𝒟.monoidal using (ρ⇒; ρ⇐) + ≃₁ : NaturalTransformation (Hom[ [𝒟×𝒟]×𝒟.U ][ [𝒟×𝒟]×𝒟.unit ,-] ∘F [F×F]×F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F ([x+y]+z.F {𝒞})) + ≃₁ = ntHelper record + { η = λ { ((X , Y) , Z) → record + { to = λ { ((f , g) , h) → F.⊗-homo.η (X + Y , Z) ∘ ((F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ h) ∘ ρ⇐ } + ; cong = λ { {(f , g) , h} {(f′ , g′) , h′} ((x , y) , z) → refl⟩∘⟨ (refl⟩∘⟨ x ⟩⊗⟨ y ⟩∘⟨refl) ⟩⊗⟨ z ⟩∘⟨refl } + } } + ; commute = λ { {(X , Y) , Z} {(X′ , Y′) , Z′} ((x , y) , z) {(f , g) , h} → commute x y z f g h } + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ; _≈_) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + commute + : {X Y Z X′ Y′ Z′ : 𝒞.Obj} + (x : 𝒞.U [ X , X′ ]) + (y : 𝒞.U [ Y , Y′ ]) + (z : 𝒞.U [ Z , Z′ ]) + (f : 𝒟.U [ 𝒟.unit , F.₀ X ]) + (g : 𝒟.U [ 𝒟.unit , F.₀ Y ]) + (h : 𝒟.U [ 𝒟.unit , F.₀ Z ]) + → F.⊗-homo.η (X′ + Y′ , Z′) ∘ (F.⊗-homo.η (X′ , Y′) ∘ (F.₁ x ∘ f ∘ id) ⊗₁ (F.₁ y ∘ g ∘ id) ∘ ρ⇐) ⊗₁ (F.₁ z ∘ h ∘ id) ∘ ρ⇐ + ≈ F.₁ ((x +₁ y) +₁ z) ∘ (F.⊗-homo.η (X + Y , Z) ∘ (F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐) ∘ id + commute {X} {Y} {Z} {X′} {Y′} {Z′} x y z f g h = begin + F.⊗-homo.η (X′ + Y′ , Z′) ∘ (F.⊗-homo.η (X′ , Y′) ∘ (F.₁ x ∘ f ∘ id) ⊗₁ (F.₁ y ∘ g ∘ id) ∘ ρ⇐) ⊗₁ (F.₁ z ∘ h ∘ id) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨ (refl⟩∘⟨ identityʳ) ⟩∘⟨refl) ⟩⊗⟨ (refl⟩∘⟨ identityʳ) ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ + Y′ , Z′) ∘ (F.⊗-homo.η (X′ , Y′) ∘ (F.₁ x ∘ f) ⊗₁ (F.₁ y ∘ g) ∘ ρ⇐) ⊗₁ (F.₁ z ∘ h) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ push-center ⊗-distrib-over-∘ ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ + Y′ , Z′) ∘ (F.⊗-homo.η (X′ , Y′) ∘ F.₁ x ⊗₁ F.₁ y ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ (F.₁ z ∘ h) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ extendʳ (F.⊗-homo.commute (x , y)) ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ + Y′ , Z′) ∘ (F.₁ (x +₁ y) ∘ F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ (F.₁ z ∘ h) ∘ ρ⇐ + ≈⟨ push-center ⊗-distrib-over-∘ ⟩ + F.⊗-homo.η (X′ + Y′ , Z′) ∘ F.₁ (x +₁ y) ⊗₁ F.₁ z ∘ (F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ + ≈⟨ extendʳ (F.⊗-homo.commute (x +₁ y , z)) ⟩ + F.₁ ((x +₁ y) +₁ z) ∘ F.⊗-homo.η (X + Y , Z) ∘ (F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ identityʳ ⟨ + F.₁ ((x +₁ y) +₁ z) ∘ (F.⊗-homo.η (X + Y , Z) ∘ (F.⊗-homo.η (X , Y) ∘ f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐) ∘ id + ∎ + ≃₂ : NaturalTransformation (Hom[ [𝒟×𝒟]×𝒟.U ][ [𝒟×𝒟]×𝒟.unit ,-] ∘F [F×F]×F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F (x+[y+z].F {𝒞})) + ≃₂ = ntHelper record + { η = λ { ((X , Y) , Z) → record + { to = λ { ((f , g) , h) → F.⊗-homo.η (X , Y + Z) ∘ (f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐)) ∘ ρ⇐ } + ; cong = λ { {(f , g) , h} {(f′ , g′) , h′} ((x , y) , z) → refl⟩∘⟨ x ⟩⊗⟨ (refl⟩∘⟨ y ⟩⊗⟨ z ⟩∘⟨refl) ⟩∘⟨refl } + } } + ; commute = λ { {(X , Y) , Z} {(X′ , Y′) , Z′} ((x , y) , z) {(f , g) , h} → commute x y z f g h } + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ; _≈_) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + commute + : {X Y Z X′ Y′ Z′ : 𝒞.Obj} + (x : 𝒞.U [ X , X′ ]) + (y : 𝒞.U [ Y , Y′ ]) + (z : 𝒞.U [ Z , Z′ ]) + (f : 𝒟.U [ 𝒟.unit , F.₀ X ]) + (g : 𝒟.U [ 𝒟.unit , F.₀ Y ]) + (h : 𝒟.U [ 𝒟.unit , F.₀ Z ]) + → F.⊗-homo.η (X′ , Y′ + Z′) ∘ (F.₁ x ∘ f ∘ id) ⊗₁ (F.⊗-homo.η (Y′ , Z′) ∘ (F.₁ y ∘ g ∘ id) ⊗₁ (F.₁ z ∘ h ∘ id) ∘ ρ⇐) ∘ ρ⇐ + ≈ F.₁ (x +₁ (y +₁ z)) ∘ (F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐) ∘ id + commute {X} {Y} {Z} {X′} {Y′} {Z′} x y z f g h = begin + F.⊗-homo.η (X′ , Y′ + Z′) ∘ (F.₁ x ∘ f ∘ id) ⊗₁ (F.⊗-homo.η (Y′ , Z′) ∘ (F.₁ y ∘ g ∘ id) ⊗₁ (F.₁ z ∘ h ∘ id) ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨ (refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨ (refl⟩∘⟨ identityʳ) ⟩∘⟨refl) ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ , Y′ + Z′) ∘ (F.₁ x ∘ f) ⊗₁ (F.⊗-homo.η (Y′ , Z′) ∘ (F.₁ y ∘ g) ⊗₁ (F.₁ z ∘ h) ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ push-center ⊗-distrib-over-∘ ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ , Y′ + Z′) ∘ (F.₁ x ∘ f) ⊗₁ (F.⊗-homo.η (Y′ , Z′) ∘ F.₁ y ⊗₁ F.₁ z ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ extendʳ (F.⊗-homo.commute (y , z)) ⟩∘⟨refl ⟩ + F.⊗-homo.η (X′ , Y′ + Z′) ∘ (F.₁ x ∘ f) ⊗₁ (F.₁ (y +₁ z) ∘ F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ push-center ⊗-distrib-over-∘ ⟩ + F.⊗-homo.η (X′ , Y′ + Z′) ∘ F.₁ x ⊗₁ F.₁ (y +₁ z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ extendʳ (F.⊗-homo.commute (x , y +₁ z)) ⟩ + F.₁ (x +₁ (y +₁ z)) ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐ + ≈⟨ refl⟩∘⟨ identityʳ ⟨ + F.₁ (x +₁ (y +₁ z)) ∘ (F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐) ∘ id + ∎ + open NaturalIsomorphism using (F⇐G) + ≃₂≋≃₁ : (F⇐G (Hom[ 𝒟.U ][ 𝒟.unit ,-] ⓘˡ (F.F ⓘˡ assoc-≃))) ∘ᵥ ≃₂ ≋ ≃₁ + ≃₂≋≃₁ {(X , Y) , Z} {(f , g) , h} = begin + F.₁ α⇐′ ∘ (F.⊗-homo.η (X , Y + Z) ∘ (f ⊗₁ _) ∘ ρ⇐) ∘ id ≈⟨ refl⟩∘⟨ identityʳ ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h ∘ ρ⇐) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ sym-assoc ⟩∘⟨refl ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ ((F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h) ∘ ρ⇐) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₂ʳ ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h) ∘ id ⊗₁ ρ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ coherence-inv₃ ⟨ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h) ∘ id ⊗₁ ρ⇐ ∘ λ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ unitorˡ-commute-to ⟨ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h) ∘ λ⇐ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ (switch-tofromˡ α coherence-inv₁) ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ f ⊗₁ (F.⊗-homo.η (Y , Z) ∘ g ⊗₁ h) ∘ α⇒ ∘ λ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₂ˡ ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ id ⊗₁ F.⊗-homo.η (Y , Z) ∘ f ⊗₁ (g ⊗₁ h) ∘ α⇒ ∘ λ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ coherence-inv₃ ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ id ⊗₁ F.⊗-homo.η (Y , Z) ∘ f ⊗₁ (g ⊗₁ h) ∘ α⇒ ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ extendʳ assoc-commute-from ⟨ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ id ⊗₁ F.⊗-homo.η (Y , Z) ∘ α⇒ ∘ (f ⊗₁ g) ⊗₁ h ∘ ρ⇐ ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ (sym split₁ʳ) ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ id ⊗₁ F.⊗-homo.η (Y , Z) ∘ α⇒ ∘ (f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩ + F.₁ α⇐′ ∘ F.⊗-homo.η (X , Y + Z) ∘ (id ⊗₁ F.⊗-homo.η (Y , Z) ∘ α⇒) ∘ (f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ ≈⟨ refl⟩∘⟨ sym-assoc ⟩ + F.₁ α⇐′ ∘ (F.⊗-homo.η (X , Y + Z) ∘ id ⊗₁ F.⊗-homo.η (Y , Z) ∘ α⇒) ∘ (f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ ≈⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ α′) F.associativity) ⟨ + F.⊗-homo.η (X + Y , Z) ∘ F.⊗-homo.η (X , Y) ⊗₁ id ∘ (f ⊗₁ g ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ (sym split₁ˡ) ⟩ + F.⊗-homo.η (X + Y , Z) ∘ (F.⊗-homo.η (X , Y) ∘ (f ⊗₁ g) ∘ ρ⇐) ⊗₁ h ∘ ρ⇐ ∎ + where + open Shorthands 𝒞.monoidal using () renaming (α⇐ to α⇐′) + open Shorthands 𝒟.monoidal using (α⇒; λ⇐) + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ; assoc-commute-from; unitorˡ-commute-to) renaming (unitorˡ to ƛ; associator to α) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + open CocartesianMonoidal 𝒞.U 𝒞.cocartesian using () renaming (associator to α′) + open ⊗-Properties 𝒟.monoidal using (coherence-inv₁; coherence-inv₃) + module Associator = Square {[𝒞×𝒞]×𝒞} {𝒞} {[𝒟×𝒟]×𝒟} {𝒟} {[F×F]×F} {F} {[x+y]+z.F {𝒞}} {x+[y+z].F {𝒞}} (assoc-≃ {𝒞}) ≃₁ ≃₂ ≃₂≋≃₁ +open LiftAssociator using (module Associator) + +module LiftBraiding where + module F = SymmetricMonoidalFunctor F + open 𝒞 using (⊥; -+_; i₁; _+_; _+₁_; ¡; +₁-cong₂; ¡-unique; -+-) + open Shorthands 𝒟.monoidal using (ρ⇒; ρ⇐) + ≃₁ : NaturalTransformation (Hom[ 𝒟×𝒟.U ][ 𝒟×𝒟.unit ,-] ∘F F×F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F (x+y.F {𝒞})) + ≃₁ = ntHelper record + { η = λ { (X , Y) → record + { to = λ { (x , y) → F.⊗-homo.η (X , Y) ∘ x ⊗₁ y ∘ ρ⇐} + ; cong = λ { {x , y} {x′ , y′} (≈x , ≈y) → refl⟩∘⟨ ≈x ⟩⊗⟨ ≈y ⟩∘⟨refl } + } } + ; commute = λ { {X , Y} {X′ , Y′} (x , y) {f , g} → + (extendʳ + (refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨ (refl⟩∘⟨ identityʳ) + ○ refl⟩∘⟨ ⊗-distrib-over-∘ + ○ extendʳ (F.⊗-homo.commute (x , y)))) + ○ refl⟩∘⟨ extendˡ (sym identityʳ) } + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ; _≈_; assoc) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + ≃₂ : NaturalTransformation (Hom[ 𝒟×𝒟.U ][ 𝒟×𝒟.unit ,-] ∘F F×F.F) (Hom[ 𝒟.U ][ 𝒟.unit ,-] ∘F F.F ∘F (y+x.F {𝒞})) + ≃₂ = ntHelper record + { η = λ { (X , Y) → record + { to = λ { (x , y) → F.⊗-homo.η (Y , X) ∘ y ⊗₁ x ∘ ρ⇐} + ; cong = λ { {x , y} {x′ , y′} (≈x , ≈y) → refl⟩∘⟨ ≈y ⟩⊗⟨ ≈x ⟩∘⟨refl } + } } + ; commute = λ { {X , Y} {X′ , Y′} (x , y) {f , g} → + (extendʳ + (refl⟩∘⟨ (refl⟩∘⟨ identityʳ) ⟩⊗⟨ (refl⟩∘⟨ identityʳ) + ○ refl⟩∘⟨ ⊗-distrib-over-∘ + ○ extendʳ (F.⊗-homo.commute (y , x)))) + ○ refl⟩∘⟨ extendˡ (sym identityʳ) } + } + where + open 𝒟.Equiv + open 𝒟 using (sym-assoc; _∘_; id; _⊗₁_; identityʳ; _≈_; assoc) + open ⊗-Reasoning 𝒟.monoidal + open ⇒-Reasoning 𝒟.U + open NaturalIsomorphism using (F⇐G) + open Symmetric 𝒞.symmetric using (braiding) + ≃₂≋≃₁ : (F⇐G (Hom[ 𝒟.U ][ 𝒟.unit ,-] ⓘˡ F.F ⓘˡ braiding)) ∘ᵥ ≃₂ ≋ ≃₁ + ≃₂≋≃₁ {X , Y} {f , g} = + refl⟩∘⟨ (identityʳ ○ sym-assoc) + ○ extendʳ + (extendʳ F.braiding-compat + ○ refl⟩∘⟨ (𝒟.braiding.⇒.commute (g , f))) + ○ refl⟩∘⟨ pullʳ (sym (switch-tofromˡ 𝒟.braiding.FX≅GX braiding-coherence-inv) ○ coherence-inv₃) + where + open σ-Properties 𝒟.braided using (braiding-coherence-inv) + open 𝒟.Equiv + open 𝒟 using (sym-assoc; identityʳ) + open ⊗-Reasoning 𝒟.monoidal + open ⊗-Properties 𝒟.monoidal using (coherence-inv₃) + open ⇒-Reasoning 𝒟.U + module Braiding = Square {𝒞×𝒞} {𝒞} {𝒟×𝒟} {𝒟} {F×F} {F} {x+y.F {𝒞}} {y+x.F {𝒞}} braiding ≃₁ ≃₂ ≃₂≋≃₁ +open LiftBraiding using (module Braiding) + +CospansMonoidal : Monoidal DecoratedCospans +CospansMonoidal = record + { ⊗ = ⊗ + ; unit = ⊥ + ; unitorˡ = [ L ]-resp-≅ ⊥+A≅A + ; unitorʳ = [ L ]-resp-≅ A+⊥≅A + ; associator = [ L ]-resp-≅ (≅.sym +-assoc) + ; unitorˡ-commute-from = λ { {X} {Y} {f} → Unitorˡ.from f } + ; unitorˡ-commute-to = λ { {X} {Y} {f} → Unitorˡ.to f } + ; unitorʳ-commute-from = λ { {X} {Y} {f} → Unitorʳ.from f } + ; unitorʳ-commute-to = λ { {X} {Y} {f} → Unitorʳ.to f } + ; assoc-commute-from = λ { {X} {X′} {f} {Y} {Y′} {g} {Z} {Z′} {h} → Associator.from _ } + ; assoc-commute-to = λ { {X} {X′} {f} {Y} {Y′} {g} {Z} {Z′} {h} → Associator.to _ } + ; triangle = triangle + ; pentagon = pentagon + } + where + open Category DecoratedCospans using (id; module Equiv; module HomReasoning) + open Equiv + open HomReasoning + open 𝒞 using (⊥; Obj; U; +-assoc) + λ⇒ = Unitorˡ.FX≅GX′.from + ρ⇒ = Unitorʳ.FX≅GX′.from + α⇒ = Associator.FX≅GX′.from + open Morphism U using (module ≅) + open Monoidal +-monoidal using () renaming (triangle to tri; pentagon to pent) + triangle : {X Y : Obj} → DecoratedCospans [ DecoratedCospans [ ⊗.₁ (id {X} , λ⇒) ∘ α⇒ ] ≈ ⊗.₁ (ρ⇒ , id {Y}) ] + triangle = sym L-resp-⊗ ⟩∘⟨refl ○ sym L.homomorphism ○ L.F-resp-≈ tri ○ L-resp-⊗ + pentagon + : {W X Y Z : Obj} + → DecoratedCospans + [ DecoratedCospans [ ⊗.₁ (id {W} , α⇒ {(X , Y) , Z}) ∘ DecoratedCospans [ α⇒ ∘ ⊗.₁ (α⇒ , id) ] ] + ≈ DecoratedCospans [ α⇒ ∘ α⇒ ] ] + pentagon = sym L-resp-⊗ ⟩∘⟨ refl ⟩∘⟨ sym L-resp-⊗ ○ refl⟩∘⟨ sym L.homomorphism ○ sym L.homomorphism ○ L.F-resp-≈ pent ○ L.homomorphism + +CospansBraided : Braided CospansMonoidal +CospansBraided = record + { braiding = niHelper record + { η = λ { (X , Y) → Braiding.FX≅GX′.from {X , Y} } + ; η⁻¹ = λ { (Y , X) → Braiding.FX≅GX′.to {Y , X} } + ; commute = λ { {X , Y} {X′ , Y′} (f , g) → Braiding.from (record { cospan = record { f₁ = f₁ f , f₁ g ; f₂ = f₂ f , f₂ g } ; decoration = decoration f , decoration g}) } + ; iso = λ { (X , Y) → Braiding.FX≅GX′.iso {X , Y} } + } + ; hexagon₁ = sym L-resp-⊗ ⟩∘⟨ refl ⟩∘⟨ sym L-resp-⊗ ○ refl⟩∘⟨ sym L.homomorphism ○ sym L.homomorphism ○ L-resp-≈ hex₁ ○ L.homomorphism ○ refl⟩∘⟨ L.homomorphism + ; hexagon₂ = sym L-resp-⊗ ⟩∘⟨refl ⟩∘⟨ sym L-resp-⊗ ○ sym L.homomorphism ⟩∘⟨refl ○ sym L.homomorphism ○ L-resp-≈ hex₂ ○ L.homomorphism ○ L.homomorphism ⟩∘⟨refl + } + where + open Symmetric 𝒞.symmetric renaming (hexagon₁ to hex₁; hexagon₂ to hex₂) + open DecoratedCospan + module Cospans = Category DecoratedCospans + open Cospans.Equiv + open Cospans.HomReasoning + open Functor L renaming (F-resp-≈ to L-resp-≈) + +CospansSymmetric : Symmetric CospansMonoidal +CospansSymmetric = record + { braided = CospansBraided + ; commutative = sym homomorphism ○ L-resp-≈ comm ○ identity + } + where + open Symmetric 𝒞.symmetric renaming (commutative to comm) + module Cospans = Category DecoratedCospans + open Cospans.Equiv + open Cospans.HomReasoning + open Functor L renaming (F-resp-≈ to L-resp-≈) diff --git a/Category/Monoidal/Instance/DecoratedCospans/Lift.agda b/Category/Monoidal/Instance/DecoratedCospans/Lift.agda new file mode 100644 index 0000000..c75f0db --- /dev/null +++ b/Category/Monoidal/Instance/DecoratedCospans/Lift.agda @@ -0,0 +1,114 @@ +{-# OPTIONS --without-K --safe #-} + +module Category.Monoidal.Instance.DecoratedCospans.Lift {o ℓ e o′ ℓ′ e′} where + +import Categories.Diagram.Pushout as PushoutDiagram +import Categories.Diagram.Pushout.Properties as PushoutProperties +import Categories.Morphism as Morphism +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Category.Diagram.Pushout as PushoutDiagram′ +import Functor.Instance.DecoratedCospan.Embed as CospanEmbed + +open import Categories.Category using (Category; _[_,_]; _[_≈_]; _[_∘_]; module Definitions) +open import Categories.Functor using (Functor; _∘F_) +open import Categories.Functor.Hom using (Hom[_][_,-]) +open import Categories.Functor.Properties using ([_]-resp-≅) +open import Categories.NaturalTransformation using (NaturalTransformation; _∘ᵥ_) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; _≃_; _ⓘˡ_) +open import Categories.NaturalTransformation.Equivalence using () renaming (_≃_ to _≋_) +open import Function.Bundles using (_⟨$⟩_) +open import Function.Construct.Composition using () renaming (function to _∘′_) +open import Functor.Exact using (RightExactFunctor; IsPushout⇒Pushout) + +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Category.Instance.Cospans using (Cospans) +open import Category.Instance.DecoratedCospans using (DecoratedCospans) +open import Category.Monoidal.Instance.Cospans.Lift {o} {ℓ} {e} using () renaming (module Square to Square′) +open import Cospan.Decorated using (DecoratedCospan) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module _ + {𝒞 : FinitelyCocompleteCategory o ℓ e} + {𝒟 : FinitelyCocompleteCategory o ℓ e} + {𝒥 : SymmetricMonoidalCategory o′ ℓ′ e′} + {𝒦 : SymmetricMonoidalCategory o′ ℓ′ e′} + {H : SymmetricMonoidalFunctor (smc 𝒞) 𝒥} + {I : SymmetricMonoidalFunctor (smc 𝒟) 𝒦} where + + module 𝒞 = FinitelyCocompleteCategory 𝒞 + module 𝒟 = FinitelyCocompleteCategory 𝒟 + module 𝒥 = SymmetricMonoidalCategory 𝒥 + module 𝒦 = SymmetricMonoidalCategory 𝒦 + module H = SymmetricMonoidalFunctor H + module I = SymmetricMonoidalFunctor I + + open CospanEmbed 𝒟 I using (L; B₁; B∘L; R∘B; ≅-L-R) + open NaturalIsomorphism using (F⇐G) + + module Square + {F G : Functor 𝒞.U 𝒟.U} + (F≃G : F ≃ G) + (⇒H≃I∘F : NaturalTransformation (Hom[ 𝒥.U ][ 𝒥.unit ,-] ∘F H.F) (Hom[ 𝒦.U ][ 𝒦.unit ,-] ∘F I.F ∘F F)) + (⇒H≃I∘G : NaturalTransformation (Hom[ 𝒥.U ][ 𝒥.unit ,-] ∘F H.F) (Hom[ 𝒦.U ][ 𝒦.unit ,-] ∘F I.F ∘F G)) + (≋ : (F⇐G (Hom[ 𝒦.U ][ 𝒦.unit ,-] ⓘˡ (I.F ⓘˡ F≃G))) ∘ᵥ ⇒H≃I∘G ≋ ⇒H≃I∘F) + where + + module F = Functor F + module G = Functor G + module ⇒H≃I∘F = NaturalTransformation ⇒H≃I∘F + module ⇒H≃I∘G = NaturalTransformation ⇒H≃I∘G + + open NaturalIsomorphism F≃G + + IF≃IG : I.F ∘F F ≃ I.F ∘F G + IF≃IG = I.F ⓘˡ F≃G + + open Morphism using (module ≅) renaming (_≅_ to _[_≅_]) + FX≅GX′ : ∀ {Z : 𝒞.Obj} → DecoratedCospans 𝒟 I [ F.₀ Z ≅ G.₀ Z ] + FX≅GX′ = [ L ]-resp-≅ FX≅GX + module FX≅GX {Z} = _[_≅_] (FX≅GX {Z}) + module FX≅GX′ {Z} = _[_≅_] (FX≅GX′ {Z}) + + module _ {X Y : 𝒞.Obj} (fg : DecoratedCospans 𝒞 H [ X , Y ]) where + + open DecoratedCospan fg renaming (f₁ to f; f₂ to g; decoration to s) + open 𝒟 using (_∘_) + + squares⇒cospan + : DecoratedCospans 𝒟 I + [ B₁ (G.₁ f ∘ FX≅GX.from) (G.₁ g ∘ FX≅GX.from) (⇒H≃I∘G.η N ⟨$⟩ s) + ≈ B₁ (F.₁ f) (F.₁ g) (⇒H≃I∘F.η N ⟨$⟩ s) + ] + squares⇒cospan = record + { cospans-≈ = Square′.squares⇒cospan F≃G cospan + ; same-deco = refl⟩∘⟨ sym 𝒦.identityʳ ○ ≋ + } + where + open 𝒦.HomReasoning + open 𝒦.Equiv + + module Cospans = Category (DecoratedCospans 𝒟 I) + + from : DecoratedCospans 𝒟 I + [ DecoratedCospans 𝒟 I [ L.₁ (⇒.η Y) ∘ B₁ (F.₁ f) (F.₁ g) (⇒H≃I∘F.η N ⟨$⟩ s) ] + ≈ DecoratedCospans 𝒟 I [ B₁ (G.₁ f) (G.₁ g) (⇒H≃I∘G.η N ⟨$⟩ s) ∘ L.₁ (⇒.η X) ] + ] + from = sym (switch-tofromˡ FX≅GX′ (refl⟩∘⟨ B∘L ○ ≅-L-R FX≅GX ⟩∘⟨refl ○ R∘B ○ squares⇒cospan)) + where + open Cospans.Equiv using (sym) + open ⇒-Reasoning (DecoratedCospans 𝒟 I) using (switch-tofromˡ) + open Cospans.HomReasoning using (refl⟩∘⟨_; _○_; _⟩∘⟨refl) + + to : DecoratedCospans 𝒟 I + [ DecoratedCospans 𝒟 I [ L.₁ (⇐.η Y) ∘ B₁ (G.₁ f) (G.₁ g) (⇒H≃I∘G.η N ⟨$⟩ s) ] ≈ DecoratedCospans 𝒟 I [ B₁ (F.₁ f) (F.₁ g) (⇒H≃I∘F.η N ⟨$⟩ s) ∘ L.₁ (⇐.η X) ] + ] + to = switch-fromtoʳ FX≅GX′ (pullʳ B∘L ○ ≅-L-R FX≅GX ⟩∘⟨refl ○ R∘B ○ squares⇒cospan) + where + open ⇒-Reasoning (DecoratedCospans 𝒟 I) using (pullʳ; switch-fromtoʳ) + open Cospans.HomReasoning using (refl⟩∘⟨_; _○_; _⟩∘⟨refl) diff --git a/Category/Monoidal/Instance/DecoratedCospans/Products.agda b/Category/Monoidal/Instance/DecoratedCospans/Products.agda new file mode 100644 index 0000000..f8ef542 --- /dev/null +++ b/Category/Monoidal/Instance/DecoratedCospans/Products.agda @@ -0,0 +1,104 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --lossy-unification #-} + +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module Category.Monoidal.Instance.DecoratedCospans.Products + {o o′ ℓ ℓ′ e e′} + (𝒞 : FinitelyCocompleteCategory o ℓ e) + {𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′} + (F : SymmetricMonoidalFunctor (smc 𝒞) 𝒟) where + +import Categories.Morphism as Morphism +import Categories.Morphism.Reasoning as ⇒-Reasoning + +open import Categories.Category using (_[_,_]; _[_≈_]; _[_∘_]; Category) +open import Categories.Category.Core using (Category) +open import Categories.Category.BinaryProducts using (BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cartesian.Bundle using (CartesianCategory) +open import Categories.Functor using (Functor; _∘F_) renaming (id to idF) +open import Categories.Functor.Monoidal.Properties using (∘-SymmetricMonoidal) +open import Categories.Functor.Monoidal.Construction.Product using (⁂-SymmetricMonoidalFunctor) +open import Categories.NaturalTransformation.Core using (NaturalTransformation; _∘ᵥ_; ntHelper) +open import Category.Instance.Properties.SymMonCat {o} {ℓ} {e} using (SymMonCat-Cartesian) +open import Category.Instance.Properties.SymMonCat {o′} {ℓ′} {e′} using () renaming (SymMonCat-Cartesian to SymMonCat-Cartesian′) +open import Category.Cartesian.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes-CC) +open import Category.Cartesian.Instance.SymMonCat {o} {ℓ} {e} using (SymMonCat-CC) +open import Data.Product.Base using (_,_) +open import Categories.Functor.Cartesian using (CartesianF) +open import Functor.Cartesian.Instance.Underlying.SymmetricMonoidal.FinitelyCocomplete {o} {ℓ} {e} using (Underlying) + +module SMC = CartesianF Underlying +module SymMonCat-CC = CartesianCategory SymMonCat-CC + +module 𝒞 = FinitelyCocompleteCategory 𝒞 +module 𝒟 = SymmetricMonoidalCategory 𝒟 + +module _ where + + open CartesianCategory FinitelyCocompletes-CC using (products) + open BinaryProducts products using (_×_) + + 𝒞×𝒞 : FinitelyCocompleteCategory o ℓ e + 𝒞×𝒞 = 𝒞 × 𝒞 + + [𝒞×𝒞]×𝒞 : FinitelyCocompleteCategory o ℓ e + [𝒞×𝒞]×𝒞 = (𝒞 × 𝒞) × 𝒞 + +module _ where + + module _ where + + open Cartesian SymMonCat-Cartesian′ using (products) + open BinaryProducts products using (_×_; _⁂_) + + 𝒟×𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′ + 𝒟×𝒟 = 𝒟 × 𝒟 + module 𝒟×𝒟 = SymmetricMonoidalCategory 𝒟×𝒟 + + [𝒟×𝒟]×𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′ + [𝒟×𝒟]×𝒟 = (𝒟 × 𝒟) × 𝒟 + module [𝒟×𝒟]×𝒟 = SymmetricMonoidalCategory [𝒟×𝒟]×𝒟 + + module _ where + + open Cartesian SymMonCat-Cartesian using (products) + open BinaryProducts products using (_×_; _⁂_) + + smc𝒞×𝒞 : SymmetricMonoidalCategory o ℓ e + smc𝒞×𝒞 = smc 𝒞 × smc 𝒞 + + smc[𝒞×𝒞]×𝒞 : SymmetricMonoidalCategory o ℓ e + smc[𝒞×𝒞]×𝒞 = (smc 𝒞×𝒞) × smc 𝒞 + + F×F′ : SymmetricMonoidalFunctor smc𝒞×𝒞 𝒟×𝒟 + F×F′ = ⁂-SymmetricMonoidalFunctor {o′} {ℓ′} {e′} {o′} {ℓ′} {e′} {𝒟} {𝒟} {o} {ℓ} {e} {o} {ℓ} {e} {smc 𝒞} {smc 𝒞} F F + + F×F : SymmetricMonoidalFunctor (smc 𝒞×𝒞) 𝒟×𝒟 + F×F = ∘-SymmetricMonoidal F×F′ from + where + open Morphism SymMonCat-CC.U using (_≅_) + ≅ : smc 𝒞×𝒞 ≅ smc𝒞×𝒞 + ≅ = SMC.×-iso 𝒞 𝒞 + open _≅_ ≅ + module F×F = SymmetricMonoidalFunctor F×F + + [F×F]×F′ : SymmetricMonoidalFunctor smc[𝒞×𝒞]×𝒞 [𝒟×𝒟]×𝒟 + [F×F]×F′ = ⁂-SymmetricMonoidalFunctor {o′} {ℓ′} {e′} {o′} {ℓ′} {e′} {𝒟×𝒟} {𝒟} {o} {ℓ} {e} {o} {ℓ} {e} {smc 𝒞×𝒞} {smc 𝒞} F×F F + + [F×F]×F : SymmetricMonoidalFunctor (smc [𝒞×𝒞]×𝒞) [𝒟×𝒟]×𝒟 + [F×F]×F = ∘-SymmetricMonoidal [F×F]×F′ from + where + open Morphism SymMonCat-CC.U using (_≅_) + ≅ : smc [𝒞×𝒞]×𝒞 ≅ smc[𝒞×𝒞]×𝒞 + ≅ = SMC.×-iso 𝒞×𝒞 𝒞 + open _≅_ ≅ + module [F×F]×F = SymmetricMonoidalFunctor [F×F]×F diff --git a/Category/Monoidal/Instance/Nat.agda b/Category/Monoidal/Instance/Nat.agda new file mode 100644 index 0000000..24b30a6 --- /dev/null +++ b/Category/Monoidal/Instance/Nat.agda @@ -0,0 +1,71 @@ +{-# OPTIONS --without-K --safe #-} + +module Category.Monoidal.Instance.Nat where + +open import Level using (0ℓ) +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory) +open import Categories.Category.Instance.Nat using (Nat; Nat-Cartesian; Nat-Cocartesian; Natop) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cocartesian using (Cocartesian; module CocartesianMonoidal; module CocartesianSymmetricMonoidal) +open import Categories.Category.Cartesian.Monoidal using (module CartesianMonoidal) +open import Categories.Category.Duality using (coCartesian⇒Cocartesian; Cocartesian⇒coCartesian) + +import Categories.Category.Cartesian.SymmetricMonoidal as CartesianSymmetricMonoidal + +Natop-Cartesian : Cartesian Natop +Natop-Cartesian = Cocartesian⇒coCartesian Nat Nat-Cocartesian + +Natop-Cocartesian : Cocartesian Natop +Natop-Cocartesian = coCartesian⇒Cocartesian Natop Nat-Cartesian + +module Monoidal where + + open MonoidalCategory + open CartesianMonoidal using () renaming (monoidal to ×-monoidal) + open CocartesianMonoidal using (+-monoidal) + + Nat,+,0 : MonoidalCategory 0ℓ 0ℓ 0ℓ + Nat,+,0 .U = Nat + Nat,+,0 .monoidal = +-monoidal Nat Nat-Cocartesian + + Nat,×,1 : MonoidalCategory 0ℓ 0ℓ 0ℓ + Nat,×,1 .U = Nat + Nat,×,1 .monoidal = ×-monoidal Nat-Cartesian + + Natop,+,0 : MonoidalCategory 0ℓ 0ℓ 0ℓ + Natop,+,0 .U = Natop + Natop,+,0 .monoidal = ×-monoidal Natop-Cartesian + + Natop,×,1 : MonoidalCategory 0ℓ 0ℓ 0ℓ + Natop,×,1 .U = Natop + Natop,×,1 .monoidal = +-monoidal Natop Natop-Cocartesian + +module Symmetric where + + open SymmetricMonoidalCategory + open CartesianMonoidal using () renaming (monoidal to ×-monoidal) + open CocartesianMonoidal using (+-monoidal) + open CartesianSymmetricMonoidal using () renaming (symmetric to ×-symmetric) + open CocartesianSymmetricMonoidal using (+-symmetric) + + Nat,+,0 : SymmetricMonoidalCategory 0ℓ 0ℓ 0ℓ + Nat,+,0 .U = Nat + Nat,+,0 .monoidal = +-monoidal Nat Nat-Cocartesian + Nat,+,0 .symmetric = +-symmetric Nat Nat-Cocartesian + + Nat,×,1 : SymmetricMonoidalCategory 0ℓ 0ℓ 0ℓ + Nat,×,1 .U = Nat + Nat,×,1 .monoidal = ×-monoidal Nat-Cartesian + Nat,×,1 .symmetric = ×-symmetric Nat Nat-Cartesian + + Natop,+,0 : SymmetricMonoidalCategory 0ℓ 0ℓ 0ℓ + Natop,+,0 .U = Natop + Natop,+,0 .monoidal = ×-monoidal Natop-Cartesian + Natop,+,0 .symmetric = ×-symmetric Natop Natop-Cartesian + + Natop,×,1 : SymmetricMonoidalCategory 0ℓ 0ℓ 0ℓ + Natop,×,1 .U = Natop + Natop,×,1 .monoidal = +-monoidal Natop Natop-Cocartesian + Natop,×,1 .symmetric = +-symmetric Natop Natop-Cocartesian + +open Symmetric public diff --git a/Cospan/Decorated.agda b/Cospan/Decorated.agda index 498a869..b5d6f07 100644 --- a/Cospan/Decorated.agda +++ b/Cospan/Decorated.agda @@ -17,7 +17,7 @@ module Cospan.Decorated module C = FinitelyCocompleteCategory C module D = SymmetricMonoidalCategory D -open import Category.Instance.Cospans C using (Cospan) +open import Category.Diagram.Cospan C using (Cospan) open import Level using (_⊔_) diff --git a/Data/CMonoid.agda b/Data/CMonoid.agda new file mode 100644 index 0000000..8aaf869 --- /dev/null +++ b/Data/CMonoid.agda @@ -0,0 +1,77 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +module Data.CMonoid {c ℓ : Level} where + +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Category.Instance.Setoids.SymmetricMonoidal {c} {ℓ} using (Setoids-×; ×-symmetric′) +open import Object.Monoid.Commutative using (CommutativeMonoid; CommutativeMonoid⇒) +open import Categories.Object.Monoid using (Monoid) + +open import Data.Monoid {c} {ℓ} using (toMonoid; fromMonoid; toMonoid⇒) + +import Algebra.Bundles as Alg + +open import Data.Setoid using (∣_∣) +open import Relation.Binary using (Setoid) +open import Function using (Func; _⟨$⟩_) +open import Data.Product using (curry′; uncurry′; _,_) +open Func + +-- A commutative monoid object in the (symmetric monoidal) category of setoids +-- is just a commutative monoid + +toCMonoid : CommutativeMonoid Setoids-×.symmetric → Alg.CommutativeMonoid c ℓ +toCMonoid M = record + { M + ; isCommutativeMonoid = record + { isMonoid = M.isMonoid + ; comm = comm + } + } + where + open CommutativeMonoid M using (monoid; commutative; μ) + module M = Alg.Monoid (toMonoid monoid) + opaque + unfolding toMonoid + comm : (x y : M.Carrier) → x M.∙ y M.≈ y M.∙ x + comm x y = commutative {x , y} + +open import Function.Construct.Constant using () renaming (function to Const) +open import Data.Setoid.Unit using (⊤ₛ) + +fromCMonoid : Alg.CommutativeMonoid c ℓ → CommutativeMonoid Setoids-×.symmetric +fromCMonoid M = record + { M + ; isCommutativeMonoid = record + { isMonoid = M.isMonoid + ; commutative = commutative + } + } + where + open Alg.CommutativeMonoid M using (monoid; comm) + module M = Monoid (fromMonoid monoid) + open Setoids-× using (_≈_; _∘_; module braiding) + opaque + unfolding toMonoid + commutative : M.μ ≈ M.μ ∘ braiding.⇒.η _ + commutative {x , y} = comm x y + +-- A morphism of monoids in the (monoidal) category of setoids is a monoid homomorphism + +module _ (M N : CommutativeMonoid Setoids-×.symmetric) where + + module M = Alg.CommutativeMonoid (toCMonoid M) + module N = Alg.CommutativeMonoid (toCMonoid N) + + open import Data.Product using (Σ; _,_) + open import Function using (_⟶ₛ_) + open import Algebra.Morphism using (IsMonoidHomomorphism) + open CommutativeMonoid + open CommutativeMonoid⇒ + + toCMonoid⇒ + : CommutativeMonoid⇒ Setoids-×.symmetric M N + → Σ (M.setoid ⟶ₛ N.setoid) (λ f + → IsMonoidHomomorphism M.rawMonoid N.rawMonoid (to f)) + toCMonoid⇒ f = toMonoid⇒ (monoid M) (monoid N) (monoid⇒ f) diff --git a/Data/Castable.agda b/Data/Castable.agda new file mode 100644 index 0000000..4f85b3d --- /dev/null +++ b/Data/Castable.agda @@ -0,0 +1,50 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Castable where + +open import Level using (Level; suc; _⊔_) +open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst; cong; module ≡-Reasoning) +open import Relation.Binary using (Sym; Trans; _⇒_) + +record IsCastable {ℓ₁ ℓ₂ : Level} {A : Set ℓ₁} (B : A → Set ℓ₂) : Set (ℓ₁ ⊔ ℓ₂) where + + field + cast : {e e′ : A} → .(e ≡ e′) → B e → B e′ + cast-trans + : {m n o : A} + → .(eq₁ : m ≡ n) + .(eq₂ : n ≡ o) + (x : B m) + → cast eq₂ (cast eq₁ x) ≡ cast (trans eq₁ eq₂) x + cast-is-id : {m : A} .(eq : m ≡ m) (x : B m) → cast eq x ≡ x + subst-is-cast : {m n : A} (eq : m ≡ n) (x : B m) → subst B eq x ≡ cast eq x + + infix 3 _≈[_]_ + _≈[_]_ : {n m : A} → B n → .(eq : n ≡ m) → B m → Set ℓ₂ + _≈[_]_ x eq y = cast eq x ≡ y + + ≈-reflexive : {n : A} → _≡_ ⇒ (λ xs ys → _≈[_]_ {n} xs refl ys) + ≈-reflexive {n} {x} {y} eq = trans (cast-is-id refl x) eq + + ≈-sym : {m n : A} .{m≡n : m ≡ n} → Sym _≈[ m≡n ]_ _≈[ sym m≡n ]_ + ≈-sym {m} {n} {m≡n} {x} {y} x≡y = begin + cast (sym m≡n) y ≡⟨ cong (cast (sym m≡n)) x≡y ⟨ + cast (sym m≡n) (cast m≡n x) ≡⟨ cast-trans m≡n (sym m≡n) x ⟩ + cast (trans m≡n (sym m≡n)) x ≡⟨ cast-is-id (trans m≡n (sym m≡n)) x ⟩ + x ∎ + where + open ≡-Reasoning + + ≈-trans : {m n o : A} .{m≡n : m ≡ n} .{n≡o : n ≡ o} → Trans _≈[ m≡n ]_ _≈[ n≡o ]_ _≈[ trans m≡n n≡o ]_ + ≈-trans {m} {n} {o} {m≡n} {n≡o} {x} {y} {z} x≡y y≡z = begin + cast (trans m≡n n≡o) x ≡⟨ cast-trans m≡n n≡o x ⟨ + cast n≡o (cast m≡n x) ≡⟨ cong (cast n≡o) x≡y ⟩ + cast n≡o y ≡⟨ y≡z ⟩ + z ∎ + where + open ≡-Reasoning + +record Castable {ℓ₁ ℓ₂ : Level} {A : Set ℓ₁} : Set (suc (ℓ₁ ⊔ ℓ₂)) where + field + B : A → Set ℓ₂ + isCastable : IsCastable B diff --git a/Data/Circuit.agda b/Data/Circuit.agda new file mode 100644 index 0000000..473e4fc --- /dev/null +++ b/Data/Circuit.agda @@ -0,0 +1,42 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module Data.Circuit {ℓ : Level} where + +open import Data.Circuit.Gate using (Gates) + +import Data.List as List +import Data.Hypergraph {ℓ} Gates as Hypergraph + +open import Data.Fin using (Fin) +open import Data.Nat using (ℕ) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (map⁺) +open import Function using (Func; _⟶ₛ_; _∘_) +open import Relation.Binary using (Setoid) + +open Func + +open Hypergraph using (Multiset∘Edgeₛ) +open Hypergraph + using (_≈_ ; mk≈ ; module Edge; edgesₛ) + renaming + ( Hypergraph to Circuit + ; Hypergraphₛ to Circuitₛ + ; mkHypergraph to mkCircuit + ; mkHypergraphₛ to mkCircuitₛ + ) + public +open List using ([]) + +map : {n m : ℕ} → (Fin n → Fin m) → Circuit n → Circuit m +map f (mkCircuit edges) = mkCircuit (List.map (Edge.map f) edges) + +discrete : (n : ℕ) → Circuit n +discrete n = mkCircuit [] + +open Edge using (Edgeₛ) + +mapₛ : {n m : ℕ} → (Fin n → Fin m) → Circuitₛ n ⟶ₛ Circuitₛ m +mapₛ f .to = map f +mapₛ {n} {m} f .cong (mk≈ x≈y) = mk≈ (map⁺ (Edgeₛ n) (Edgeₛ m) (cong (Edge.mapₛ f)) x≈y) diff --git a/Data/Circuit/Convert.agda b/Data/Circuit/Convert.agda new file mode 100644 index 0000000..d5abd35 --- /dev/null +++ b/Data/Circuit/Convert.agda @@ -0,0 +1,181 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Circuit.Convert where + +open import Level using (0ℓ) + +import Data.Vec as Vec +import Data.Vec.Relation.Binary.Equality.Cast as VecCast +import Data.List.Relation.Binary.Permutation.Propositional as L +import Data.Vec.Functional.Relation.Binary.Permutation as V +import DecorationFunctor.Hypergraph.Labeled {0ℓ} {0ℓ} as LabeledHypergraph + +open import Data.Nat.Base using (ℕ) +open import Data.Circuit.Gate using (Gate; Gates; cast-gate; cast-gate-is-id; subst-is-cast-gate) +open import Data.Circuit {0ℓ} {0ℓ} using (Circuit; Circuitₛ; _≈_; mkCircuit; module Edge; mk≈) +open import Data.Fin.Base using (Fin) +open import Data.Product.Base using (_,_) +open import Data.Permutation using (fromList-↭; toList-↭) +open import Data.List using (length) +open import Data.Vec.Functional using (toVec; fromVec; toList; fromList) +open import Function.Bundles using (Equivalence; _↔_) +open import Function.Base using (_∘_; id) +open import Data.Vec.Properties using (tabulate-cong; tabulate-∘; map-cast) +open import Data.Fin.Base using () renaming (cast to fincast) +open import Data.Fin.Properties using () renaming (cast-trans to fincast-trans; cast-is-id to fincast-is-id) +open import Data.List.Relation.Binary.Permutation.Homogeneous using (Permutation) +open import Data.Product.Base using (proj₁; proj₂; _×_) +open import Data.Fin.Permutation using (flip; _⟨$⟩ˡ_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) + +open LabeledHypergraph using (Hypergraph-same) renaming (Hypergraph to Hypergraph′; Hypergraph-setoid to Hypergraph-Setoid′) + +to : {v : ℕ} → Circuit v → Hypergraph′ v +to C = record + { h = length edges + ; a = arity ∘ fromList edges + ; j = fromVec ∘ ports ∘ fromList edges + ; l = label ∘ fromList edges + } + where + open Edge.Edge using (arity; ports; label) + open Circuit C + +from : {v : ℕ} → Hypergraph′ v → Circuit v +from {v} H = record + { edges = toList asEdge + } + where + open Hypergraph′ H + asEdge : Fin h → Edge.Edge v + asEdge e = record { label = l e ; ports = toVec (j e) } + +to-cong : {v : ℕ} {H H′ : Circuit v} → H ≈ H′ → Hypergraph-same (to H) (to H′) +to-cong {v} {H} {H′} ≈H = record + { ↔h = flip ρ + ; ≗a = ≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + open Edge.Edge using (arity; ports; label) + open _≈_ ≈H + open import Data.Fin.Permutation using (_⟨$⟩ʳ_; _⟨$⟩ˡ_; Permutation′; inverseʳ) + open import Data.Fin.Base using (cast) + open import Data.Fin.Properties using (cast-is-id) + ρ : Fin (length H′.edges) ↔ Fin (length H.edges) + ρ = proj₁ (fromList-↭ ↭-edges) + + open ≡.≡-Reasoning + edges≗ρ∘edges′ : (i : Fin (length H.edges)) → fromList H.edges i ≡ fromList H′.edges (ρ ⟨$⟩ˡ i) + edges≗ρ∘edges′ i = begin + fromList H.edges i ≡⟨ ≡.cong (fromList H.edges) (inverseʳ ρ) ⟨ + fromList H.edges (ρ ⟨$⟩ʳ (ρ ⟨$⟩ˡ i)) ≡⟨ proj₂ (fromList-↭ ↭-edges) (ρ ⟨$⟩ˡ i) ⟩ + fromList H′.edges (ρ ⟨$⟩ˡ i) ∎ + + ≗a : (e : Fin (Hypergraph′.h (to H))) + → Hypergraph′.a (to H) e + ≡ arity (fromList H′.edges (ρ ⟨$⟩ˡ e)) + ≗a = ≡.cong arity ∘ edges≗ρ∘edges′ + + ≗j : (e : Fin (Hypergraph′.h (to H))) + (i : Fin (Hypergraph′.a (to H) e)) + → fromVec (ports (fromList H.edges e)) i + ≡ fromVec (ports (fromList H′.edges (ρ ⟨$⟩ˡ e))) (cast (≗a e) i) + ≗j e i + rewrite edges≗ρ∘edges′ e + rewrite cast-is-id ≡.refl i = ≡.refl + + ≗l : (e : Fin (Hypergraph′.h (to H))) + → label (fromList H.edges e) + ≡ cast-gate (≡.sym (≗a e)) (label (fromList H′.edges (ρ ⟨$⟩ˡ e))) + ≗l e + rewrite edges≗ρ∘edges′ e + rewrite cast-gate-is-id ≡.refl (label (fromList H′.edges (ρ ⟨$⟩ˡ e))) = + ≡.refl + +module _ {v : ℕ} where + open import Data.Hypergraph.Label using (HypergraphLabel) + open HypergraphLabel Gates using (isCastable) + open import Data.Castable using (IsCastable) + open IsCastable isCastable using (≈-reflexive; ≈-sym; ≈-trans) + from-cong + : {H H′ : Hypergraph′ v} + → Hypergraph-same H H′ + → from H ≈ from H′ + from-cong {H} {H′} ≈H = mk≈ (toList-↭ (flip ↔h , H∘ρ≗H′)) + where + + module H = Hypergraph′ H + module H′ = Hypergraph′ H′ + open Hypergraph′ + open Hypergraph-same ≈H using (↔h; ≗a; ≗l; ≗j; inverseˡ) renaming (from to f; to to t) + asEdge : (H : Hypergraph′ v) → Fin (h H) → Edge.Edge v + asEdge H e = record { label = l H e ; ports = toVec (j H e) } + + to-from : (e : Fin H′.h) → t (f e) ≡ e + to-from e = inverseˡ ≡.refl + + a∘to-from : (e : Fin H′.h) → H′.a (t (f e)) ≡ H′.a e + a∘to-from = ≡.cong H′.a ∘ to-from + + ≗a′ : (e : Fin H′.h) → H.a (f e) ≡ H′.a e + ≗a′ e = ≡.trans (≗a (f e)) (a∘to-from e) + + l≗ : (e : Fin H.h) → cast-gate (≗a e) (H.l e) ≡ H′.l (t e) + l≗ e = ≈-sym (≡.sym (≗l e)) + + l∘to-from : (e : Fin H′.h) → cast-gate (a∘to-from e) (H′.l (t (f e))) ≡ H′.l e + l∘to-from e rewrite to-from e = ≈-reflexive ≡.refl + + ≗l′ : (e : Fin H′.h) → cast-gate (≗a′ e) (H.l (f e)) ≡ H′.l e + ≗l′ e = ≈-trans {H.a _} (l≗ (f e)) (l∘to-from e) + + j∘to-from + : (e : Fin H′.h) (i : Fin (H′.a (t (f e)))) + → H′.j (t (f e)) i + ≡ H′.j e (fincast (a∘to-from e) i) + j∘to-from e i rewrite to-from e = ≡.cong (H′.j e) (≡.sym (fincast-is-id ≡.refl i)) + + open ≡.≡-Reasoning + + ≗j′ : (e : Fin H′.h) (i : Fin (H.a (f e))) → H.j (f e) i ≡ H′.j e (fincast (≗a′ e) i) + ≗j′ e i = begin + H.j (f e) i ≡⟨ ≗j (f e) i ⟩ + H′.j (t (f e)) (fincast _ i) ≡⟨ j∘to-from e (fincast _ i) ⟩ + H′.j e (fincast (a∘to-from e) (fincast _ i)) ≡⟨ ≡.cong (H′.j e) (fincast-trans (≗a (f e)) _ i) ⟩ + H′.j e (fincast (≗a′ e) i) ∎ + + cast-toVec + : {n m : ℕ} + {A : Set} + (m≡n : m ≡ n) + (f : Fin n → A) + → Vec.cast m≡n (toVec (f ∘ fincast m≡n)) ≡ toVec f + cast-toVec m≡n f rewrite m≡n = begin + Vec.cast _ (toVec (f ∘ (fincast _))) ≡⟨ VecCast.cast-is-id ≡.refl (toVec (f ∘ fincast ≡.refl)) ⟩ + toVec (f ∘ fincast _) ≡⟨ tabulate-∘ f (fincast ≡.refl) ⟩ + Vec.map f (toVec (fincast _)) ≡⟨ ≡.cong (Vec.map f) (tabulate-cong (fincast-is-id ≡.refl)) ⟩ + Vec.map f (toVec id) ≡⟨ tabulate-∘ f id ⟨ + toVec f ∎ + + ≗p′ : (e : Fin H′.h) → Vec.cast (≗a′ e) (toVec (H.j (f e))) ≡ toVec (H′.j e) + ≗p′ e = begin + Vec.cast (≗a′ e) (toVec (H.j (f e))) ≡⟨ ≡.cong (Vec.cast (≗a′ e)) (tabulate-cong (≗j′ e)) ⟩ + Vec.cast _ (toVec (H′.j e ∘ fincast _)) ≡⟨ cast-toVec (≗a′ e) (H′.j e) ⟩ + toVec (H′.j e) ∎ + + H∘ρ≗H′ : (e : Fin H′.h) → asEdge H (↔h ⟨$⟩ˡ e) ≡ asEdge H′ e + H∘ρ≗H′ e = Edge.≈⇒≡ record + { ≡arity = ≗a′ e + ; ≡label = ≗l′ e + ; ≡ports = ≗p′ e + } + +equiv : (v : ℕ) → Equivalence (Circuitₛ v) (Hypergraph-Setoid′ v) +equiv v = record + { to = to + ; from = from + ; to-cong = to-cong + ; from-cong = from-cong + } diff --git a/Data/Circuit/Gate.agda b/Data/Circuit/Gate.agda new file mode 100644 index 0000000..f4b55de --- /dev/null +++ b/Data/Circuit/Gate.agda @@ -0,0 +1,137 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Circuit.Gate where + +open import Level using (0ℓ) +open import Data.Castable using (Castable) +open import Data.Hypergraph.Label using (HypergraphLabel) +open import Data.String using (String) +open import Data.Nat.Base using (ℕ; _≤_) +open import Data.Nat.Properties using (≤-refl; ≤-trans; ≤-antisym; ≤-total) +open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst; isEquivalence; cong) +import Relation.Binary.PropositionalEquality as ≡ + +import Data.Nat as Nat +import Data.Fin as Fin + +data Gate : ℕ → Set where + ZERO : Gate 1 + ONE : Gate 1 + ID : Gate 2 + NOT : Gate 2 + AND : Gate 3 + OR : Gate 3 + XOR : Gate 3 + NAND : Gate 3 + NOR : Gate 3 + XNOR : Gate 3 + +cast-gate : {e e′ : ℕ} → .(e ≡ e′) → Gate e → Gate e′ +cast-gate {1} {1} eq g = g +cast-gate {2} {2} eq g = g +cast-gate {3} {3} eq g = g + +cast-gate-trans + : {m n o : ℕ} + → .(eq₁ : m ≡ n) + .(eq₂ : n ≡ o) + (g : Gate m) + → cast-gate eq₂ (cast-gate eq₁ g) ≡ cast-gate (trans eq₁ eq₂) g +cast-gate-trans {1} {1} {1} eq₁ eq₂ g = refl +cast-gate-trans {2} {2} {2} eq₁ eq₂ g = refl +cast-gate-trans {3} {3} {3} eq₁ eq₂ g = refl + +cast-gate-is-id : {m : ℕ} .(eq : m ≡ m) (g : Gate m) → cast-gate eq g ≡ g +cast-gate-is-id {1} eq g = refl +cast-gate-is-id {2} eq g = refl +cast-gate-is-id {3} eq g = refl + +subst-is-cast-gate : {m n : ℕ} (eq : m ≡ n) (g : Gate m) → subst Gate eq g ≡ cast-gate eq g +subst-is-cast-gate refl g = sym (cast-gate-is-id refl g) + +GateCastable : Castable +GateCastable = record + { B = Gate + ; isCastable = record + { cast = cast-gate + ; cast-trans = cast-gate-trans + ; cast-is-id = cast-gate-is-id + ; subst-is-cast = subst-is-cast-gate + } + } + +showGate : (n : ℕ) → Gate n → String +showGate _ ZERO = "ZERO" +showGate _ ONE = "ONE" +showGate _ ID = "ID" +showGate _ NOT = "NOT" +showGate _ AND = "AND" +showGate _ OR = "OR" +showGate _ XOR = "XOR" +showGate _ NAND = "NAND" +showGate _ NOR = "NOR" +showGate _ XNOR = "XNOR" + +toℕ : (n : ℕ) → Gate n → ℕ +toℕ 1 ZERO = 0 +toℕ 1 ONE = 1 +toℕ 2 ID = 0 +toℕ 2 NOT = 1 +toℕ 3 AND = 0 +toℕ 3 OR = 1 +toℕ 3 XOR = 2 +toℕ 3 NAND = 3 +toℕ 3 NOR = 4 +toℕ 3 XNOR = 5 + +toℕ-injective : {n : ℕ} {x y : Gate n} → toℕ n x ≡ toℕ n y → x ≡ y +toℕ-injective {1} {ZERO} {ZERO} refl = refl +toℕ-injective {1} {ONE} {ONE} refl = refl +toℕ-injective {2} {ID} {ID} refl = refl +toℕ-injective {2} {NOT} {NOT} refl = refl +toℕ-injective {3} {AND} {AND} refl = refl +toℕ-injective {3} {OR} {OR} refl = refl +toℕ-injective {3} {XOR} {XOR} refl = refl +toℕ-injective {3} {NAND} {NAND} refl = refl +toℕ-injective {3} {NOR} {NOR} refl = refl +toℕ-injective {3} {XNOR} {XNOR} refl = refl + +open import Relation.Binary using (Rel; Decidable; DecidableEquality) +import Relation.Nullary.Decidable as Dec + +_[_≤_] : (n : ℕ) → Rel (Gate n) 0ℓ +_[_≤_] n x y = toℕ n x ≤ toℕ n y + +_≟_ : {n : ℕ} → DecidableEquality (Gate n) +_≟_ {n} x y = Dec.map′ toℕ-injective (cong (toℕ n)) (toℕ n x Nat.≟ toℕ n y) + +_≤?_ : {n : ℕ} → Decidable (n [_≤_]) +_≤?_ {n} x y = toℕ n x Nat.≤? toℕ n y + +Gates : HypergraphLabel +Gates = record + { Label = Gate + ; showLabel = showGate + ; isCastable = record + { cast = cast-gate + ; cast-trans = cast-gate-trans + ; cast-is-id = cast-gate-is-id + ; subst-is-cast = subst-is-cast-gate + } + ; _[_≤_] = λ n x y → toℕ n x ≤ toℕ n y + ; isDecTotalOrder = λ n → record + { isTotalOrder = record + { isPartialOrder = record + { isPreorder = record + { isEquivalence = isEquivalence + ; reflexive = λ { refl → ≤-refl } + ; trans = ≤-trans + } + ; antisym = λ i≤j j≤i → toℕ-injective (≤-antisym i≤j j≤i) + } + ; total = λ { x y → ≤-total (toℕ n x) (toℕ n y) } + } + ; _≟_ = _≟_ + ; _≤?_ = _≤?_ + } + } diff --git a/Data/Circuit/Merge.agda b/Data/Circuit/Merge.agda new file mode 100644 index 0000000..9cf180a --- /dev/null +++ b/Data/Circuit/Merge.agda @@ -0,0 +1,427 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Circuit.Merge where + +open import Data.Nat.Base using (ℕ) +open import Data.Fin.Base using (Fin; pinch; punchIn; punchOut; splitAt) +open import Data.Fin.Properties using (punchInᵢ≢i; punchIn-punchOut) +open import Data.Bool.Properties using (if-eta) +open import Data.Bool using (Bool; if_then_else_) +open import Data.Circuit.Value using (Value; join; join-comm; join-assoc) +open import Data.Sum.Properties using ([,]-cong; [,-]-cong; [-,]-cong; [,]-∘; [,]-map) +open import Data.Subset.Functional + using + ( Subset + ; ⁅_⁆ ; ⊥ ; ⁅⁆∘ρ + ; foldl ; foldl-cong₁ ; foldl-cong₂ + ; foldl-[] ; foldl-suc + ; foldl-⊥ ; foldl-⁅⁆ + ; foldl-fusion + ) +open import Data.Vector as V using (Vector; head; tail; removeAt) +open import Data.Vec.Functional using (_++_) +open import Data.Fin.Permutation + using + ( Permutation + ; Permutation′ + ; _⟨$⟩ˡ_ ; _⟨$⟩ʳ_ + ; inverseˡ ; inverseʳ + ; id + ; flip + ; insert ; remove + ; punchIn-permute + ) +open import Data.Product using (Σ-syntax; _,_) +open import Data.Fin.Preimage using (preimage; preimage-cong₁; preimage-cong₂) +open import Function.Base using (∣_⟩-_; _∘_; case_of_; _$_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≢_; _≗_; module ≡-Reasoning) + +open Value using (U) +open ℕ +open Fin +open Bool + +open ≡-Reasoning + +_when_ : Value → Bool → Value +x when b = if b then x else U + +opaque + merge-with : {A : ℕ} → Value → Vector Value A → Subset A → Value + merge-with e v = foldl (∣ join ⟩- v) e + + merge-with-cong : {A : ℕ} {v₁ v₂ : Vector Value A} (e : Value) → v₁ ≗ v₂ → merge-with e v₁ ≗ merge-with e v₂ + merge-with-cong e v₁≗v₂ = foldl-cong₁ (λ x → ≡.cong (join x) ∘ v₁≗v₂) e + + merge-with-cong₂ : {A : ℕ} (e : Value) (v : Vector Value A) {S₁ S₂ : Subset A} → S₁ ≗ S₂ → merge-with e v S₁ ≡ merge-with e v S₂ + merge-with-cong₂ e v = foldl-cong₂ (∣ join ⟩- v) e + + merge-with-⊥ : {A : ℕ} (e : Value) (v : Vector Value A) → merge-with e v ⊥ ≡ e + merge-with-⊥ e v = foldl-⊥ (∣ join ⟩- v) e + + merge-with-[] : (e : Value) (v : Vector Value 0) (S : Subset 0) → merge-with e v S ≡ e + merge-with-[] e v = foldl-[] (∣ join ⟩- v) e + + merge-with-suc + : {A : ℕ} (e : Value) (v : Vector Value (suc A)) (S : Subset (suc A)) + → merge-with e v S + ≡ merge-with (if head S then join e (head v) else e) (tail v) (tail S) + merge-with-suc e v = foldl-suc (∣ join ⟩- v) e + + merge-with-join + : {A : ℕ} + (x y : Value) + (v : Vector Value A) + → merge-with (join x y) v ≗ join x ∘ merge-with y v + merge-with-join {A} x y v S = ≡.sym (foldl-fusion (join x) fuse y S) + where + fuse : (acc : Value) (k : Fin A) → join x (join acc (v k)) ≡ join (join x acc) (v k) + fuse acc k = ≡.sym (join-assoc x acc (v k)) + + merge-with-⁅⁆ : {A : ℕ} (e : Value) (v : Vector Value A) (x : Fin A) → merge-with e v ⁅ x ⁆ ≡ join e (v x) + merge-with-⁅⁆ e v = foldl-⁅⁆ (∣ join ⟩- v) e + +merge-with-U : {A : ℕ} (e : Value) (S : Subset A) → merge-with e (λ _ → U) S ≡ e +merge-with-U {zero} e S = merge-with-[] e (λ _ → U) S +merge-with-U {suc A} e S = begin + merge-with e (λ _ → U) S ≡⟨ merge-with-suc e (λ _ → U) S ⟩ + merge-with + (if head S then join e U else e) + (tail (λ _ → U)) (tail S) ≡⟨ ≡.cong (λ h → merge-with (if head S then h else e) _ _) (join-comm e U) ⟩ + merge-with + (if head S then e else e) + (tail (λ _ → U)) (tail S) ≡⟨ ≡.cong (λ h → merge-with h (λ _ → U) (tail S)) (if-eta (head S)) ⟩ + merge-with e (tail (λ _ → U)) (tail S) ≡⟨⟩ + merge-with e (λ _ → U) (tail S) ≡⟨ merge-with-U e (tail S) ⟩ + e ∎ + +merge : {A : ℕ} → Vector Value A → Subset A → Value +merge v = merge-with U v + +merge-cong₁ : {A : ℕ} {v₁ v₂ : Vector Value A} → v₁ ≗ v₂ → merge v₁ ≗ merge v₂ +merge-cong₁ = merge-with-cong U + +merge-cong₂ : {A : ℕ} (v : Vector Value A) {S₁ S₂ : Subset A} → S₁ ≗ S₂ → merge v S₁ ≡ merge v S₂ +merge-cong₂ = merge-with-cong₂ U + +merge-⊥ : {A : ℕ} (v : Vector Value A) → merge v ⊥ ≡ U +merge-⊥ = merge-with-⊥ U + +merge-[] : (v : Vector Value 0) (S : Subset 0) → merge v S ≡ U +merge-[] = merge-with-[] U + +merge-[]₂ : {v₁ v₂ : Vector Value 0} {S₁ S₂ : Subset 0} → merge v₁ S₁ ≡ merge v₂ S₂ +merge-[]₂ {v₁} {v₂} {S₁} {S₂} = ≡.trans (merge-[] v₁ S₁) (≡.sym (merge-[] v₂ S₂)) + +merge-⁅⁆ : {A : ℕ} (v : Vector Value A) (x : Fin A) → merge v ⁅ x ⁆ ≡ v x +merge-⁅⁆ = merge-with-⁅⁆ U + +join-merge : {A : ℕ} (e : Value) (v : Vector Value A) (S : Subset A) → join e (merge v S) ≡ merge-with e v S +join-merge e v S = ≡.sym (≡.trans (≡.cong (λ h → merge-with h v S) (join-comm U e)) (merge-with-join e U v S)) + +merge-suc + : {A : ℕ} (v : Vector Value (suc A)) (S : Subset (suc A)) + → merge v S + ≡ merge-with (head v when head S) (tail v) (tail S) +merge-suc = merge-with-suc U + +insert-f0-0 + : {A B : ℕ} + (f : Fin (suc A) → Fin (suc B)) + → Σ[ ρ ∈ Permutation′ (suc B) ] (ρ ⟨$⟩ʳ (f zero) ≡ zero) +insert-f0-0 f = insert (f zero) zero id , help + where + open import Data.Fin using (_≟_) + open import Relation.Nullary.Decidable using (yes; no) + help : insert (f zero) zero id ⟨$⟩ʳ f zero ≡ zero + help with f zero ≟ f zero + ... | yes _ = ≡.refl + ... | no f0≢f0 with () ← f0≢f0 ≡.refl + +merge-removeAt + : {A : ℕ} + (k : Fin (suc A)) + (v : Vector Value (suc A)) + (S : Subset (suc A)) + → merge v S ≡ join (v k when S k) (merge (removeAt v k) (removeAt S k)) +merge-removeAt {A} zero v S = begin + merge-with U v S ≡⟨ merge-suc v S ⟩ + merge-with (head v when head S) (tail v) (tail S) ≡⟨ join-merge (head v when head S) (tail v) (tail S) ⟨ + join (head v when head S) (merge-with U (tail v) (tail S)) ∎ +merge-removeAt {suc A} (suc k) v S = begin + merge-with U v S ≡⟨ merge-suc v S ⟩ + merge-with v0? (tail v) (tail S) ≡⟨ join-merge _ (tail v) (tail S) ⟨ + join v0? (merge (tail v) (tail S)) ≡⟨ ≡.cong (join v0?) (merge-removeAt k (tail v) (tail S)) ⟩ + join v0? (join vk? (merge (tail v-) (tail S-))) ≡⟨ join-assoc (head v when head S) _ _ ⟨ + join (join v0? vk?) (merge (tail v-) (tail S-)) ≡⟨ ≡.cong (λ h → join h (merge (tail v-) (tail S-))) (join-comm (head v- when head S-) _) ⟩ + join (join vk? v0?) (merge (tail v-) (tail S-)) ≡⟨ join-assoc (tail v k when tail S k) _ _ ⟩ + join vk? (join v0? (merge (tail v-) (tail S-))) ≡⟨ ≡.cong (join vk?) (join-merge _ (tail v-) (tail S-)) ⟩ + join vk? (merge-with v0? (tail v-) (tail S-)) ≡⟨ ≡.cong (join vk?) (merge-suc v- S-) ⟨ + join vk? (merge v- S-) ∎ + where + v0? vk? : Value + v0? = head v when head S + vk? = tail v k when tail S k + v- : Vector Value (suc A) + v- = removeAt v (suc k) + S- : Subset (suc A) + S- = removeAt S (suc k) + +import Function.Structures as FunctionStructures +open module FStruct {A B : Set} = FunctionStructures {_} {_} {_} {_} {A} _≡_ {B} _≡_ using (IsInverse) +open IsInverse using () renaming (inverseˡ to invˡ; inverseʳ to invʳ) + +merge-preimage-ρ + : {A B : ℕ} + → (ρ : Permutation A B) + → (v : Vector Value A) + (S : Subset B) + → merge v (preimage (ρ ⟨$⟩ʳ_) S) ≡ merge (v ∘ (ρ ⟨$⟩ˡ_)) S +merge-preimage-ρ {zero} {zero} ρ v S = merge-[]₂ +merge-preimage-ρ {zero} {suc B} ρ v S with () ← ρ ⟨$⟩ˡ zero +merge-preimage-ρ {suc A} {zero} ρ v S with () ← ρ ⟨$⟩ʳ zero +merge-preimage-ρ {suc A} {suc B} ρ v S = begin + merge v (preimage ρʳ S) ≡⟨ merge-removeAt (head ρˡ) v (preimage ρʳ S) ⟩ + join + (head (v ∘ ρˡ) when S (ρʳ (ρˡ zero))) + (merge v- [preimageρʳS]-) ≡⟨ ≡.cong (λ h → join h (merge v- [preimageρʳS]-)) ≡vρˡ0? ⟩ + join vρˡ0? (merge v- [preimageρʳS]-) ≡⟨ ≡.cong (join vρˡ0?) (merge-cong₂ v- preimage-) ⟩ + join vρˡ0? (merge v- (preimage ρʳ- S-)) ≡⟨ ≡.cong (join vρˡ0?) (merge-preimage-ρ ρ- v- S-) ⟩ + join vρˡ0? (merge (v- ∘ ρˡ-) S-) ≡⟨ ≡.cong (join vρˡ0?) (merge-cong₁ v∘ρˡ- S-) ⟩ + join vρˡ0? (merge (tail (v ∘ ρˡ)) S-) ≡⟨ join-merge vρˡ0? (tail (v ∘ ρˡ)) S- ⟩ + merge-with vρˡ0? (tail (v ∘ ρˡ)) S- ≡⟨ merge-suc (v ∘ ρˡ) S ⟨ + merge (v ∘ ρˡ) S ∎ + where + ρˡ : Fin (suc B) → Fin (suc A) + ρˡ = ρ ⟨$⟩ˡ_ + ρʳ : Fin (suc A) → Fin (suc B) + ρʳ = ρ ⟨$⟩ʳ_ + ρ- : Permutation A B + ρ- = remove (head ρˡ) ρ + ρˡ- : Fin B → Fin A + ρˡ- = ρ- ⟨$⟩ˡ_ + ρʳ- : Fin A → Fin B + ρʳ- = ρ- ⟨$⟩ʳ_ + v- : Vector Value A + v- = removeAt v (head ρˡ) + [preimageρʳS]- : Subset A + [preimageρʳS]- = removeAt (preimage ρʳ S) (head ρˡ) + S- : Subset B + S- = tail S + vρˡ0? : Value + vρˡ0? = head (v ∘ ρˡ) when head S + ≡vρˡ0? : head (v ∘ ρˡ) when S (ρʳ (head ρˡ)) ≡ head (v ∘ ρˡ) when head S + ≡vρˡ0? = ≡.cong ((head (v ∘ ρˡ) when_) ∘ S) (inverseʳ ρ) + v∘ρˡ- : v- ∘ ρˡ- ≗ tail (v ∘ ρˡ) + v∘ρˡ- x = begin + v- (ρˡ- x) ≡⟨⟩ + v (punchIn (head ρˡ) (punchOut {A} {head ρˡ} _)) ≡⟨ ≡.cong v (punchIn-punchOut _) ⟩ + v (ρˡ (punchIn (ρʳ (ρˡ zero)) x)) ≡⟨ ≡.cong (λ h → v (ρˡ (punchIn h x))) (inverseʳ ρ) ⟩ + v (ρˡ (punchIn zero x)) ≡⟨⟩ + v (ρˡ (suc x)) ≡⟨⟩ + tail (v ∘ ρˡ) x ∎ + preimage- : [preimageρʳS]- ≗ preimage ρʳ- S- + preimage- x = begin + [preimageρʳS]- x ≡⟨⟩ + removeAt (preimage ρʳ S) (head ρˡ) x ≡⟨⟩ + S (ρʳ (punchIn (head ρˡ) x)) ≡⟨ ≡.cong S (punchIn-permute ρ (head ρˡ) x) ⟩ + S (punchIn (ρʳ (head ρˡ)) (ρʳ- x)) ≡⟨⟩ + S (punchIn (ρʳ (ρˡ zero)) (ρʳ- x)) ≡⟨ ≡.cong (λ h → S (punchIn h (ρʳ- x))) (inverseʳ ρ) ⟩ + S (punchIn zero (ρʳ- x)) ≡⟨⟩ + S (suc (ρʳ- x)) ≡⟨⟩ + preimage ρʳ- S- x ∎ + +push-with : {A B : ℕ} → (e : Value) → Vector Value A → (Fin A → Fin B) → Vector Value B +push-with e v f = merge-with e v ∘ preimage f ∘ ⁅_⁆ + +push : {A B : ℕ} → Vector Value A → (Fin A → Fin B) → Vector Value B +push = push-with U + +mutual + merge-preimage + : {A B : ℕ} + (f : Fin A → Fin B) + → (v : Vector Value A) + (S : Subset B) + → merge v (preimage f S) ≡ merge (push v f) S + merge-preimage {zero} {zero} f v S = merge-[]₂ + merge-preimage {zero} {suc B} f v S = begin + merge v (preimage f S) ≡⟨ merge-[] v (preimage f S) ⟩ + U ≡⟨ merge-with-U U S ⟨ + merge (λ _ → U) S ≡⟨ merge-cong₁ (λ x → ≡.sym (merge-[] v (⁅ x ⁆ ∘ f))) S ⟩ + merge (push v f) S ∎ + merge-preimage {suc A} {zero} f v S with () ← f zero + merge-preimage {suc A} {suc B} f v S with insert-f0-0 f + ... | ρ , ρf0≡0 = begin + merge v (preimage f S) ≡⟨ merge-cong₂ v (preimage-cong₁ (λ x → inverseˡ ρ {f x}) S) ⟨ + merge v (preimage (ρˡ ∘ ρʳ ∘ f) S) ≡⟨⟩ + merge v (preimage (ρʳ ∘ f) (preimage ρˡ S)) ≡⟨ merge-preimage-f0≡0 (ρʳ ∘ f) ρf0≡0 v (preimage ρˡ S) ⟩ + merge (merge v ∘ preimage (ρʳ ∘ f) ∘ ⁅_⁆) (preimage ρˡ S) ≡⟨ merge-preimage-ρ (flip ρ) (merge v ∘ preimage (ρʳ ∘ f) ∘ ⁅_⁆) S ⟩ + merge (merge v ∘ preimage (ρʳ ∘ f) ∘ ⁅_⁆ ∘ ρʳ) S ≡⟨ merge-cong₁ (merge-cong₂ v ∘ preimage-cong₂ (ρʳ ∘ f) ∘ ⁅⁆∘ρ ρ) S ⟩ + merge (merge v ∘ preimage (ρʳ ∘ f) ∘ preimage ρˡ ∘ ⁅_⁆) S ≡⟨⟩ + merge (merge v ∘ preimage (ρˡ ∘ ρʳ ∘ f) ∘ ⁅_⁆) S ≡⟨ merge-cong₁ (merge-cong₂ v ∘ preimage-cong₁ (λ y → inverseˡ ρ {f y}) ∘ ⁅_⁆) S ⟩ + merge (merge v ∘ preimage f ∘ ⁅_⁆) S ∎ + where + ρʳ ρˡ : Fin (ℕ.suc B) → Fin (ℕ.suc B) + ρʳ = ρ ⟨$⟩ʳ_ + ρˡ = ρ ⟨$⟩ˡ_ + + merge-preimage-f0≡0 + : {A B : ℕ} + (f : Fin (ℕ.suc A) → Fin (ℕ.suc B)) + → f Fin.zero ≡ Fin.zero + → (v : Vector Value (ℕ.suc A)) + (S : Subset (ℕ.suc B)) + → merge v (preimage f S) ≡ merge (merge v ∘ preimage f ∘ ⁅_⁆) S + merge-preimage-f0≡0 {A} {B} f f0≡0 v S + using S0 , S- ← head S , tail S + using v0 , v- ← head v , tail v + using _ , f- ← head f , tail f + = begin + merge v f⁻¹[S] ≡⟨ merge-suc v f⁻¹[S] ⟩ + merge-with v0? v- f⁻¹[S]- ≡⟨ join-merge v0? v- f⁻¹[S]- ⟨ + join v0? (merge v- f⁻¹[S]-) ≡⟨ ≡.cong (join v0?) (merge-preimage f- v- S) ⟩ + join v0? (merge f[v-] S) ≡⟨ join-merge v0? f[v-] S ⟩ + merge-with v0? f[v-] S ≡⟨ merge-with-suc v0? f[v-] S ⟩ + merge-with v0?+[f[v-]0?] f[v-]- S- ≡⟨ ≡.cong (λ h → merge-with h f[v-]- S-) ≡f[v]0 ⟩ + merge-with f[v]0? f[v-]- S- ≡⟨ merge-with-cong f[v]0? ≡f[v]- S- ⟩ + merge-with f[v]0? f[v]- S- ≡⟨ merge-suc f[v] S ⟨ + merge f[v] S ∎ + where + f⁻¹[S] : Subset (suc A) + f⁻¹[S] = preimage f S + f⁻¹[S]- : Subset A + f⁻¹[S]- = tail f⁻¹[S] + f⁻¹[S]0 : Bool + f⁻¹[S]0 = head f⁻¹[S] + f[v] : Vector Value (suc B) + f[v] = push v f + f[v]- : Vector Value B + f[v]- = tail f[v] + f[v]0 : Value + f[v]0 = head f[v] + f[v-] : Vector Value (suc B) + f[v-] = push v- f- + f[v-]- : Vector Value B + f[v-]- = tail f[v-] + f[v-]0 : Value + f[v-]0 = head f[v-] + f⁻¹⁅0⁆ : Subset (suc A) + f⁻¹⁅0⁆ = preimage f ⁅ zero ⁆ + f⁻¹⁅0⁆- : Subset A + f⁻¹⁅0⁆- = tail f⁻¹⁅0⁆ + v0? v0?+[f[v-]0?] f[v]0? : Value + v0? = v0 when f⁻¹[S]0 + v0?+[f[v-]0?] = (if S0 then join v0? f[v-]0 else v0?) + f[v]0? = f[v]0 when S0 + ≡f[v]0 : v0?+[f[v-]0?] ≡ f[v]0? + ≡f[v]0 rewrite f0≡0 with S0 + ... | true = begin + join v0 (merge v- f⁻¹⁅0⁆-) ≡⟨ join-merge v0 v- (tail (preimage f ⁅ zero ⁆)) ⟩ + merge-with v0 v- f⁻¹⁅0⁆- ≡⟨ ≡.cong (λ h → merge-with (v0 when ⁅ zero ⁆ h) v- f⁻¹⁅0⁆-) f0≡0 ⟨ + merge-with v0?′ v- f⁻¹⁅0⁆- ≡⟨ merge-suc v (preimage f ⁅ zero ⁆) ⟨ + merge v f⁻¹⁅0⁆ ∎ + where + v0?′ : Value + v0?′ = v0 when head f⁻¹⁅0⁆ + ... | false = ≡.refl + ≡f[v]- : f[v-]- ≗ f[v]- + ≡f[v]- x = begin + push v- f- (suc x) ≡⟨ ≡.cong (λ h → merge-with (v0 when ⁅ suc x ⁆ h) v- (preimage f- ⁅ suc x ⁆)) f0≡0 ⟨ + push-with v0?′ v- f- (suc x) ≡⟨ merge-suc v (preimage f ⁅ suc x ⁆) ⟨ + push v f (suc x) ∎ + where + v0?′ : Value + v0?′ = v0 when head (preimage f ⁅ suc x ⁆) + +merge-++ + : {n m : ℕ} + (xs : Vector Value n) + (ys : Vector Value m) + (S₁ : Subset n) + (S₂ : Subset m) + → merge (xs ++ ys) (S₁ ++ S₂) + ≡ join (merge xs S₁) (merge ys S₂) +merge-++ {zero} {m} xs ys S₁ S₂ = begin + merge (xs ++ ys) (S₁ ++ S₂) ≡⟨ merge-cong₂ (xs ++ ys) (λ _ → ≡.refl) ⟩ + merge (xs ++ ys) S₂ ≡⟨ merge-cong₁ (λ _ → ≡.refl) S₂ ⟩ + merge ys S₂ ≡⟨ ≡.cong (λ h → join h (merge ys S₂)) (merge-[] xs S₁) ⟨ + join (merge xs S₁) (merge ys S₂) ∎ +merge-++ {suc n} {m} xs ys S₁ S₂ = begin + merge (xs ++ ys) (S₁ ++ S₂) ≡⟨ merge-suc (xs ++ ys) (S₁ ++ S₂) ⟩ + merge-with (head xs when head S₁) (tail (xs ++ ys)) (tail (S₁ ++ S₂)) ≡⟨ join-merge (head xs when head S₁) (tail (xs ++ ys)) (tail (S₁ ++ S₂)) ⟨ + join (head xs when head S₁) (merge (tail (xs ++ ys)) (tail (S₁ ++ S₂))) + ≡⟨ ≡.cong (join (head xs when head S₁)) (merge-cong₁ ([,]-map ∘ splitAt n) (tail (S₁ ++ S₂))) ⟩ + join (head xs when head S₁) (merge (tail xs ++ ys) (tail (S₁ ++ S₂))) + ≡⟨ ≡.cong (join (head xs when head S₁)) (merge-cong₂ (tail xs ++ ys) ([,]-map ∘ splitAt n)) ⟩ + join (head xs when head S₁) (merge (tail xs ++ ys) (tail S₁ ++ S₂)) ≡⟨ ≡.cong (join (head xs when head S₁)) (merge-++ (tail xs) ys (tail S₁) S₂) ⟩ + join (head xs when head S₁) (join (merge (tail xs) (tail S₁)) (merge ys S₂)) ≡⟨ join-assoc (head xs when head S₁) (merge (tail xs) (tail S₁)) _ ⟨ + join (join (head xs when head S₁) (merge (tail xs) (tail S₁))) (merge ys S₂) + ≡⟨ ≡.cong (λ h → join h (merge ys S₂)) (join-merge (head xs when head S₁) (tail xs) (tail S₁)) ⟩ + join (merge-with (head xs when head S₁) (tail xs) (tail S₁)) (merge ys S₂) ≡⟨ ≡.cong (λ h → join h (merge ys S₂)) (merge-suc xs S₁) ⟨ + join (merge xs S₁) (merge ys S₂) ∎ + +open import Function using (Equivalence) +open Equivalence +open import Data.Nat using (_+_) +open import Data.Fin using (_↑ˡ_; _↑ʳ_; _≟_) +open import Data.Fin.Properties using (↑ˡ-injective; ↑ʳ-injective; splitAt⁻¹-↑ˡ; splitAt-↑ˡ; splitAt⁻¹-↑ʳ; splitAt-↑ʳ) +open import Relation.Nullary.Decidable using (does; does-⇔; dec-false) + +open Fin +⁅⁆-≟ : {n : ℕ} (x y : Fin n) → ⁅ x ⁆ y ≡ does (x ≟ y) +⁅⁆-≟ zero zero = ≡.refl +⁅⁆-≟ zero (suc y) = ≡.refl +⁅⁆-≟ (suc x) zero = ≡.refl +⁅⁆-≟ (suc x) (suc y) = ⁅⁆-≟ x y + +open import Data.Sum using ([_,_]′; inj₁; inj₂) +⁅⁆-++ + : {n′ m′ : ℕ} + (i : Fin (n′ + m′)) + → [ (λ x → ⁅ x ⁆ ++ ⊥) , (λ x → ⊥ ++ ⁅ x ⁆) ]′ (splitAt n′ i) ≗ ⁅ i ⁆ +⁅⁆-++ {n′} {m′} i x with splitAt n′ i in eq₁ +... | inj₁ i′ with splitAt n′ x in eq₂ +... | inj₁ x′ = begin + ⁅ i′ ⁆ x′ ≡⟨ ⁅⁆-≟ i′ x′ ⟩ + does (i′ ≟ x′) ≡⟨ does-⇔ ⇔ (i′ ≟ x′) (i′ ↑ˡ m′ ≟ x′ ↑ˡ m′) ⟩ + does (i′ ↑ˡ m′ ≟ x′ ↑ˡ m′) ≡⟨ ⁅⁆-≟ (i′ ↑ˡ m′) (x′ ↑ˡ m′) ⟨ + ⁅ i′ ↑ˡ m′ ⁆ (x′ ↑ˡ m′) ≡⟨ ≡.cong₂ ⁅_⁆ (splitAt⁻¹-↑ˡ eq₁) (splitAt⁻¹-↑ˡ eq₂) ⟩ + ⁅ i ⁆ x ∎ + where + ⇔ : Equivalence (≡.setoid (i′ ≡ x′)) (≡.setoid (i′ ↑ˡ m′ ≡ x′ ↑ˡ m′)) + ⇔ .to = ≡.cong (_↑ˡ m′) + ⇔ .from = ↑ˡ-injective m′ i′ x′ + ⇔ .to-cong ≡.refl = ≡.refl + ⇔ .from-cong ≡.refl = ≡.refl +... | inj₂ x′ = begin + false ≡⟨ dec-false (i′ ↑ˡ m′ ≟ n′ ↑ʳ x′) ↑ˡ≢↑ʳ ⟨ + does (i′ ↑ˡ m′ ≟ n′ ↑ʳ x′) ≡⟨ ⁅⁆-≟ (i′ ↑ˡ m′) (n′ ↑ʳ x′) ⟨ + ⁅ i′ ↑ˡ m′ ⁆ (n′ ↑ʳ x′) ≡⟨ ≡.cong₂ ⁅_⁆ (splitAt⁻¹-↑ˡ eq₁) (splitAt⁻¹-↑ʳ eq₂) ⟩ + ⁅ i ⁆ x ∎ + where + ↑ˡ≢↑ʳ : i′ ↑ˡ m′ ≢ n′ ↑ʳ x′ + ↑ˡ≢↑ʳ ≡ = case ≡.trans (≡.sym (splitAt-↑ˡ n′ i′ m′)) (≡.trans (≡.cong (splitAt n′) ≡) (splitAt-↑ʳ n′ m′ x′)) of λ { () } +⁅⁆-++ {n′} i x | inj₂ i′ with splitAt n′ x in eq₂ +⁅⁆-++ {n′} {m′} i x | inj₂ i′ | inj₁ x′ = begin + [ ⊥ , ⁅ i′ ⁆ ]′ (splitAt n′ x) ≡⟨ ≡.cong ([ ⊥ , ⁅ i′ ⁆ ]′) eq₂ ⟩ + false ≡⟨ dec-false (n′ ↑ʳ i′ ≟ x′ ↑ˡ m′) ↑ʳ≢↑ˡ ⟨ + does (n′ ↑ʳ i′ ≟ x′ ↑ˡ m′) ≡⟨ ⁅⁆-≟ (n′ ↑ʳ i′) (x′ ↑ˡ m′) ⟨ + ⁅ n′ ↑ʳ i′ ⁆ (x′ ↑ˡ m′) ≡⟨ ≡.cong₂ ⁅_⁆ (splitAt⁻¹-↑ʳ eq₁) (splitAt⁻¹-↑ˡ eq₂) ⟩ + ⁅ i ⁆ x ∎ + where + ↑ʳ≢↑ˡ : n′ ↑ʳ i′ ≢ x′ ↑ˡ m′ + ↑ʳ≢↑ˡ ≡ = case ≡.trans (≡.sym (splitAt-↑ʳ n′ m′ i′)) (≡.trans (≡.cong (splitAt n′) ≡) (splitAt-↑ˡ n′ x′ m′)) of λ { () } +⁅⁆-++ {n′} i x | inj₂ i′ | inj₂ x′ = begin + [ ⊥ , ⁅ i′ ⁆ ]′ (splitAt n′ x) ≡⟨ ≡.cong [ ⊥ , ⁅ i′ ⁆ ]′ eq₂ ⟩ + ⁅ i′ ⁆ x′ ≡⟨ ⁅⁆-≟ i′ x′ ⟩ + does (i′ ≟ x′) ≡⟨ does-⇔ ⇔ (i′ ≟ x′) (n′ ↑ʳ i′ ≟ n′ ↑ʳ x′) ⟩ + does (n′ ↑ʳ i′ ≟ n′ ↑ʳ x′) ≡⟨ ⁅⁆-≟ (n′ ↑ʳ i′) (n′ ↑ʳ x′) ⟨ + ⁅ n′ ↑ʳ i′ ⁆ (n′ ↑ʳ x′) ≡⟨ ≡.cong₂ ⁅_⁆ (splitAt⁻¹-↑ʳ eq₁) (splitAt⁻¹-↑ʳ eq₂) ⟩ + ⁅ i ⁆ x ∎ + where + ⇔ : Equivalence (≡.setoid (i′ ≡ x′)) (≡.setoid (n′ ↑ʳ i′ ≡ n′ ↑ʳ x′)) + ⇔ .to = ≡.cong (n′ ↑ʳ_) + ⇔ .from = ↑ʳ-injective n′ i′ x′ + ⇔ .to-cong ≡.refl = ≡.refl + ⇔ .from-cong ≡.refl = ≡.refl diff --git a/Data/Circuit/Typecheck.agda b/Data/Circuit/Typecheck.agda new file mode 100644 index 0000000..e34ea44 --- /dev/null +++ b/Data/Circuit/Typecheck.agda @@ -0,0 +1,78 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Circuit.Typecheck where + +open import Data.SExp using (SExp) +open import Data.Circuit.Gate using (GateLabel; Gate) +open import Data.Hypergraph.Label using (HypergraphLabel) +open import Data.Hypergraph.Edge GateLabel using (Edge) +open import Data.Hypergraph.Base GateLabel using (Hypergraph) + +open import Data.List using (List; length) renaming (map to mapL) +open import Data.List.Effectful using () renaming (module TraversableA to ListTraversable) +open import Data.Maybe using (Maybe) renaming (map to mapM) +open import Data.Nat using (ℕ; _<?_; _≟_) +open import Data.String using (String) +open import Data.Product using (_×_; _,_; Σ) +open import Data.Vec using (Vec; []; _∷_; fromList) renaming (map to mapV) +open import Data.Vec.Effectful using () renaming (module TraversableA to VecTraversable) +open import Data.Maybe.Effectful using (applicative) +open import Data.Fin using (Fin; #_; fromℕ<) +open import Level using (0ℓ) + +import Relation.Binary.PropositionalEquality as ≡ + +open List +open SExp +open Gate +open Maybe + +gate : {n a : ℕ} (g : Gate a) → Vec (Fin n) a → Edge n +gate g p = record { label = g; ports = p } + +typeCheckGateLabel : SExp → Maybe (Σ ℕ Gate) +typeCheckGateLabel (Atom "one") = just (1 , ONE) +typeCheckGateLabel (Atom "zero") = just (1 , ZERO) +typeCheckGateLabel (Atom "not") = just (2 , NOT) +typeCheckGateLabel (Atom "id") = just (2 , ID) +typeCheckGateLabel (Atom "and") = just (3 , AND) +typeCheckGateLabel (Atom "or") = just (3 , OR) +typeCheckGateLabel (Atom "xor") = just (3 , XOR) +typeCheckGateLabel (Atom "nand") = just (3 , NAND) +typeCheckGateLabel (Atom "nor") = just (3 , NOR) +typeCheckGateLabel (Atom "xnor") = just (3 , XNOR) +typeCheckGateLabel _ = nothing + +open import Relation.Nullary.Decidable using (Dec; yes; no) +open Dec +open VecTraversable {0ℓ} applicative using () renaming (sequenceA to vecSequenceA) +open ListTraversable {0ℓ} applicative using () renaming (sequenceA to listSequenceA) + +typeCheckPort : (v : ℕ) → SExp → Maybe (Fin v) +typeCheckPort v (Nat n) with n <? v +... | yes n<v = just (fromℕ< n<v) +... | no _ = nothing +typeCheckPort _ _ = nothing + +typeCheckPorts : (v n : ℕ) → List SExp → Maybe (Vec (Fin v) n) +typeCheckPorts v n xs with length xs ≟ n +... | yes ≡.refl = vecSequenceA (mapV (typeCheckPort v) (fromList xs)) +... | no _ = nothing + +typeCheckGate : (v : ℕ) → SExp → Maybe (Edge v) +typeCheckGate v (SExps (labelString ∷ ports)) with typeCheckGateLabel labelString +... | just (n , label) = mapM (gate label) (typeCheckPorts v n ports) +... | nothing = nothing +typeCheckGate v _ = nothing + +typeCheckHeader : SExp → Maybe ℕ +typeCheckHeader (SExps (Atom "hypergraph" ∷ Nat n ∷ [])) = just n +typeCheckHeader _ = nothing + +typeCheckHypergraph : SExp → Maybe (Σ ℕ Hypergraph) +typeCheckHypergraph (SExps (x ∷ xs)) with typeCheckHeader x +... | nothing = nothing +... | just n with listSequenceA (mapL (typeCheckGate n) xs) +... | just e = just (n , record { edges = e }) +... | nothing = nothing +typeCheckHypergraph _ = nothing diff --git a/Data/Circuit/Value.agda b/Data/Circuit/Value.agda new file mode 100644 index 0000000..b135c35 --- /dev/null +++ b/Data/Circuit/Value.agda @@ -0,0 +1,180 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Circuit.Value where + +import Relation.Binary.Lattice.Properties.BoundedJoinSemilattice as LatticeProp + +open import Algebra.Bundles using (CommutativeMonoid) +open import Algebra.Structures using (IsCommutativeMonoid; IsMonoid; IsSemigroup; IsMagma) +open import Data.Product.Base using (_×_; _,_) +open import Data.String.Base using (String) +open import Level using (0ℓ) +open import Relation.Binary.Lattice.Bundles using (BoundedJoinSemilattice) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) + +open CommutativeMonoid +open IsCommutativeMonoid +open IsMagma +open IsMonoid +open IsSemigroup + +data Value : Set where + U T F X : Value + +data ≤-Value : Value → Value → Set where + v≤v : {v : Value} → ≤-Value v v + U≤T : ≤-Value U T + U≤F : ≤-Value U F + U≤X : ≤-Value U X + T≤X : ≤-Value T X + F≤X : ≤-Value F X + +≤-reflexive : {x y : Value} → x ≡ y → ≤-Value x y +≤-reflexive ≡.refl = v≤v + +≤-transitive : {i j k : Value} → ≤-Value i j → ≤-Value j k → ≤-Value i k +≤-transitive v≤v y = y +≤-transitive x v≤v = x +≤-transitive U≤T T≤X = U≤X +≤-transitive U≤F F≤X = U≤X + +≤-antisymmetric : {i j : Value} → ≤-Value i j → ≤-Value j i → i ≡ j +≤-antisymmetric v≤v _ = ≡.refl + +showValue : Value → String +showValue U = "U" +showValue T = "T" +showValue F = "F" +showValue X = "X" + +join : Value → Value → Value +join U y = y +join x U = x +join T T = T +join T F = X +join F T = X +join F F = F +join X _ = X +join _ X = X + +≤-supremum + : (x y : Value) + → ≤-Value x (join x y) + × ≤-Value y (join x y) + × ((z : Value) → ≤-Value x z → ≤-Value y z → ≤-Value (join x y) z) +≤-supremum U U = v≤v , v≤v , λ _ U≤z _ → U≤z +≤-supremum U T = U≤T , v≤v , λ { z x≤z y≤z → y≤z } +≤-supremum U F = U≤F , v≤v , λ { z x≤z y≤z → y≤z } +≤-supremum U X = U≤X , v≤v , λ { z x≤z y≤z → y≤z } +≤-supremum T U = v≤v , U≤T , λ { z x≤z y≤z → x≤z } +≤-supremum T T = v≤v , v≤v , λ { z x≤z y≤z → x≤z } +≤-supremum T F = T≤X , F≤X , λ { X x≤z y≤z → v≤v } +≤-supremum T X = T≤X , v≤v , λ { z x≤z y≤z → y≤z } +≤-supremum F U = v≤v , U≤F , λ { z x≤z y≤z → x≤z } +≤-supremum F T = F≤X , T≤X , λ { X x≤z y≤z → v≤v } +≤-supremum F F = v≤v , v≤v , λ { z x≤z y≤z → x≤z } +≤-supremum F X = F≤X , v≤v , λ { z x≤z y≤z → y≤z } +≤-supremum X U = v≤v , U≤X , λ { z x≤z y≤z → x≤z } +≤-supremum X T = v≤v , T≤X , λ { z x≤z y≤z → x≤z } +≤-supremum X F = v≤v , F≤X , λ { z x≤z y≤z → x≤z } +≤-supremum X X = v≤v , v≤v , λ { z x≤z y≤z → x≤z } + +join-comm : (x y : Value) → join x y ≡ join y x +join-comm U U = ≡.refl +join-comm U T = ≡.refl +join-comm U F = ≡.refl +join-comm U X = ≡.refl +join-comm T U = ≡.refl +join-comm T T = ≡.refl +join-comm T F = ≡.refl +join-comm T X = ≡.refl +join-comm F U = ≡.refl +join-comm F T = ≡.refl +join-comm F F = ≡.refl +join-comm F X = ≡.refl +join-comm X U = ≡.refl +join-comm X T = ≡.refl +join-comm X F = ≡.refl +join-comm X X = ≡.refl + +join-assoc : (x y z : Value) → join (join x y) z ≡ join x (join y z) +join-assoc U y z = ≡.refl +join-assoc T U z = ≡.refl +join-assoc T T U = ≡.refl +join-assoc T T T = ≡.refl +join-assoc T T F = ≡.refl +join-assoc T T X = ≡.refl +join-assoc T F U = ≡.refl +join-assoc T F T = ≡.refl +join-assoc T F F = ≡.refl +join-assoc T F X = ≡.refl +join-assoc T X U = ≡.refl +join-assoc T X T = ≡.refl +join-assoc T X F = ≡.refl +join-assoc T X X = ≡.refl +join-assoc F U z = ≡.refl +join-assoc F T U = ≡.refl +join-assoc F T T = ≡.refl +join-assoc F T F = ≡.refl +join-assoc F T X = ≡.refl +join-assoc F F U = ≡.refl +join-assoc F F T = ≡.refl +join-assoc F F F = ≡.refl +join-assoc F F X = ≡.refl +join-assoc F X U = ≡.refl +join-assoc F X T = ≡.refl +join-assoc F X F = ≡.refl +join-assoc F X X = ≡.refl +join-assoc X U z = ≡.refl +join-assoc X T U = ≡.refl +join-assoc X T T = ≡.refl +join-assoc X T F = ≡.refl +join-assoc X T X = ≡.refl +join-assoc X F U = ≡.refl +join-assoc X F T = ≡.refl +join-assoc X F F = ≡.refl +join-assoc X F X = ≡.refl +join-assoc X X U = ≡.refl +join-assoc X X T = ≡.refl +join-assoc X X F = ≡.refl +join-assoc X X X = ≡.refl + +Lattice : BoundedJoinSemilattice 0ℓ 0ℓ 0ℓ +Lattice = record + { Carrier = Value + ; _≈_ = _≡_ + ; _≤_ = ≤-Value + ; _∨_ = join + ; ⊥ = U + ; isBoundedJoinSemilattice = record + { isJoinSemilattice = record + { isPartialOrder = record + { isPreorder = record + { isEquivalence = ≡.isEquivalence + ; reflexive = ≤-reflexive + ; trans = ≤-transitive + } + ; antisym = ≤-antisymmetric + } + ; supremum = ≤-supremum + } + ; minimum = λ where + U → v≤v + T → U≤T + F → U≤F + X → U≤X + } + } + +module Lattice = BoundedJoinSemilattice Lattice + +Monoid : CommutativeMonoid 0ℓ 0ℓ +Monoid .Carrier = Lattice.Carrier +Monoid ._≈_ = Lattice._≈_ +Monoid ._∙_ = Lattice._∨_ +Monoid .ε = Lattice.⊥ +Monoid .isCommutativeMonoid .isMonoid .isSemigroup .isMagma .isEquivalence = ≡.isEquivalence +Monoid .isCommutativeMonoid .isMonoid .isSemigroup .isMagma .∙-cong = ≡.cong₂ join +Monoid .isCommutativeMonoid .isMonoid .isSemigroup .assoc = join-assoc +Monoid .isCommutativeMonoid .isMonoid .identity = LatticeProp.identity Lattice +Monoid .isCommutativeMonoid .comm = join-comm diff --git a/Data/Fin/Preimage.agda b/Data/Fin/Preimage.agda new file mode 100644 index 0000000..4f7ea43 --- /dev/null +++ b/Data/Fin/Preimage.agda @@ -0,0 +1,48 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Fin.Preimage where + +open import Data.Nat.Base using (ℕ) +open import Data.Fin.Base using (Fin) +open import Function.Base using (∣_⟩-_; _∘_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) +open import Data.Subset.Functional using (Subset; foldl; _∪_; ⊥; ⁅_⁆) + +private + variable A B C : ℕ + +preimage : (Fin A → Fin B) → Subset B → Subset A +preimage f Y = Y ∘ f + +⋃ : Subset A → (Fin A → Subset B) → Subset B +⋃ S f = foldl (∣ _∪_ ⟩- f) ⊥ S + +image : (Fin A → Fin B) → Subset A → Subset B +image f X = ⋃ X (⁅_⁆ ∘ f) + +preimage-cong₁ + : {f g : Fin A → Fin B} + → f ≗ g + → (S : Subset B) + → preimage f S ≗ preimage g S +preimage-cong₁ f≗g S x = ≡.cong S (f≗g x) + +preimage-cong₂ + : (f : Fin A → Fin B) + {S₁ S₂ : Subset B} + → S₁ ≗ S₂ + → preimage f S₁ ≗ preimage f S₂ +preimage-cong₂ f S₁≗S₂ = S₁≗S₂ ∘ f + +preimage-∘ + : (f : Fin A → Fin B) + (g : Fin B → Fin C) + (z : Fin C) + → preimage (g ∘ f) ⁅ z ⁆ ≗ preimage f (preimage g ⁅ z ⁆) +preimage-∘ f g S x = ≡.refl + +preimage-⊥ + : {n m : ℕ} + (f : Fin n → Fin m) + → preimage f ⊥ ≗ ⊥ +preimage-⊥ f x = ≡.refl diff --git a/Data/Hypergraph.agda b/Data/Hypergraph.agda new file mode 100644 index 0000000..532e082 --- /dev/null +++ b/Data/Hypergraph.agda @@ -0,0 +1,86 @@ +{-# OPTIONS --without-K --safe #-} + +open import Data.Hypergraph.Label using (HypergraphLabel) +open import Level using (Level; 0ℓ) + +module Data.Hypergraph {ℓ : Level} (HL : HypergraphLabel) where + +import Data.List.Relation.Binary.Pointwise as PW +import Function.Reasoning as →-Reasoning +import Relation.Binary.PropositionalEquality as ≡ +import Data.Hypergraph.Edge {ℓ} HL as Hyperedge +import Data.List.Relation.Binary.Permutation.Propositional as List-↭ +import Data.List.Relation.Binary.Permutation.Setoid as ↭ + +open HypergraphLabel HL using (Label) public + +open import Data.List using (List; map) +open import Data.Nat using (ℕ) +open import Data.String using (String; unlines) +open import Function using (_∘_; _⟶ₛ_; Func) +open import Relation.Binary using (Setoid) + +module Edge = Hyperedge +open Edge using (Edge; Edgeₛ) + +-- A hypergraph is a list of edges +record Hypergraph (v : ℕ) : Set ℓ where + constructor mkHypergraph + field + edges : List (Edge v) + +module _ {v : ℕ} where + + open ↭ (Edgeₛ v) using (_↭_; ↭-refl; ↭-sym; ↭-trans) + + show : Hypergraph v → String + show (mkHypergraph e) = unlines (map Edge.show e) + + -- an equivalence relation on hypergraphs + record _≈_ (H H′ : Hypergraph v) : Set ℓ where + + constructor mk≈ + + module H = Hypergraph H + module H′ = Hypergraph H′ + + field + ↭-edges : H.edges ↭ H′.edges + + infixr 4 _≈_ + + ≈-refl : {H : Hypergraph v} → H ≈ H + ≈-refl = mk≈ ↭-refl + + ≈-sym : {H H′ : Hypergraph v} → H ≈ H′ → H′ ≈ H + ≈-sym (mk≈ ≡n) = mk≈ (↭-sym ≡n) + + ≈-trans : {H H′ H″ : Hypergraph v} → H ≈ H′ → H′ ≈ H″ → H ≈ H″ + ≈-trans (mk≈ ≡n₁) (mk≈ ≡n₂) = mk≈ (↭-trans ≡n₁ ≡n₂) + +-- The setoid of labeled hypergraphs with v nodes +Hypergraphₛ : ℕ → Setoid ℓ ℓ +Hypergraphₛ v = record + { Carrier = Hypergraph v + ; _≈_ = _≈_ + ; isEquivalence = record + { refl = ≈-refl + ; sym = ≈-sym + ; trans = ≈-trans + } + } + +open Func + +open ↭ using (↭-setoid) + +Multiset∘Edgeₛ : (n : ℕ) → Setoid ℓ ℓ +Multiset∘Edgeₛ = ↭-setoid ∘ Edgeₛ + +mkHypergraphₛ : {n : ℕ} → Multiset∘Edgeₛ n ⟶ₛ Hypergraphₛ n +mkHypergraphₛ .to = mkHypergraph +mkHypergraphₛ .cong = mk≈ + +edgesₛ : {n : ℕ} → Hypergraphₛ n ⟶ₛ Multiset∘Edgeₛ n +edgesₛ .to = Hypergraph.edges +edgesₛ .cong = _≈_.↭-edges diff --git a/Data/Hypergraph/Edge.agda b/Data/Hypergraph/Edge.agda new file mode 100644 index 0000000..447f008 --- /dev/null +++ b/Data/Hypergraph/Edge.agda @@ -0,0 +1,113 @@ +{-# OPTIONS --without-K --safe #-} + +open import Data.Hypergraph.Label using (HypergraphLabel) + +open import Level using (Level; 0ℓ) +module Data.Hypergraph.Edge {ℓ : Level} (HL : HypergraphLabel) where + +import Data.Vec.Functional as Vec +import Data.Vec.Functional.Relation.Binary.Equality.Setoid as PW +import Data.Fin.Properties as FinProp + +open import Data.Fin as Fin using (Fin) +open import Data.Fin.Show using () renaming (show to showFin) +open import Data.Nat using (ℕ) +open import Data.String using (String; _<+>_) +open import Data.Vec.Show using () renaming (show to showVec) +open import Level using (0ℓ) +open import Relation.Binary using (Setoid; IsEquivalence) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; module ≡-Reasoning) +open import Function using (_⟶ₛ_; Func; _∘_) + +module HL = HypergraphLabel HL + +open HL using (Label) +open Vec using (Vector) +open Func + +record Edge (v : ℕ) : Set ℓ where + constructor mkEdge + field + {arity} : ℕ + label : Label arity + ports : Fin arity → Fin v + +map : {n m : ℕ} → (Fin n → Fin m) → Edge n → Edge m +map f edge = record + { label = label + ; ports = Vec.map f ports + } + where + open Edge edge + +module _ {v : ℕ} where + + open PW (≡.setoid (Fin v)) using (_≋_) + + -- an equivalence relation on edges with v nodes + record _≈_ (E E′ : Edge v) : Set ℓ where + constructor mk≈ + module E = Edge E + module E′ = Edge E′ + field + ≡arity : E.arity ≡ E′.arity + ≡label : HL.cast ≡arity E.label ≡ E′.label + ≡ports : E.ports ≋ E′.ports ∘ Fin.cast ≡arity + + ≈-refl : {x : Edge v} → x ≈ x + ≈-refl {x} = record + { ≡arity = ≡.refl + ; ≡label = HL.≈-reflexive ≡.refl + ; ≡ports = ≡.cong (Edge.ports x) ∘ ≡.sym ∘ FinProp.cast-is-id _ + } + + ≈-sym : {x y : Edge v} → x ≈ y → y ≈ x + ≈-sym x≈y = record + { ≡arity = ≡.sym ≡arity + ; ≡label = HL.≈-sym ≡label + ; ≡ports = ≡.sym ∘ ≡ports-sym + } + where + open _≈_ x≈y + open ≡-Reasoning + ≡ports-sym : (i : Fin E′.arity) → E.ports (Fin.cast _ i) ≡ E′.ports i + ≡ports-sym i = begin + E.ports (Fin.cast _ i) ≡⟨ ≡ports (Fin.cast _ i) ⟩ + E′.ports (Fin.cast ≡arity (Fin.cast _ i)) + ≡⟨ ≡.cong E′.ports (FinProp.cast-involutive ≡arity _ i) ⟩ + E′.ports i ∎ + + ≈-trans : {x y z : Edge v} → x ≈ y → y ≈ z → x ≈ z + ≈-trans {x} {y} {z} x≈y y≈z = record + { ≡arity = ≡.trans x≈y.≡arity y≈z.≡arity + ; ≡label = HL.≈-trans x≈y.≡label y≈z.≡label + ; ≡ports = ≡-ports + } + where + module x≈y = _≈_ x≈y + module y≈z = _≈_ y≈z + open ≡-Reasoning + ≡-ports : (i : Fin x≈y.E.arity) → x≈y.E.ports i ≡ y≈z.E′.ports (Fin.cast _ i) + ≡-ports i = begin + x≈y.E.ports i ≡⟨ x≈y.≡ports i ⟩ + y≈z.E.ports (Fin.cast _ i) ≡⟨ y≈z.≡ports (Fin.cast _ i) ⟩ + y≈z.E′.ports (Fin.cast y≈z.≡arity (Fin.cast _ i)) + ≡⟨ ≡.cong y≈z.E′.ports (FinProp.cast-trans _ y≈z.≡arity i) ⟩ + y≈z.E′.ports (Fin.cast _ i) ∎ + + ≈-IsEquivalence : IsEquivalence _≈_ + ≈-IsEquivalence = record + { refl = ≈-refl + ; sym = ≈-sym + ; trans = ≈-trans + } + + show : Edge v → String + show (mkEdge {a} l p) = HL.showLabel a l <+> showVec showFin (Vec.toVec p) + +Edgeₛ : (v : ℕ) → Setoid ℓ ℓ +Edgeₛ v = record { isEquivalence = ≈-IsEquivalence {v} } + +mapₛ : {n m : ℕ} → (Fin n → Fin m) → Edgeₛ n ⟶ₛ Edgeₛ m +mapₛ f .to = map f +mapₛ f .cong (mk≈ ≡a ≡l ≡p) = mk≈ ≡a ≡l (≡.cong f ∘ ≡p) diff --git a/Data/Hypergraph/Edge/Order.agda b/Data/Hypergraph/Edge/Order.agda new file mode 100644 index 0000000..4b3c1e8 --- /dev/null +++ b/Data/Hypergraph/Edge/Order.agda @@ -0,0 +1,280 @@ +{-# OPTIONS --without-K --safe #-} + +open import Data.Hypergraph.Label using (HypergraphLabel) + +module Data.Hypergraph.Edge.Order (HL : HypergraphLabel) where + +import Data.List.Sort as ListSort +import Data.Fin as Fin +import Data.Fin.Properties as FinProp +import Data.Vec as Vec +import Data.Vec.Relation.Binary.Equality.Cast as VecCast +import Data.Vec.Relation.Binary.Lex.Strict as Lex +import Relation.Binary.PropositionalEquality as ≡ +import Relation.Binary.Properties.DecTotalOrder as DTOP +import Relation.Binary.Properties.StrictTotalOrder as STOP + +open import Data.Hypergraph.Edge HL using (Edge; ≈-Edge; ≈-Edge-IsEquivalence) +open import Data.Fin using (Fin) +open import Data.Nat using (ℕ; _<_) +open import Data.Nat.Properties using (<-irrefl; <-trans; <-resp₂-≡; <-cmp) +open import Data.Product.Base using (_,_; proj₁; proj₂) +open import Data.Vec.Relation.Binary.Pointwise.Inductive using (Pointwise-≡⇒≡) +open import Level using (0ℓ) +open import Relation.Binary + using + ( Rel + ; Tri ; Trichotomous + ; IsStrictTotalOrder ; IsEquivalence + ; _Respects_ + ; DecTotalOrder ; StrictTotalOrder + ) +open import Relation.Nullary using (¬_) + +module HL = HypergraphLabel HL +open HL using (Label; cast; cast-is-id) +open Vec using (Vec) + +open ≡ using (_≡_) + +open HL using (_[_<_]) +_<<_ : {v a : ℕ} → Rel (Vec (Fin v) a) 0ℓ +_<<_ {v} = Lex.Lex-< _≡_ (Fin._<_ {v}) + +data <-Edge {v : ℕ} : Edge v → Edge v → Set where + <-arity + : {x y : Edge v} + → Edge.arity x < Edge.arity y + → <-Edge x y + <-label + : {x y : Edge v} + (≡a : Edge.arity x ≡ Edge.arity y) + → Edge.arity y [ cast ≡a (Edge.label x) < Edge.label y ] + → <-Edge x y + <-ports + : {x y : Edge v} + (≡a : Edge.arity x ≡ Edge.arity y) + (≡l : Edge.label x HL.≈[ ≡a ] Edge.label y) + → Vec.cast ≡a (Edge.ports x) << Edge.ports y + → <-Edge x y + +<-Edge-irrefl : {v : ℕ} {x y : Edge v} → ≈-Edge x y → ¬ <-Edge x y +<-Edge-irrefl record { ≡arity = ≡a } (<-arity n<m) = <-irrefl ≡a n<m +<-Edge-irrefl record { ≡label = ≡l } (<-label _ (_ , x≉y)) = x≉y ≡l +<-Edge-irrefl record { ≡ports = ≡p } (<-ports ≡.refl ≡l x<y) + = Lex.<-irrefl FinProp.<-irrefl (≡⇒Pointwise-≡ ≡p) x<y + +<-Edge-trans : {v : ℕ} {i j k : Edge v} → <-Edge i j → <-Edge j k → <-Edge i k +<-Edge-trans (<-arity i<j) (<-arity j<k) = <-arity (<-trans i<j j<k) +<-Edge-trans (<-arity i<j) (<-label ≡.refl j<k) = <-arity i<j +<-Edge-trans (<-arity i<j) (<-ports ≡.refl _ j<k) = <-arity i<j +<-Edge-trans (<-label ≡.refl i<j) (<-arity j<k) = <-arity j<k +<-Edge-trans {_} {i} (<-label ≡.refl i<j) (<-label ≡.refl j<k) + = <-label ≡.refl (<-label-trans i<j (<-respˡ-≈ (HL.≈-reflexive ≡.refl) j<k)) + where + open DTOP (HL.decTotalOrder (Edge.arity i)) using (<-respˡ-≈) renaming (<-trans to <-label-trans) +<-Edge-trans {k = k} (<-label ≡.refl i<j) (<-ports ≡.refl ≡.refl _) + = <-label ≡.refl (<-respʳ-≈ (≡.sym (HL.≈-reflexive ≡.refl)) i<j) + where + open DTOP (HL.decTotalOrder (Edge.arity k)) using (<-respʳ-≈) +<-Edge-trans (<-ports ≡.refl _ _) (<-arity j<k) = <-arity j<k +<-Edge-trans {k = k} (<-ports ≡.refl ≡.refl _) (<-label ≡.refl j<k) + = <-label ≡.refl (<-respˡ-≈ (≡.cong (cast _) (HL.≈-reflexive ≡.refl)) j<k) + where + open DTOP (HL.decTotalOrder (Edge.arity k)) using (<-respˡ-≈) +<-Edge-trans {j = j} (<-ports ≡.refl ≡l₁ i<j) (<-ports ≡.refl ≡l₂ j<k) + rewrite (VecCast.cast-is-id ≡.refl (Edge.ports j)) + = <-ports ≡.refl + (HL.≈-trans ≡l₁ ≡l₂) + (Lex.<-trans ≡-isPartialEquivalence FinProp.<-resp₂-≡ FinProp.<-trans i<j j<k) + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + +<-Edge-respˡ-≈ : {v : ℕ} {y : Edge v} → (λ x → <-Edge x y) Respects ≈-Edge +<-Edge-respˡ-≈ ≈x (<-arity x₁<y) = <-arity (proj₂ <-resp₂-≡ ≡arity x₁<y) + where + open ≈-Edge ≈x using (≡arity) +<-Edge-respˡ-≈ {_} {y} record { ≡arity = ≡.refl ; ≡label = ≡.refl } (<-label ≡.refl x₁<y) + = <-label ≡.refl (<-respˡ-≈ (≡.sym (HL.≈-reflexive ≡.refl)) x₁<y) + where + module y = Edge y + open DTOP (HL.decTotalOrder y.arity) using (<-respˡ-≈) +<-Edge-respˡ-≈ record { ≡arity = ≡.refl ; ≡label = ≡.refl; ≡ports = ≡.refl} (<-ports ≡.refl ≡.refl x₁<y) + = <-ports + ≡.refl + (≡.cong (cast _) (HL.≈-reflexive ≡.refl)) + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (≡.sym (VecCast.≈-reflexive ≡.refl))) + x₁<y) + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + +<-Edge-respʳ-≈ : {v : ℕ} {x : Edge v} → <-Edge x Respects ≈-Edge +<-Edge-respʳ-≈ record { ≡arity = ≡a } (<-arity x<y₁) = <-arity (proj₁ <-resp₂-≡ ≡a x<y₁) +<-Edge-respʳ-≈ {_} {x} record { ≡arity = ≡.refl ; ≡label = ≡.refl } (<-label ≡.refl x<y₁) + = <-label ≡.refl (<-respʳ-≈ (≡.sym (HL.≈-reflexive ≡.refl)) x<y₁) + where + module x = Edge x + open DTOP (HL.decTotalOrder x.arity) using (<-respʳ-≈) +<-Edge-respʳ-≈ record { ≡arity = ≡.refl ; ≡label = ≡.refl; ≡ports = ≡.refl} (<-ports ≡.refl ≡.refl x<y₁) + = <-ports + ≡.refl + (≡.cong (cast _) (≡.sym (HL.≈-reflexive ≡.refl))) + (Lex.<-respectsʳ + ≡-isPartialEquivalence + FinProp.<-respʳ-≡ + (≡⇒Pointwise-≡ (≡.sym (VecCast.≈-reflexive ≡.refl))) + x<y₁) + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + +open Tri +open ≈-Edge +tri : {v : ℕ} → Trichotomous (≈-Edge {v}) (<-Edge {v}) +tri x y with <-cmp x.arity y.arity + where + module x = Edge x + module y = Edge y +tri x y | tri< x<y x≢y y≮x = tri< (<-arity x<y) (λ x≡y → x≢y (≡arity x≡y)) ¬y<x + where + ¬y<x : ¬ <-Edge y x + ¬y<x (<-arity y<x) = y≮x y<x + ¬y<x (<-label ≡a _) = x≢y (≡.sym ≡a) + ¬y<x (<-ports ≡a _ _) = x≢y (≡.sym ≡a) +tri x y | tri≈ x≮y ≡.refl y≮x = compare-label + where + module x = Edge x + module y = Edge y + open StrictTotalOrder (HL.strictTotalOrder x.arity) using (compare) + import Relation.Binary.Properties.DecTotalOrder + open DTOP (HL.decTotalOrder x.arity) using (<-respˡ-≈) + compare-label : Tri (<-Edge x y) (≈-Edge x y) (<-Edge y x) + compare-label with compare x.label y.label + ... | tri< x<y x≢y y≮x′ = tri< + (<-label ≡.refl (<-respˡ-≈ (≡.sym (HL.≈-reflexive ≡.refl)) x<y)) + (λ x≡y → x≢y (≡.trans (≡.sym (HL.≈-reflexive ≡.refl)) (≡label x≡y))) + ¬y<x + where + ¬y<x : ¬ <-Edge y x + ¬y<x (<-arity y<x) = y≮x y<x + ¬y<x (<-label _ y<x) = y≮x′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) y<x) + ¬y<x (<-ports _ ≡l _) = x≢y (≡.trans (≡.sym ≡l) (cast-is-id ≡.refl y.label)) + ... | tri≈ x≮y′ x≡y′ y≮x′ = compare-ports + where + compare-ports : Tri (<-Edge x y) (≈-Edge x y) (<-Edge y x) + compare-ports with Lex.<-cmp ≡.sym FinProp.<-cmp x.ports y.ports + ... | tri< x<y x≢y y≮x″ = + tri< + (<-ports ≡.refl + (HL.≈-reflexive x≡y′) + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (≡.sym (VecCast.≈-reflexive ≡.refl))) + x<y)) + (λ x≡y → x≢y (≡⇒Pointwise-≡ (≡.trans (≡.sym (VecCast.≈-reflexive ≡.refl)) (≡ports x≡y)))) + ¬y<x + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + ¬y<x : ¬ <-Edge y x + ¬y<x (<-arity y<x) = y≮x y<x + ¬y<x (<-label _ y<x) = y≮x′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) y<x) + ¬y<x (<-ports _ _ y<x) = + y≮x″ + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (VecCast.≈-reflexive ≡.refl)) + y<x) + ... | tri≈ x≮y″ x≡y″ y≮x″ = tri≈ + ¬x<y + (record { ≡arity = ≡.refl ; ≡label = HL.≈-reflexive x≡y′ ; ≡ports = VecCast.≈-reflexive (Pointwise-≡⇒≡ x≡y″) }) + ¬y<x + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + ¬x<y : ¬ <-Edge x y + ¬x<y (<-arity x<y) = x≮y x<y + ¬x<y (<-label _ x<y) = x≮y′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) x<y) + ¬x<y (<-ports _ _ x<y) = + x≮y″ + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (VecCast.≈-reflexive ≡.refl)) + x<y) + ¬y<x : ¬ <-Edge y x + ¬y<x (<-arity y<x) = y≮x y<x + ¬y<x (<-label _ y<x) = y≮x′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) y<x) + ¬y<x (<-ports _ _ y<x) = + y≮x″ + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (VecCast.≈-reflexive ≡.refl)) + y<x) + + ... | tri> x≮y″ x≢y y<x = + tri> + ¬x<y + (λ x≡y → x≢y (≡⇒Pointwise-≡ (≡.trans (≡.sym (VecCast.≈-reflexive ≡.refl)) (≡ports x≡y)))) + (<-ports + ≡.refl + (HL.≈-sym (HL.≈-reflexive x≡y′)) + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (≡.sym (VecCast.≈-reflexive ≡.refl))) + y<x)) + where + open IsEquivalence ≡.isEquivalence using () renaming (isPartialEquivalence to ≡-isPartialEquivalence) + ¬x<y : ¬ <-Edge x y + ¬x<y (<-arity x<y) = x≮y x<y + ¬x<y (<-label _ x<y) = x≮y′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) x<y) + ¬x<y (<-ports _ _ x<y) = + x≮y″ + (Lex.<-respectsˡ + ≡-isPartialEquivalence + FinProp.<-respˡ-≡ + (≡⇒Pointwise-≡ (VecCast.≈-reflexive ≡.refl)) + x<y) + ... | tri> x≮y′ x≢y y<x = tri> + ¬x<y + (λ x≡y → x≢y (≡.trans (≡.sym (HL.≈-reflexive ≡.refl)) (≡label x≡y))) + (<-label ≡.refl (<-respˡ-≈ (≡.sym (HL.≈-reflexive ≡.refl)) y<x)) + where + ¬x<y : ¬ <-Edge x y + ¬x<y (<-arity x<y) = x≮y x<y + ¬x<y (<-label ≡a x<y) = x≮y′ (<-respˡ-≈ (HL.≈-reflexive ≡.refl) x<y) + ¬x<y (<-ports _ ≡l _) = x≢y (≡.trans (≡.sym (HL.≈-reflexive ≡.refl)) ≡l) +tri x y | tri> x≮y x≢y y<x = tri> ¬x<y (λ x≡y → x≢y (≡arity x≡y)) (<-arity y<x) + where + ¬x<y : ¬ <-Edge x y + ¬x<y (<-arity x<y) = x≮y x<y + ¬x<y (<-label ≡a x<y) = x≢y ≡a + ¬x<y (<-ports ≡a _ _) = x≢y ≡a + +isStrictTotalOrder : {v : ℕ} → IsStrictTotalOrder (≈-Edge {v}) (<-Edge {v}) +isStrictTotalOrder = record + { isStrictPartialOrder = record + { isEquivalence = ≈-Edge-IsEquivalence + ; irrefl = <-Edge-irrefl + ; trans = <-Edge-trans + ; <-resp-≈ = <-Edge-respʳ-≈ , <-Edge-respˡ-≈ + } + ; compare = tri + } + +strictTotalOrder : {v : ℕ} → StrictTotalOrder 0ℓ 0ℓ 0ℓ +strictTotalOrder {v} = record + { Carrier = Edge v + ; _≈_ = ≈-Edge {v} + ; _<_ = <-Edge {v} + ; isStrictTotalOrder = isStrictTotalOrder {v} + } + +open module STOP′ {v} = STOP (strictTotalOrder {v}) using (decTotalOrder) public + +module Sort {v} = ListSort (decTotalOrder {v}) +open Sort using (sort) public diff --git a/Data/Hypergraph/Label.agda b/Data/Hypergraph/Label.agda new file mode 100644 index 0000000..ca95002 --- /dev/null +++ b/Data/Hypergraph/Label.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --without-K --safe #-} +module Data.Hypergraph.Label where + +open import Data.Castable using (IsCastable) +open import Data.Nat.Base using (ℕ) +open import Data.String using (String) +open import Level using (Level; 0ℓ) +open import Relation.Binary using (Rel; IsDecTotalOrder) +open import Relation.Binary.Bundles using (DecTotalOrder; StrictTotalOrder) +open import Relation.Binary.Properties.DecTotalOrder using (<-strictTotalOrder; _<_) +open import Relation.Binary.PropositionalEquality using (_≡_) + +record HypergraphLabel : Set₁ where + + field + Label : ℕ → Set + showLabel : (n : ℕ) → Label n → String + isCastable : IsCastable Label + _[_≤_] : (n : ℕ) → Rel (Label n) 0ℓ + isDecTotalOrder : (n : ℕ) → IsDecTotalOrder _≡_ (n [_≤_]) + + decTotalOrder : (n : ℕ) → DecTotalOrder 0ℓ 0ℓ 0ℓ + decTotalOrder n = record + { Carrier = Label n + ; _≈_ = _≡_ + ; _≤_ = n [_≤_] + ; isDecTotalOrder = isDecTotalOrder n + } + + _[_<_] : (n : ℕ) → Rel (Label n) 0ℓ + _[_<_] n = _<_ (decTotalOrder n) + + strictTotalOrder : (n : ℕ) → StrictTotalOrder 0ℓ 0ℓ 0ℓ + strictTotalOrder n = <-strictTotalOrder (decTotalOrder n) + + open IsCastable isCastable public diff --git a/Data/Monoid.agda b/Data/Monoid.agda new file mode 100644 index 0000000..1ba0af4 --- /dev/null +++ b/Data/Monoid.agda @@ -0,0 +1,86 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +module Data.Monoid {c ℓ : Level} where + +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Category.Instance.Setoids.SymmetricMonoidal {c} {ℓ} using (Setoids-×; ×-monoidal′) +open import Categories.Object.Monoid using (Monoid; Monoid⇒) + +import Algebra.Bundles as Alg + +open import Data.Setoid using (∣_∣) +open import Relation.Binary using (Setoid) +open import Function using (Func) +open import Data.Product using (curry′; uncurry′; _,_) +open Func + +-- A monoid object in the (monoidal) category of setoids is just a monoid + +open import Function.Construct.Constant using () renaming (function to Const) +open import Data.Setoid.Unit using (⊤ₛ) + +opaque + unfolding ×-monoidal′ + toMonoid : Monoid Setoids-×.monoidal → Alg.Monoid c ℓ + toMonoid M = record + { Carrier = Carrier + ; _≈_ = _≈_ + ; _∙_ = curry′ (to μ) + ; ε = to η _ + ; isMonoid = record + { isSemigroup = record + { isMagma = record + { isEquivalence = isEquivalence + ; ∙-cong = curry′ (cong μ) + } + ; assoc = λ x y z → assoc {(x , y) , z} + } + ; identity = (λ x → sym (identityˡ {_ , x}) ) , λ x → sym (identityʳ {x , _}) + } + } + where + open Monoid M renaming (Carrier to A) + open Setoid A + + fromMonoid : Alg.Monoid c ℓ → Monoid Setoids-×.monoidal + fromMonoid M = record + { Carrier = setoid + ; isMonoid = record + { μ = record { to = uncurry′ _∙_ ; cong = uncurry′ ∙-cong } + ; η = Const ⊤ₛ setoid ε + ; assoc = λ { {(x , y) , z} → assoc x y z } + ; identityˡ = λ { {_ , x} → sym (identityˡ x) } + ; identityʳ = λ { {x , _} → sym (identityʳ x) } + } + } + where + open Alg.Monoid M + +-- A morphism of monoids in the (monoidal) category of setoids is a monoid homomorphism + +module _ (M N : Monoid Setoids-×.monoidal) where + + module M = Alg.Monoid (toMonoid M) + module N = Alg.Monoid (toMonoid N) + + open import Data.Product using (Σ; _,_) + open import Function using (_⟶ₛ_) + open import Algebra.Morphism using (IsMonoidHomomorphism) + open Monoid⇒ + + opaque + + unfolding toMonoid + + toMonoid⇒ + : Monoid⇒ Setoids-×.monoidal M N + → Σ (M.setoid ⟶ₛ N.setoid) (λ f + → IsMonoidHomomorphism M.rawMonoid N.rawMonoid (to f)) + toMonoid⇒ f = arr f , record + { isMagmaHomomorphism = record + { isRelHomomorphism = record { cong = cong (arr f) } + ; homo = λ x y → preserves-μ f {x , y} + } + ; ε-homo = preserves-η f + } diff --git a/Data/Opaque/List.agda b/Data/Opaque/List.agda new file mode 100644 index 0000000..83a2fe3 --- /dev/null +++ b/Data/Opaque/List.agda @@ -0,0 +1,165 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Opaque.List where + +import Data.List as L +import Function.Construct.Constant as Const +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning + +open import Algebra.Bundles using (Monoid) +open import Algebra.Morphism using (IsMonoidHomomorphism) +open import Data.List.Effectful.Foldable using (foldable; ++-homo) +open import Data.List.Relation.Binary.Pointwise as PW using (++⁺; map⁺) +open import Data.Product using (_,_; curry′; uncurry′) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid using (∣_∣) +open import Data.Unit.Polymorphic using (⊤) +open import Effect.Foldable using (RawFoldable) +open import Function using (_⟶ₛ_; Func; _⟨$⟩_; id) +open import Level using (Level; _⊔_) +open import Relation.Binary using (Setoid) + +open Func + +private + + variable + a c ℓ : Level + A B : Set a + Aₛ Bₛ : Setoid c ℓ + + ⊤ₛ : Setoid c ℓ + ⊤ₛ = record { Carrier = ⊤ ; _≈_ = λ _ _ → ⊤ } + +opaque + + List : Set a → Set a + List = L.List + + [] : List A + [] = L.[] + + _∷_ : A → List A → List A + _∷_ = L._∷_ + + map : (A → B) → List A → List B + map = L.map + + _++_ : List A → List A → List A + _++_ = L._++_ + + Listₛ : Setoid c ℓ → Setoid c (c ⊔ ℓ) + Listₛ = PW.setoid + + []ₛ : ⊤ₛ {c} {c ⊔ ℓ} ⟶ₛ Listₛ {c} {ℓ} Aₛ + []ₛ = Const.function ⊤ₛ (Listₛ _) [] + + ∷ₛ : Aₛ ×ₛ Listₛ Aₛ ⟶ₛ Listₛ Aₛ + ∷ₛ .to = uncurry′ _∷_ + ∷ₛ .cong = uncurry′ PW._∷_ + + [-]ₛ : Aₛ ⟶ₛ Listₛ Aₛ + [-]ₛ .to = L.[_] + [-]ₛ .cong y = y PW.∷ PW.[] + + mapₛ : (Aₛ ⟶ₛ Bₛ) → Listₛ Aₛ ⟶ₛ Listₛ Bₛ + mapₛ f .to = map (to f) + mapₛ f .cong xs≈ys = map⁺ (to f) (to f) (PW.map (cong f) xs≈ys) + + cartesianProduct : ∣ Listₛ Aₛ ∣ → ∣ Listₛ Bₛ ∣ → ∣ Listₛ (Aₛ ×ₛ Bₛ) ∣ + cartesianProduct = L.cartesianProduct + + cartesian-product-cong + : {xs xs′ : ∣ Listₛ Aₛ ∣} + {ys ys′ : ∣ Listₛ Bₛ ∣} + → (let open Setoid (Listₛ Aₛ) in xs ≈ xs′) + → (let open Setoid (Listₛ Bₛ) in ys ≈ ys′) + → (let open Setoid (Listₛ (Aₛ ×ₛ Bₛ)) in cartesianProduct xs ys ≈ cartesianProduct xs′ ys′) + cartesian-product-cong PW.[] ys≋ys′ = PW.[] + cartesian-product-cong {Aₛ = Aₛ} {Bₛ = Bₛ} {xs = x₀ L.∷ xs} {xs′ = x₀′ L.∷ xs′} (x₀≈x₀′ PW.∷ xs≋xs′) ys≋ys′ = + ++⁺ + (map⁺ (x₀ ,_) (x₀′ ,_) (PW.map (x₀≈x₀′ ,_) ys≋ys′)) + (cartesian-product-cong {Aₛ = Aₛ} {Bₛ = Bₛ} xs≋xs′ ys≋ys′) + + pairsₛ : Listₛ Aₛ ×ₛ Listₛ Bₛ ⟶ₛ Listₛ (Aₛ ×ₛ Bₛ) + pairsₛ .to = uncurry′ L.cartesianProduct + pairsₛ {Aₛ = Aₛ} {Bₛ = Bₛ} .cong = uncurry′ (cartesian-product-cong {Aₛ = Aₛ} {Bₛ = Bₛ}) + + ++ₛ : Listₛ Aₛ ×ₛ Listₛ Aₛ ⟶ₛ Listₛ Aₛ + ++ₛ .to = uncurry′ _++_ + ++ₛ .cong = uncurry′ ++⁺ + + foldr : (Aₛ ×ₛ Bₛ ⟶ₛ Bₛ) → ∣ Bₛ ∣ → ∣ Listₛ Aₛ ∣ → ∣ Bₛ ∣ + foldr f = L.foldr (curry′ (to f)) + + foldr-cong + : (f : Aₛ ×ₛ Bₛ ⟶ₛ Bₛ) + → (e : ∣ Bₛ ∣) + → (let module [A]ₛ = Setoid (Listₛ Aₛ)) + → {xs ys : ∣ Listₛ Aₛ ∣} + → (xs [A]ₛ.≈ ys) + → (open Setoid Bₛ) + → foldr f e xs ≈ foldr f e ys + foldr-cong {Bₛ = Bₛ} f e PW.[] = Setoid.refl Bₛ + foldr-cong f e (x≈y PW.∷ xs≋ys) = cong f (x≈y , foldr-cong f e xs≋ys) + + foldrₛ : (Aₛ ×ₛ Bₛ ⟶ₛ Bₛ) → ∣ Bₛ ∣ → Listₛ Aₛ ⟶ₛ Bₛ + foldrₛ f e .to = foldr f e + foldrₛ {Bₛ = Bₛ} f e .cong = foldr-cong f e + +module _ (M : Monoid c ℓ) where + + open Monoid M renaming (setoid to Mₛ) + + opaque + unfolding List + fold : ∣ Listₛ Mₛ ∣ → ∣ Mₛ ∣ + fold = RawFoldable.fold foldable rawMonoid + + fold-cong + : {xs ys : ∣ Listₛ Mₛ ∣} + → (let module [M]ₛ = Setoid (Listₛ Mₛ)) + → (xs [M]ₛ.≈ ys) + → fold xs ≈ fold ys + fold-cong = PW.rec (λ {xs} {ys} _ → fold xs ≈ fold ys) ∙-cong refl + + foldₛ : Listₛ Mₛ ⟶ₛ Mₛ + foldₛ .to = fold + foldₛ .cong = fold-cong + + opaque + unfolding fold + ++ₛ-homo + : (xs ys : ∣ Listₛ Mₛ ∣) + → foldₛ ⟨$⟩ (++ₛ ⟨$⟩ (xs , ys)) ≈ (foldₛ ⟨$⟩ xs) ∙ (foldₛ ⟨$⟩ ys) + ++ₛ-homo xs ys = ++-homo M id xs + + []ₛ-homo : foldₛ ⟨$⟩ ([]ₛ ⟨$⟩ _) ≈ ε + []ₛ-homo = refl + +module _ (M N : Monoid c ℓ) where + + module M = Monoid M + module N = Monoid N + + open IsMonoidHomomorphism + + opaque + unfolding fold + + fold-mapₛ + : (f : M.setoid ⟶ₛ N.setoid) + → IsMonoidHomomorphism M.rawMonoid N.rawMonoid (to f) + → {xs : ∣ Listₛ M.setoid ∣} + → foldₛ N ⟨$⟩ (mapₛ f ⟨$⟩ xs) N.≈ f ⟨$⟩ (foldₛ M ⟨$⟩ xs) + fold-mapₛ f isMH {L.[]} = N.sym (ε-homo isMH) + fold-mapₛ f isMH {x L.∷ xs} = begin + f′ x ∙ fold N (map f′ xs) ≈⟨ N.∙-cong N.refl (fold-mapₛ f isMH {xs}) ⟩ + f′ x ∙ f′ (fold M xs) ≈⟨ homo isMH x (fold M xs) ⟨ + f′ (x ∘ fold M xs) ∎ + where + open N using (_∙_) + open M using () renaming (_∙_ to _∘_) + open ≈-Reasoning N.setoid + f′ : M.Carrier → N.Carrier + f′ = to f diff --git a/Data/Opaque/Multiset.agda b/Data/Opaque/Multiset.agda new file mode 100644 index 0000000..99501d6 --- /dev/null +++ b/Data/Opaque/Multiset.agda @@ -0,0 +1,131 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --hidden-argument-puns #-} + +module Data.Opaque.Multiset where + +import Data.List as L +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning +import Data.Opaque.List as List + +open import Algebra.Bundles using (CommutativeMonoid) +open import Data.List.Effectful.Foldable using (foldable; ++-homo) +open import Data.List.Relation.Binary.Permutation.Setoid as ↭ using (↭-setoid; ↭-refl) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (map⁺; ++⁺; ++-comm) +open import Algebra.Morphism using (IsMonoidHomomorphism) +open import Data.Product using (_,_; uncurry′) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid using (∣_∣) +open import Data.Setoid.Unit using (⊤ₛ) +open import Effect.Foldable using (RawFoldable) +open import Function using (_⟶ₛ_; Func; _⟨$⟩_; id) +open import Function.Construct.Constant using () renaming (function to Const) +open import Level using (Level; _⊔_) +open import Relation.Binary using (Setoid) + +open Func + +private + + variable + a c ℓ : Level + A B : Set a + Aₛ Bₛ : Setoid c ℓ + +opaque + + Multiset : Set a → Set a + Multiset = L.List + + [] : Multiset A + [] = L.[] + + _∷_ : A → Multiset A → Multiset A + _∷_ = L._∷_ + + map : (A → B) → Multiset A → Multiset B + map = L.map + + _++_ : Multiset A → Multiset A → Multiset A + _++_ = L._++_ + + Multisetₛ : Setoid c ℓ → Setoid c (c ⊔ ℓ) + Multisetₛ = ↭-setoid + + []ₛ : ⊤ₛ {c} {c ⊔ ℓ} ⟶ₛ Multisetₛ {c} {ℓ} Aₛ + []ₛ {Aₛ} = Const ⊤ₛ (Multisetₛ Aₛ) [] + + ∷ₛ : Aₛ ×ₛ Multisetₛ Aₛ ⟶ₛ Multisetₛ Aₛ + ∷ₛ .to = uncurry′ _∷_ + ∷ₛ .cong = uncurry′ ↭.prep + + [-]ₛ : Aₛ ⟶ₛ Multisetₛ Aₛ + [-]ₛ .to = L.[_] + [-]ₛ {Aₛ} .cong y = ↭.prep y (↭-refl Aₛ) + + mapₛ : (Aₛ ⟶ₛ Bₛ) → Multisetₛ Aₛ ⟶ₛ Multisetₛ Bₛ + mapₛ f .to = map (to f) + mapₛ {Aₛ} {Bₛ} f .cong xs≈ys = map⁺ Aₛ Bₛ (cong f) xs≈ys + + ++ₛ : Multisetₛ Aₛ ×ₛ Multisetₛ Aₛ ⟶ₛ Multisetₛ Aₛ + ++ₛ .to = uncurry′ _++_ + ++ₛ {Aₛ} .cong = uncurry′ (++⁺ Aₛ) + + ++ₛ-comm + : (open Setoid (Multisetₛ Aₛ)) + → (xs ys : Carrier) + → ++ₛ ⟨$⟩ (xs , ys) ≈ ++ₛ ⟨$⟩ (ys , xs) + ++ₛ-comm {Aₛ} xs ys = ++-comm Aₛ xs ys + +module _ (M : CommutativeMonoid c ℓ) where + + open CommutativeMonoid M renaming (setoid to Mₛ) + + opaque + unfolding Multiset List.fold-cong + fold : ∣ Multisetₛ Mₛ ∣ → ∣ Mₛ ∣ + fold = List.fold monoid -- RawFoldable.fold foldable rawMonoid + + fold-cong + : {xs ys : ∣ Multisetₛ Mₛ ∣} + → (let module [M]ₛ = Setoid (Multisetₛ Mₛ)) + → (xs [M]ₛ.≈ ys) + → fold xs ≈ fold ys + fold-cong (↭.refl x) = cong (List.foldₛ monoid) x + fold-cong (↭.prep x≈y xs↭ys) = ∙-cong x≈y (fold-cong xs↭ys) + fold-cong + {x₀ L.∷ x₁ L.∷ xs} + {y₀ L.∷ y₁ L.∷ ys} + (↭.swap x₀≈y₁ x₁≈y₀ xs↭ys) = trans + (sym (assoc x₀ x₁ (fold xs))) (trans + (∙-cong (trans (∙-cong x₀≈y₁ x₁≈y₀) (comm y₁ y₀)) (fold-cong xs↭ys)) + (assoc y₀ y₁ (fold ys))) + fold-cong {xs} {ys} (↭.trans xs↭zs zs↭ys) = trans (fold-cong xs↭zs) (fold-cong zs↭ys) + + foldₛ : Multisetₛ Mₛ ⟶ₛ Mₛ + foldₛ .to = fold + foldₛ .cong = fold-cong + + opaque + unfolding fold + ++ₛ-homo + : (xs ys : ∣ Multisetₛ Mₛ ∣) + → foldₛ ⟨$⟩ (++ₛ ⟨$⟩ (xs , ys)) ≈ (foldₛ ⟨$⟩ xs) ∙ (foldₛ ⟨$⟩ ys) + ++ₛ-homo xs ys = ++-homo monoid id xs + + []ₛ-homo : foldₛ ⟨$⟩ ([]ₛ ⟨$⟩ _) ≈ ε + []ₛ-homo = refl + +module _ (M N : CommutativeMonoid c ℓ) where + + module M = CommutativeMonoid M + module N = CommutativeMonoid N + + opaque + unfolding fold + + fold-mapₛ + : (f : M.setoid ⟶ₛ N.setoid) + → IsMonoidHomomorphism M.rawMonoid N.rawMonoid (to f) + → {xs : ∣ Multisetₛ M.setoid ∣} + → foldₛ N ⟨$⟩ (mapₛ f ⟨$⟩ xs) N.≈ f ⟨$⟩ (foldₛ M ⟨$⟩ xs) + fold-mapₛ = List.fold-mapₛ M.monoid N.monoid diff --git a/Data/Permutation.agda b/Data/Permutation.agda new file mode 100644 index 0000000..55c8748 --- /dev/null +++ b/Data/Permutation.agda @@ -0,0 +1,217 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module Data.Permutation {ℓ : Level} {A : Set ℓ} where + +import Data.Fin as Fin +import Data.Fin.Properties as FinProp +import Data.Fin.Permutation as ↔-Fin +import Data.List as List +import Data.List.Properties as ListProp +import Data.List.Relation.Binary.Permutation.Propositional as ↭-List +import Data.Nat as Nat +import Data.Vec.Functional as Vector +import Data.Vec.Functional.Properties as VectorProp +import Data.Vec.Functional.Relation.Binary.Permutation as ↭-Vec +import Data.Vec.Functional.Relation.Binary.Permutation.Properties as ↭-VecProp + +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) +open import Data.Product.Base using (_,_; proj₁; proj₂) +open import Data.Vec.Functional.Relation.Unary.Any using (Any) +open import Function.Base using (_∘_) + +open ↭-Vec using () renaming (_↭_ to _↭′_) +open ↭-List using (_↭_; ↭-sym; module PermutationReasoning) +open ↔-Fin using (Permutation; _⟨$⟩ˡ_; _⟨$⟩ʳ_) +open List using (List) hiding (module List) +open Fin using (Fin; cast) hiding (module Fin) +open Vector using (Vector; toList; fromList) +open Fin.Fin +open Nat using (ℕ; zero; suc) +open _↭_ + +module _ where + + open Fin using (#_) + open ↔-Fin using (lift₀; insert) + open ↭-VecProp using (↭-refl; ↭-trans) + + -- convert a List permutation to a Vector permutation + fromList-↭ + : {xs ys : List A} + → xs ↭ ys + → fromList xs ↭′ fromList ys + fromList-↭ refl = ↭-refl + fromList-↭ (prep _ xs↭ys) + with n↔m , xs↭ys′ ← fromList-↭ xs↭ys = lift₀ n↔m , λ where + zero → ≡.refl + (suc i) → xs↭ys′ i + fromList-↭ (swap x y xs↭ys) + with n↔m , xs↭ys′ ← fromList-↭ xs↭ys = insert (# 0) (# 1) (lift₀ n↔m) , λ where + zero → ≡.refl + (suc zero) → ≡.refl + (suc (suc i)) → xs↭ys′ i + fromList-↭ (trans {xs} xs↭ys ys↭zs) = + ↭-trans {_} {_} {_} {i = fromList xs} (fromList-↭ xs↭ys) (fromList-↭ ys↭zs) + +-- witness for an element in a Vector +_∈_ : A → {n : ℕ} → Vector A n → Set ℓ +x ∈ xs = Any (x ≡_) xs + +-- apply a permutation to a witness +apply + : {n m : ℕ} + {xs : Vector A n} + {ys : Vector A m} + {x : A} + → (xs↭ys : xs ↭′ ys) + → x ∈ xs + → x ∈ ys +apply {suc n} {zero} (↔n , _) (i , _) with () ← ↔n ⟨$⟩ˡ i +apply {suc n} {suc m} {xs} {ys} {x} (↔n , ↔≗) (i , x≡xs-i) = i′ , x≡ys-i′ + where + i′ : Fin (suc m) + i′ = ↔n ⟨$⟩ˡ i + open ≡.≡-Reasoning + x≡ys-i′ : x ≡ ys i′ + x≡ys-i′ = begin + x ≡⟨ x≡xs-i ⟩ + xs i ≡⟨ ≡.cong xs (↔-Fin.inverseʳ ↔n) ⟨ + xs (↔n ⟨$⟩ʳ (↔n ⟨$⟩ˡ i)) ≡⟨⟩ + xs (↔n ⟨$⟩ʳ i′) ≡⟨ ↔≗ i′ ⟩ + ys i′ ∎ + +-- remove an element from a Vector +remove : {n : ℕ} {x : A} (xs : Vector A (suc n)) → x ∈ xs → Vector A n +remove xs (i , _) = Vector.removeAt xs i + +-- remove an element and its image from a permutation +↭-remove + : {n m : ℕ} + {xs : Vector A (suc n)} + {ys : Vector A (suc m)} + {x : A} + → (xs↭ys : xs ↭′ ys) + → (x∈xs : x ∈ xs) + → let x∈ys = apply xs↭ys x∈xs in + remove xs x∈xs ↭′ remove ys x∈ys +↭-remove {n} {m} {xs} {ys} {x} xs↭ys@(ρ , ↔≗) x∈xs@(k , _) = ρ⁻ , ↔≗⁻ + where + k′ : Fin (suc m) + k′ = ρ ⟨$⟩ˡ k + x∈ys : x ∈ ys + x∈ys = apply xs↭ys x∈xs + ρ⁻ : Permutation m n + ρ⁻ = ↔-Fin.remove k′ ρ + xs⁻ : Vector A n + xs⁻ = remove xs x∈xs + ys⁻ : Vector A m + ys⁻ = remove ys x∈ys + open ≡.≡-Reasoning + open Fin using (punchIn) + ↔≗⁻ : (i : Fin m) → xs⁻ (ρ⁻ ⟨$⟩ʳ i) ≡ ys⁻ i + ↔≗⁻ i = begin + xs⁻ (ρ⁻ ⟨$⟩ʳ i) ≡⟨⟩ + remove xs x∈xs (ρ⁻ ⟨$⟩ʳ i) ≡⟨⟩ + xs (punchIn k (ρ⁻ ⟨$⟩ʳ i)) ≡⟨ ≡.cong xs (↔-Fin.punchIn-permute′ ρ k i) ⟨ + xs (ρ ⟨$⟩ʳ (punchIn k′ i)) ≡⟨ ↔≗ (punchIn k′ i) ⟩ + ys (punchIn k′ i) ≡⟨⟩ + remove ys x∈ys i ≡⟨⟩ + ys⁻ i ∎ + +open List.List +open List using (length; insertAt) + +-- build a permutation which moves the first element of a list to an arbitrary position +↭-insert-half + : {x₀ : A} + {xs : List A} + → (i : Fin (suc (length xs))) + → x₀ ∷ xs ↭ insertAt xs i x₀ +↭-insert-half zero = refl +↭-insert-half {x₀} {x₁ ∷ xs} (suc i) = trans (swap x₀ x₁ refl) (prep x₁ (↭-insert-half i)) + +-- add a mapping to a permutation, given a value and its start and end positions +↭-insert + : {xs ys : List A} + → xs ↭ ys + → (i : Fin (suc (length xs))) + (j : Fin (suc (length ys))) + (x : A) + → insertAt xs i x ↭ insertAt ys j x +↭-insert {xs} {ys} xs↭ys i j x = xs↭ys⁺ + where + open PermutationReasoning + xs↭ys⁺ : insertAt xs i x ↭ insertAt ys j x + xs↭ys⁺ = begin + insertAt xs i x ↭⟨ ↭-insert-half i ⟨ + x ∷ xs <⟨ xs↭ys ⟩ + x ∷ ys ↭⟨ ↭-insert-half j ⟩ + insertAt ys j x ∎ + +open ListProp using (length-tabulate; tabulate-cong) +insertAt-toList + : {n : ℕ} + {v : Vector A n} + (i : Fin (suc (length (toList v)))) + (i′ : Fin (suc n)) + → i ≡ cast (≡.cong suc (≡.sym (length-tabulate v))) i′ + → (x : A) + → insertAt (toList v) i x + ≡ toList (Vector.insertAt v i′ x) +insertAt-toList zero zero _ x = ≡.refl +insertAt-toList {suc n} {v} (suc i) (suc i′) ≡.refl x = + ≡.cong (v zero ∷_) (insertAt-toList i i′ ≡.refl x) + +-- convert a Vector permutation to a List permutation +toList-↭ + : {n m : ℕ} + {xs : Vector A n} + {ys : Vector A m} + → xs ↭′ ys + → toList xs ↭ toList ys +toList-↭ {zero} {zero} _ = refl +toList-↭ {zero} {suc m} (ρ , _) with () ← ρ ⟨$⟩ʳ zero +toList-↭ {suc n} {zero} (ρ , _) with () ← ρ ⟨$⟩ˡ zero +toList-↭ {suc n} {suc m} {xs} {ys} xs↭ys′ = xs↭ys + where + -- first element and its image + x₀ : A + x₀ = xs zero + x₀∈xs : x₀ ∈ xs + x₀∈xs = zero , ≡.refl + x₀∈ys : x₀ ∈ ys + x₀∈ys = apply xs↭ys′ x₀∈xs + -- reduce the problem by removing those elements + xs⁻ : Vector A n + xs⁻ = remove xs x₀∈xs + ys⁻ : Vector A m + ys⁻ = remove ys x₀∈ys + xs↭ys⁻ : xs⁻ ↭′ ys⁻ + xs↭ys⁻ = ↭-remove xs↭ys′ x₀∈xs + -- recursion on the reduced problem + xs↭ys⁻′ : toList xs⁻ ↭ toList ys⁻ + xs↭ys⁻′ = toList-↭ xs↭ys⁻ + -- indices for working with vectors + i : Fin (suc n) + i = proj₁ x₀∈xs + j : Fin (suc m) + j = proj₁ x₀∈ys + i′ : Fin (suc (length (toList xs⁻))) + i′ = cast (≡.sym (≡.cong suc (length-tabulate xs⁻))) i + j′ : Fin (suc (length (toList ys⁻))) + j′ = cast (≡.sym (≡.cong suc (length-tabulate ys⁻))) j + -- main construction + open VectorProp using (insertAt-removeAt) + open PermutationReasoning + open Vector using () renaming (insertAt to insertAt′) + xs↭ys : toList xs ↭ toList ys + xs↭ys = begin + toList xs ≡⟨ tabulate-cong (insertAt-removeAt xs i) ⟨ + toList (insertAt′ xs⁻ i x₀) ≡⟨ insertAt-toList i′ i ≡.refl x₀ ⟨ + insertAt (toList xs⁻) i′ x₀ ↭⟨ ↭-insert xs↭ys⁻′ i′ j′ x₀ ⟩ + insertAt (toList ys⁻) j′ x₀ ≡⟨ insertAt-toList j′ j ≡.refl x₀ ⟩ + toList (insertAt′ ys⁻ j x₀) ≡⟨ ≡.cong (toList ∘ insertAt′ ys⁻ j) (proj₂ x₀∈ys) ⟩ + toList (insertAt′ ys⁻ j (ys j)) ≡⟨ tabulate-cong (insertAt-removeAt ys j) ⟩ + toList ys ∎ diff --git a/Data/Permutation/Sort.agda b/Data/Permutation/Sort.agda new file mode 100644 index 0000000..8d097f2 --- /dev/null +++ b/Data/Permutation/Sort.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --without-K --safe #-} + +open import Relation.Binary.Bundles using (DecTotalOrder) + +module Data.Permutation.Sort {ℓ₁ ℓ₂ ℓ₃} (dto : DecTotalOrder ℓ₁ ℓ₂ ℓ₃) where + +open DecTotalOrder dto using (module Eq; totalOrder) renaming (Carrier to A) + +open import Data.List.Base using (List) +open import Data.List.Relation.Binary.Equality.Setoid Eq.setoid using (_≋_) +open import Data.List.Relation.Binary.Permutation.Setoid Eq.setoid using (_↭_; module PermutationReasoning) +open import Data.List.Relation.Unary.Sorted.TotalOrder.Properties using (↗↭↗⇒≋) +open import Data.List.Sort dto using (sortingAlgorithm) +open import Data.List.Sort.Base totalOrder using (SortingAlgorithm) +open SortingAlgorithm sortingAlgorithm using (sort; sort-↭ₛ; sort-↗) + +sorted-≋ + : {xs ys : List A} + → xs ↭ ys + → sort xs ≋ sort ys +sorted-≋ {xs} {ys} xs↭ys = ↗↭↗⇒≋ totalOrder (sort-↗ xs) (sort-↗ ys) sort-xs↭sort-ys + where + open PermutationReasoning + sort-xs↭sort-ys : sort xs ↭ sort ys + sort-xs↭sort-ys = begin + sort xs ↭⟨ sort-↭ₛ xs ⟩ + xs ↭⟨ xs↭ys ⟩ + ys ↭⟨ sort-↭ₛ ys ⟨ + sort ys ∎ diff --git a/Data/SExp.agda b/Data/SExp.agda new file mode 100644 index 0000000..adf4325 --- /dev/null +++ b/Data/SExp.agda @@ -0,0 +1,75 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.SExp where + +open import Data.List using (List) +open import Data.Nat.Base using (ℕ) +open import Data.Nat.Show using () renaming (show to showNat) +open import Data.String.Base as String using (String; parens; intersperse; _<+>_) + +open List + +module ListBased where + + data SExp : Set where + Atom : String → SExp + Nat : ℕ → SExp + SExps : List SExp → SExp + + mutual + showSExp : SExp → String + showSExp (Atom s) = s + showSExp (Nat n) = showNat n + showSExp (SExps xs) = parens (intersperse " " (showList xs)) + + -- expanded out map for termination checker + showList : List SExp → List String + showList [] = [] + showList (x ∷ xs) = showSExp x ∷ showList xs + +module PairBased where + + data SExp : Set where + Atom : String → SExp + Nat : ℕ → SExp + Nil : SExp + Pair : SExp → SExp → SExp + + mutual + showSExp : SExp → String + showSExp (Atom s) = s + showSExp (Nat n) = showNat n + showSExp Nil = "()" + showSExp (Pair l r) = parens (showPair l r) + + showPair : SExp → SExp → String + showPair x (Atom s) = showSExp x <+> "." <+> s + showPair x (Nat n) = showSExp x <+> "." <+> showNat n + showPair x Nil = showSExp x + showPair x (Pair l r) = showSExp x <+> showPair l r + +open ListBased public +open PairBased + +mutual + sexpL→sexpP : ListBased.SExp → PairBased.SExp + sexpL→sexpP (Atom s) = Atom s + sexpL→sexpP (Nat n) = Nat n + sexpL→sexpP (SExps xs) = [sexpL]→sexpP xs + + [sexpL]→sexpP : List (ListBased.SExp) → PairBased.SExp + [sexpL]→sexpP [] = Nil + [sexpL]→sexpP (x ∷ xs) = Pair (sexpL→sexpP x) ([sexpL]→sexpP xs) + +mutual + sexpP→sexpL : PairBased.SExp → ListBased.SExp + sexpP→sexpL (Atom s) = Atom s + sexpP→sexpL (Nat n) = Nat n + sexpP→sexpL Nil = SExps [] + sexpP→sexpL (Pair x y) = SExps (sexpP→sexpL x ∷ sexpP→[sexpL] y) + + sexpP→[sexpL] : PairBased.SExp → List (ListBased.SExp) + sexpP→[sexpL] (Atom _) = [] + sexpP→[sexpL] (Nat _) = [] + sexpP→[sexpL] Nil = [] + sexpP→[sexpL] (Pair x y) = sexpP→sexpL x ∷ sexpP→[sexpL] y diff --git a/Data/SExp/Parser.agda b/Data/SExp/Parser.agda new file mode 100644 index 0000000..ec1b6a5 --- /dev/null +++ b/Data/SExp/Parser.agda @@ -0,0 +1,73 @@ +{-# OPTIONS --without-K --guardedness --safe #-} + +open import Text.Parser.Types.Core using (Parameters) + +module Data.SExp.Parser {l} {P : Parameters l} where + +import Data.String.Base as String +open import Data.List as List using (List) + +open import Data.Char.Base using (Char) +open import Data.List.Sized.Interface using (Sized) +open import Data.List.NonEmpty as List⁺ using (List⁺) renaming (map to map⁺) +open import Data.Subset using (Subset; into) +open import Effect.Monad using (RawMonadPlus) +open import Function.Base using (_∘_; _$_) +open import Induction.Nat.Strong using (fix; map; □_) +open import Relation.Binary.PropositionalEquality.Decidable using (DecidableEquality) +open import Relation.Unary using (IUniversal; _⇒_) +open import Level.Bounded using (theSet; [_]) +open import Text.Parser.Types P using (Parser) +open import Text.Parser.Combinators {P = P} using (_<$>_; list⁺; _<|>_; _<$_; _<&_; _<&?_; box) +open import Text.Parser.Combinators.Char {P = P} using (alpha; parens; withSpaces; spaces; char) +open import Text.Parser.Combinators.Numbers {P = P} using (decimalℕ) + +open import Data.SExp using (SExp) +open SExp + +open Parameters P using (M; Tok; Toks) + +module _ + {{𝕄 : RawMonadPlus M}} + {{𝕊 : Sized Tok Toks}} + {{𝔻 : DecidableEquality (theSet Tok)}} + {{ℂ : Subset Char (theSet Tok)}} + {{ℂ⁻ : Subset (theSet Tok) Char}} + where + + -- An atom is just a non-empty list of alphabetical characters. + -- We use `<$>` to turn that back into a string and apply the `Atom` constructor. + atom : ∀[ Parser [ SExp ] ] + atom = Atom ∘ String.fromList⁺ ∘ map⁺ (into ℂ⁻) <$> list⁺ alpha + + -- Natural number parser + nat : ∀[ Parser [ SExp ] ] + nat = Nat <$> decimalℕ + + -- Empty list parser + nil : ∀[ Parser [ SExp ] ] + nil = SExps List.[] <$ char '(' <&? box spaces <& box (char ')') + + -- Parser for a list of S-Expressions + list : ∀[ Parser [ SExp ] ⇒ Parser [ List SExp ] ] + list sexp = List⁺.toList <$> list⁺ (withSpaces sexp) + + -- Compound S-Expression parser + compound : ∀[ □ Parser [ SExp ] ⇒ Parser [ SExp ] ] + compound rec = nil <|> SExps <$> parens (map list rec) + + -- S-Expression parser + sexp : ∀[ □ Parser [ SExp ] ⇒ Parser [ SExp ] ] + sexp rec = atom <|> nat <|> compound rec + + -- Build the parser as a fixpoint since SExp is an inductive type + sexp-top : ∀[ Parser [ SExp ] ] + sexp-top = fix (Parser [ SExp ]) sexp + + -- Top-level parser for list of S-Expressions + LIST : ∀[ Parser [ SExp ] ] + LIST = SExps <$> list sexp-top + + -- Top-level S-Expression parser + SEXP : ∀[ Parser [ SExp ] ] + SEXP = withSpaces sexp-top diff --git a/Data/Setoid.agda b/Data/Setoid.agda new file mode 100644 index 0000000..454d276 --- /dev/null +++ b/Data/Setoid.agda @@ -0,0 +1,8 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Setoid where + +open import Relation.Binary using (Setoid) +open import Function.Construct.Setoid using () renaming (setoid to infixr 0 _⇒ₛ_) public + +open Setoid renaming (Carrier to ∣_∣) public diff --git a/Data/Setoid/Unit.agda b/Data/Setoid/Unit.agda new file mode 100644 index 0000000..8b8edde --- /dev/null +++ b/Data/Setoid/Unit.agda @@ -0,0 +1,11 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module Data.Setoid.Unit {c ℓ : Level} where + +open import Categories.Category.Instance.SingletonSet using (SingletonSetoid) +open import Relation.Binary using (Setoid) + +⊤ₛ : Setoid c ℓ +⊤ₛ = SingletonSetoid {c} {ℓ} diff --git a/Data/Subset/Functional.agda b/Data/Subset/Functional.agda new file mode 100644 index 0000000..33a6d8e --- /dev/null +++ b/Data/Subset/Functional.agda @@ -0,0 +1,162 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Subset.Functional where + +open import Data.Bool.Base using (Bool; _∨_; _∧_; if_then_else_) +open import Data.Bool.Properties using (if-float) +open import Data.Fin.Base using (Fin) +open import Data.Fin.Permutation using (Permutation′; _⟨$⟩ʳ_; _⟨$⟩ˡ_; inverseˡ; inverseʳ) +open import Data.Fin.Properties using (suc-injective; 0≢1+n) +open import Data.Nat.Base using (ℕ) +open import Function.Base using (∣_⟩-_; _∘_) +open import Function.Definitions using (Injective) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; _≢_) +open import Data.Vector as V using (Vector; head; tail) + +open Bool +open Fin +open ℕ + +Subset : ℕ → Set +Subset = Vector Bool + +private + variable n A : ℕ + variable B C : Set + +⊥ : Subset n +⊥ _ = false + +_∪_ : Subset n → Subset n → Subset n +(A ∪ B) k = A k ∨ B k + +_∩_ : Subset n → Subset n → Subset n +(A ∩ B) k = A k ∧ B k + +⁅_⁆ : Fin n → Subset n +⁅_⁆ zero zero = true +⁅_⁆ zero (suc _) = false +⁅_⁆ (suc k) zero = false +⁅_⁆ (suc k) (suc i) = ⁅ k ⁆ i + +⁅⁆-refl : (k : Fin n) → ⁅ k ⁆ k ≡ true +⁅⁆-refl zero = ≡.refl +⁅⁆-refl (suc k) = ⁅⁆-refl k + +⁅x⁆y≡true + : (x y : Fin n) + → .(⁅ x ⁆ y ≡ true) + → x ≡ y +⁅x⁆y≡true zero zero prf = ≡.refl +⁅x⁆y≡true (suc x) (suc y) prf = ≡.cong suc (⁅x⁆y≡true x y prf) + +⁅x⁆y≡false + : (x y : Fin n) + → .(⁅ x ⁆ y ≡ false) + → x ≢ y +⁅x⁆y≡false zero (suc y) prf = 0≢1+n +⁅x⁆y≡false (suc x) zero prf = 0≢1+n ∘ ≡.sym +⁅x⁆y≡false (suc x) (suc y) prf = ⁅x⁆y≡false x y prf ∘ suc-injective + +f-⁅⁆ + : {n m : ℕ} + (f : Fin n → Fin m) + → Injective _≡_ _≡_ f + → (x y : Fin n) + → ⁅ x ⁆ y ≡ ⁅ f x ⁆ (f y) +f-⁅⁆ f f-inj zero zero = ≡.sym (⁅⁆-refl (f zero)) +f-⁅⁆ f f-inj zero (suc y) with ⁅ f zero ⁆ (f (suc y)) in eq +... | true with () ← f-inj (⁅x⁆y≡true (f zero) (f (suc y)) eq) +... | false = ≡.refl +f-⁅⁆ f f-inj (suc x) zero with ⁅ f (suc x) ⁆ (f zero) in eq +... | true with () ← f-inj (⁅x⁆y≡true (f (suc x)) (f zero) eq) +... | false = ≡.refl +f-⁅⁆ f f-inj (suc x) (suc y) = f-⁅⁆ (f ∘ suc) (suc-injective ∘ f-inj) x y + +⁅⁆∘ρ + : (ρ : Permutation′ (suc n)) + (x : Fin (suc n)) + → ⁅ ρ ⟨$⟩ʳ x ⁆ ≗ ⁅ x ⁆ ∘ (ρ ⟨$⟩ˡ_) +⁅⁆∘ρ {n} ρ x y = begin + ⁅ ρ ⟨$⟩ʳ x ⁆ y ≡⟨ f-⁅⁆ (ρ ⟨$⟩ˡ_) ρˡ-inj (ρ ⟨$⟩ʳ x) y ⟩ + ⁅ ρ ⟨$⟩ˡ (ρ ⟨$⟩ʳ x) ⁆ (ρ ⟨$⟩ˡ y) ≡⟨ ≡.cong (λ h → ⁅ h ⁆ (ρ ⟨$⟩ˡ y)) (inverseˡ ρ) ⟩ + ⁅ x ⁆ (ρ ⟨$⟩ˡ y) ∎ + where + open ≡.≡-Reasoning + ρˡ-inj : {x y : Fin (suc n)} → ρ ⟨$⟩ˡ x ≡ ρ ⟨$⟩ˡ y → x ≡ y + ρˡ-inj {x} {y} ρˡx≡ρˡy = begin + x ≡⟨ inverseʳ ρ ⟨ + ρ ⟨$⟩ʳ (ρ ⟨$⟩ˡ x) ≡⟨ ≡.cong (ρ ⟨$⟩ʳ_) ρˡx≡ρˡy ⟩ + ρ ⟨$⟩ʳ (ρ ⟨$⟩ˡ y) ≡⟨ inverseʳ ρ ⟩ + y ∎ + +opaque + -- TODO dependent fold + foldl : (B → Fin A → B) → B → Subset A → B + foldl {B = B} f = V.foldl (λ _ → B) (λ { {k} acc b → if b then f acc k else acc }) + + foldl-cong₁ + : {f g : B → Fin A → B} + → (∀ x y → f x y ≡ g x y) + → (e : B) + → (S : Subset A) + → foldl f e S ≡ foldl g e S + foldl-cong₁ {B = B} f≗g e S = V.foldl-cong (λ _ → B) (λ { {k} x y → ≡.cong (if y then_else x) (f≗g x k) }) e S + + foldl-cong₂ + : (f : B → Fin A → B) + (e : B) + {S₁ S₂ : Subset A} + → (S₁ ≗ S₂) + → foldl f e S₁ ≡ foldl f e S₂ + foldl-cong₂ {B = B} f e S₁≗S₂ = V.foldl-cong-arg (λ _ → B) (λ {n} acc b → if b then f acc n else acc) e S₁≗S₂ + + foldl-suc + : (f : B → Fin (suc A) → B) + → (e : B) + → (S : Subset (suc A)) + → foldl f e S ≡ foldl (∣ f ⟩- suc) (if head S then f e zero else e) (tail S) + foldl-suc f e S = ≡.refl + + foldl-⊥ + : {A : ℕ} + {B : Set} + (f : B → Fin A → B) + (e : B) + → foldl f e ⊥ ≡ e + foldl-⊥ {zero} _ _ = ≡.refl + foldl-⊥ {suc A} f e = foldl-⊥ (∣ f ⟩- suc) e + + foldl-⁅⁆ + : (f : B → Fin A → B) + (e : B) + (k : Fin A) + → foldl f e ⁅ k ⁆ ≡ f e k + foldl-⁅⁆ f e zero = foldl-⊥ f (f e zero) + foldl-⁅⁆ f e (suc k) = foldl-⁅⁆ (∣ f ⟩- suc) e k + + foldl-fusion + : (h : C → B) + {f : C → Fin A → C} + {g : B → Fin A → B} + → (∀ x n → h (f x n) ≡ g (h x) n) + → (e : C) + → h ∘ foldl f e ≗ foldl g (h e) + foldl-fusion {C = C} {A = A} h {f} {g} fuse e = V.foldl-fusion h ≡.refl fuse′ + where + open ≡.≡-Reasoning + fuse′ + : {k : Fin A} + (acc : C) + (b : Bool) + → h (if b then f acc k else acc) ≡ (if b then g (h acc) k else h acc) + fuse′ {k} acc b = begin + h (if b then f acc k else acc) ≡⟨ if-float h b ⟩ + (if b then h (f acc k) else h acc) ≡⟨ ≡.cong (if b then_else h acc) (fuse acc k) ⟩ + (if b then g (h acc) k else h acc) ∎ + +Subset0≗⊥ : (S : Subset 0) → S ≗ ⊥ +Subset0≗⊥ S () + +foldl-[] : (f : B → Fin 0 → B) (e : B) (S : Subset 0) → foldl f e S ≡ e +foldl-[] f e S = ≡.trans (foldl-cong₂ f e (Subset0≗⊥ S)) (foldl-⊥ f e) diff --git a/Data/System.agda b/Data/System.agda new file mode 100644 index 0000000..d71c2a8 --- /dev/null +++ b/Data/System.agda @@ -0,0 +1,144 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; 0ℓ; suc) + +module Data.System {ℓ : Level} where + +import Relation.Binary.Properties.Preorder as PreorderProperties +import Relation.Binary.Reasoning.Setoid as ≈-Reasoning + +open import Categories.Category using (Category) +open import Data.Circuit.Value using (Monoid) +open import Data.Nat using (ℕ) +open import Data.Setoid using (_⇒ₛ_; ∣_∣) +open import Data.Setoid.Unit using (⊤ₛ) +open import Data.System.Values Monoid using (Values; _≋_; module ≋; <ε>) +open import Function using (Func; _⟨$⟩_; flip) +open import Function.Construct.Constant using () renaming (function to Const) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (_∙_) +open import Level using (Level; 0ℓ; suc) +open import Relation.Binary as Rel using (Reflexive; Transitive; Preorder; Setoid; Rel) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) + +open Func + +module _ (n : ℕ) where + + record System : Set₁ where + + field + S : Setoid 0ℓ 0ℓ + fₛ : ∣ Values n ⇒ₛ S ⇒ₛ S ∣ + fₒ : ∣ S ⇒ₛ Values n ∣ + + fₛ′ : ∣ Values n ∣ → ∣ S ∣ → ∣ S ∣ + fₛ′ i = to (to fₛ i) + + fₒ′ : ∣ S ∣ → ∣ Values n ∣ + fₒ′ = to fₒ + + module S = Setoid S + + open System + + discrete : System + discrete .S = ⊤ₛ + discrete .fₛ = Const (Values n) (⊤ₛ ⇒ₛ ⊤ₛ) (Id ⊤ₛ) + discrete .fₒ = Const ⊤ₛ (Values n) <ε> + +module _ {n : ℕ} where + + record _≤_ (a b : System n) : Set ℓ where + + private + module A = System a + module B = System b + + open B using (S) + + field + ⇒S : ∣ A.S ⇒ₛ B.S ∣ + ≗-fₛ : (i : ∣ Values n ∣) (s : ∣ A.S ∣) → ⇒S ⟨$⟩ (A.fₛ′ i s) S.≈ B.fₛ′ i (⇒S ⟨$⟩ s) + ≗-fₒ : (s : ∣ A.S ∣) → A.fₒ′ s ≋ B.fₒ′ (⇒S ⟨$⟩ s) + + infix 4 _≤_ + +open System + +private + + module _ {n : ℕ} where + + open _≤_ + + ≤-refl : Reflexive (_≤_ {n}) + ⇒S ≤-refl = Id _ + ≗-fₛ (≤-refl {x}) _ _ = S.refl x + ≗-fₒ ≤-refl _ = ≋.refl + + ≡⇒≤ : _≡_ Rel.⇒ _≤_ + ≡⇒≤ ≡.refl = ≤-refl + + ≤-trans : Transitive _≤_ + ⇒S (≤-trans a b) = ⇒S b ∙ ⇒S a + ≗-fₛ (≤-trans {x} {y} {z} a b) i s = let open ≈-Reasoning (S z) in begin + ⇒S b ⟨$⟩ (⇒S a ⟨$⟩ (fₛ′ x i s)) ≈⟨ cong (⇒S b) (≗-fₛ a i s) ⟩ + ⇒S b ⟨$⟩ (fₛ′ y i (⇒S a ⟨$⟩ s)) ≈⟨ ≗-fₛ b i (⇒S a ⟨$⟩ s) ⟩ + fₛ′ z i (⇒S b ⟨$⟩ (⇒S a ⟨$⟩ s)) ∎ + ≗-fₒ (≤-trans {x} {y} {z} a b) s = let open ≈-Reasoning (Values n) in begin + fₒ′ x s ≈⟨ ≗-fₒ a s ⟩ + fₒ′ y (⇒S a ⟨$⟩ s) ≈⟨ ≗-fₒ b (⇒S a ⟨$⟩ s) ⟩ + fₒ′ z (⇒S b ⟨$⟩ (⇒S a ⟨$⟩ s)) ∎ + + variable + A B C : System n + + _≈_ : Rel (A ≤ B) 0ℓ + _≈_ {A} {B} ≤₁ ≤₂ = ⇒S ≤₁ A⇒B.≈ ⇒S ≤₂ + where + module A⇒B = Setoid (S A ⇒ₛ S B) + + open Rel.IsEquivalence + + ≈-isEquiv : Rel.IsEquivalence (_≈_ {A} {B}) + ≈-isEquiv {B = B} .refl = S.refl B + ≈-isEquiv {B = B} .sym a = S.sym B a + ≈-isEquiv {B = B} .trans a b = S.trans B a b + + ≤-resp-≈ : {f h : B ≤ C} {g i : A ≤ B} → f ≈ h → g ≈ i → ≤-trans g f ≈ ≤-trans i h + ≤-resp-≈ {_} {C} {_} {f} {h} {g} {i} f≈h g≈i {x} = begin + ⇒S f ⟨$⟩ (⇒S g ⟨$⟩ x) ≈⟨ f≈h ⟩ + ⇒S h ⟨$⟩ (⇒S g ⟨$⟩ x) ≈⟨ cong (⇒S h) g≈i ⟩ + ⇒S h ⟨$⟩ (⇒S i ⟨$⟩ x) ∎ + where + open ≈-Reasoning (System.S C) + +System-≤ : ℕ → Preorder (suc 0ℓ) (suc 0ℓ) ℓ +System-≤ n = record + { _≲_ = _≤_ {n} + ; isPreorder = record + { isEquivalence = ≡.isEquivalence + ; reflexive = ≡⇒≤ + ; trans = ≤-trans + } + } + +Systemₛ : ℕ → Setoid (suc 0ℓ) ℓ +Systemₛ n = PreorderProperties.InducedEquivalence (System-≤ n) + +Systems : ℕ → Category (suc 0ℓ) ℓ 0ℓ +Systems n = record + { Obj = System n + ; _⇒_ = _≤_ + ; _≈_ = _≈_ + ; id = ≤-refl + ; _∘_ = flip ≤-trans + ; assoc = λ {D = D} → S.refl D + ; sym-assoc = λ {D = D} → S.refl D + ; identityˡ = λ {B = B} → S.refl B + ; identityʳ = λ {B = B} → S.refl B + ; identity² = λ {A = A} → S.refl A + ; equiv = ≈-isEquiv + ; ∘-resp-≈ = λ {f = f} {h} {g} {i} → ≤-resp-≈ {f = f} {h} {g} {i} + } diff --git a/Data/System/Values.agda b/Data/System/Values.agda new file mode 100644 index 0000000..545a835 --- /dev/null +++ b/Data/System/Values.agda @@ -0,0 +1,146 @@ +{-# OPTIONS --without-K --safe #-} + +open import Algebra.Bundles using (CommutativeMonoid) +open import Level using (0ℓ) + +module Data.System.Values (A : CommutativeMonoid 0ℓ 0ℓ) where + +import Algebra.Properties.CommutativeMonoid.Sum as Sum +import Data.Vec.Functional.Relation.Binary.Equality.Setoid as Pointwise +import Relation.Binary.PropositionalEquality as ≡ + +open import Data.Fin using (_↑ˡ_; _↑ʳ_; zero; suc; splitAt) +open import Data.Nat using (ℕ; _+_; suc) +open import Data.Product using (_,_) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid using (∣_∣) +open import Data.Sum using (inj₁; inj₂) +open import Data.Vec.Functional as Vec using (Vector; zipWith; replicate) +open import Function using (Func; _⟶ₛ_; _⟨$⟩_; _∘_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid; IsEquivalence) + +open CommutativeMonoid A renaming (Carrier to Value) +open Func +open Sum A using (sum) + +open Pointwise setoid using (≋-setoid; ≋-isEquivalence) + +opaque + + Values : ℕ → Setoid 0ℓ 0ℓ + Values = ≋-setoid + +module _ {n : ℕ} where + + opaque + + unfolding Values + + merge : ∣ Values n ∣ → Value + merge = sum + + _⊕_ : ∣ Values n ∣ → ∣ Values n ∣ → ∣ Values n ∣ + xs ⊕ ys = zipWith _∙_ xs ys + + <ε> : ∣ Values n ∣ + <ε> = replicate n ε + + head : ∣ Values (suc n) ∣ → Value + head xs = xs zero + + tail : ∣ Values (suc n) ∣ → ∣ Values n ∣ + tail xs = xs ∘ suc + + module ≋ = Setoid (Values n) + + _≋_ : Rel ∣ Values n ∣ 0ℓ + _≋_ = ≋._≈_ + + infix 4 _≋_ + +opaque + + unfolding Values + open Setoid + ≋-isEquiv : ∀ n → IsEquivalence (_≈_ (Values n)) + ≋-isEquiv = ≋-isEquivalence + +module _ {n : ℕ} where + + opaque + unfolding _⊕_ + + ⊕-cong : {x y u v : ≋.Carrier {n}} → x ≋ y → u ≋ v → x ⊕ u ≋ y ⊕ v + ⊕-cong x≋y u≋v i = ∙-cong (x≋y i) (u≋v i) + + ⊕-assoc : (x y z : ≋.Carrier {n}) → (x ⊕ y) ⊕ z ≋ x ⊕ (y ⊕ z) + ⊕-assoc x y z i = assoc (x i) (y i) (z i) + + ⊕-identityˡ : (x : ≋.Carrier {n}) → <ε> ⊕ x ≋ x + ⊕-identityˡ x i = identityˡ (x i) + + ⊕-identityʳ : (x : ≋.Carrier {n}) → x ⊕ <ε> ≋ x + ⊕-identityʳ x i = identityʳ (x i) + + ⊕-comm : (x y : ≋.Carrier {n}) → x ⊕ y ≋ y ⊕ x + ⊕-comm x y i = comm (x i) (y i) + +open CommutativeMonoid +Valuesₘ : ℕ → CommutativeMonoid 0ℓ 0ℓ +Valuesₘ n .Carrier = ∣ Values n ∣ +Valuesₘ n ._≈_ = _≋_ +Valuesₘ n ._∙_ = _⊕_ +Valuesₘ n .ε = <ε> +Valuesₘ n .isCommutativeMonoid = record + { isMonoid = record + { isSemigroup = record + { isMagma = record + { isEquivalence = ≋-isEquiv n + ; ∙-cong = ⊕-cong + } + ; assoc = ⊕-assoc + } + ; identity = ⊕-identityˡ , ⊕-identityʳ + } + ; comm = ⊕-comm + } + +opaque + + unfolding Values + + [] : ∣ Values 0 ∣ + [] = Vec.[] + + []-unique : (xs ys : ∣ Values 0 ∣) → xs ≋ ys + []-unique xs ys () + +module _ {n m : ℕ} where + + opaque + + unfolding Values + + _++_ : ∣ Values n ∣ → ∣ Values m ∣ → ∣ Values (n + m) ∣ + _++_ = Vec._++_ + + infixr 5 _++_ + + ++-cong + : (xs xs′ : ∣ Values n ∣) + {ys ys′ : ∣ Values m ∣} + → xs ≋ xs′ + → ys ≋ ys′ + → xs ++ ys ≋ xs′ ++ ys′ + ++-cong xs xs′ xs≋xs′ ys≋ys′ i with splitAt n i + ... | inj₁ i = xs≋xs′ i + ... | inj₂ i = ys≋ys′ i + + splitₛ : Values (n + m) ⟶ₛ Values n ×ₛ Values m + to splitₛ v = v ∘ (_↑ˡ m) , v ∘ (n ↑ʳ_) + cong splitₛ v₁≋v₂ = v₁≋v₂ ∘ (_↑ˡ m) , v₁≋v₂ ∘ (n ↑ʳ_) + + ++ₛ : Values n ×ₛ Values m ⟶ₛ Values (n + m) + to ++ₛ (xs , ys) = xs ++ ys + cong ++ₛ (≗xs , ≗ys) = ++-cong _ _ ≗xs ≗ys diff --git a/Data/Vector.agda b/Data/Vector.agda new file mode 100644 index 0000000..874f0b2 --- /dev/null +++ b/Data/Vector.agda @@ -0,0 +1,126 @@ +{-# OPTIONS --without-K --safe #-} + +module Data.Vector where + +open import Data.Nat.Base using (ℕ) +open import Data.Vec.Functional using (Vector; head; tail; []; removeAt; map; _++_) public +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; module ≡-Reasoning) +open import Function.Base using (∣_⟩-_; _∘_) +open import Data.Fin.Base using (Fin; toℕ; _↑ˡ_; _↑ʳ_) +open ℕ +open Fin + +foldl + : ∀ {n : ℕ} {A : Set} (B : ℕ → Set) + → (∀ {k : Fin n} → B (toℕ k) → A → B (suc (toℕ k))) + → B zero + → Vector A n + → B n +foldl {zero} B ⊕ e v = e +foldl {suc n} B ⊕ e v = foldl (B ∘ suc) ⊕ (⊕ e (head v)) (tail v) + +foldl-cong + : {n : ℕ} {A : Set} + (B : ℕ → Set) + {f g : ∀ {k : Fin n} → B (toℕ k) → A → B (suc (toℕ k))} + → (∀ {k} → ∀ x y → f {k} x y ≡ g {k} x y) + → (e : B zero) + → (v : Vector A n) + → foldl B f e v ≡ foldl B g e v +foldl-cong {zero} B f≗g e v = ≡.refl +foldl-cong {suc n} B {g = g} f≗g e v rewrite (f≗g e (head v)) = foldl-cong (B ∘ suc) f≗g (g e (head v)) (tail v) + +foldl-cong-arg + : {n : ℕ} {A : Set} + (B : ℕ → Set) + (f : ∀ {k : Fin n} → B (toℕ k) → A → B (suc (toℕ k))) + → (e : B zero) + → {v w : Vector A n} + → v ≗ w + → foldl B f e v ≡ foldl B f e w +foldl-cong-arg {zero} B f e v≗w = ≡.refl +foldl-cong-arg {suc n} B f e {w = w} v≗w rewrite v≗w zero = foldl-cong-arg (B ∘ suc) f (f e (head w)) (v≗w ∘ suc) + +foldl-map + : {n : ℕ} {A : ℕ → Set} {B C : Set} + (f : ∀ {k : Fin n} → A (toℕ k) → B → A (suc (toℕ k))) + (g : C → B) + (x : A zero) + (xs : Vector C n) + → foldl A f x (map g xs) + ≡ foldl A (∣ f ⟩- g) x xs +foldl-map {zero} f g e xs = ≡.refl +foldl-map {suc n} f g e xs = foldl-map f g (f e (g (head xs))) (tail xs) + +foldl-fusion + : {n : ℕ} + {A : Set} {B C : ℕ → Set} + (h : {k : ℕ} → B k → C k) + → {f : {k : Fin n} → B (toℕ k) → A → B (suc (toℕ k))} {d : B zero} + → {g : {k : Fin n} → C (toℕ k) → A → C (suc (toℕ k))} {e : C zero} + → (h d ≡ e) + → ({k : Fin n} (b : B (toℕ k)) (x : A) → h (f {k} b x) ≡ g (h b) x) + → h ∘ foldl B f d ≗ foldl C g e +foldl-fusion {zero} _ base _ _ = base +foldl-fusion {suc n} {A} h {f} {d} {g} {e} base fuse xs = foldl-fusion h eq fuse (tail xs) + where + x₀ : A + x₀ = head xs + open ≡.≡-Reasoning + eq : h (f d x₀) ≡ g e x₀ + eq = begin + h (f d x₀) ≡⟨ fuse d x₀ ⟩ + g (h d) x₀ ≡⟨ ≡.cong-app (≡.cong g base) x₀ ⟩ + g e x₀ ∎ + +foldl-[] + : {A : Set} + (B : ℕ → Set) + (f : {k : Fin 0} → B (toℕ k) → A → B (suc (toℕ k))) + {e : B 0} + → foldl B f e [] ≡ e +foldl-[] _ _ = ≡.refl + +open import Data.Sum using ([_,_]′) +open import Data.Sum.Properties using ([,-]-cong; [-,]-cong; [,]-∘) +open import Data.Fin.Properties using (splitAt-↑ˡ; splitAt-↑ʳ) +open import Data.Fin using (splitAt) +open import Data.Nat using (_+_) +++-↑ˡ + : {n m : ℕ} + {A : Set} + (X : Vector A n) + (Y : Vector A m) + → (X ++ Y) ∘ (_↑ˡ m) ≗ X +++-↑ˡ {n} {m} X Y i = ≡.cong [ X , Y ]′ (splitAt-↑ˡ n i m) + +++-↑ʳ + : {n m : ℕ} + {A : Set} + (X : Vector A n) + (Y : Vector A m) + → (X ++ Y) ∘ (n ↑ʳ_) ≗ Y +++-↑ʳ {n} {m} X Y i = ≡.cong [ X , Y ]′ (splitAt-↑ʳ n m i) + ++-assocʳ : {m n o : ℕ} → Fin (m + (n + o)) → Fin (m + n + o) ++-assocʳ {m} {n} {o} = [ (λ x → x ↑ˡ n ↑ˡ o) , [ (λ x → (m ↑ʳ x) ↑ˡ o) , m + n ↑ʳ_ ]′ ∘ splitAt n ]′ ∘ splitAt m + +open ≡-Reasoning +++-assoc + : {m n o : ℕ} + {A : Set} + (X : Vector A m) + (Y : Vector A n) + (Z : Vector A o) + → ((X ++ Y) ++ Z) ∘ +-assocʳ {m} ≗ X ++ (Y ++ Z) +++-assoc {m} {n} {o} X Y Z i = begin + ((X ++ Y) ++ Z) (+-assocʳ {m} i) ≡⟨⟩ + ((X ++ Y) ++ Z) ([ (λ x → x ↑ˡ n ↑ˡ o) , _ ]′ (splitAt m i)) ≡⟨ [,]-∘ ((X ++ Y) ++ Z) (splitAt m i) ⟩ + [ ((X ++ Y) ++ Z) ∘ (λ x → x ↑ˡ n ↑ˡ o) , _ ]′ (splitAt m i) ≡⟨ [-,]-cong (++-↑ˡ (X ++ Y) Z ∘ (_↑ˡ n)) (splitAt m i) ⟩ + [ (X ++ Y) ∘ (_↑ˡ n) , _ ]′ (splitAt m i) ≡⟨ [-,]-cong (++-↑ˡ X Y) (splitAt m i) ⟩ + [ X , ((X ++ Y) ++ Z) ∘ [ _ , _ ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,-]-cong ([,]-∘ ((X ++ Y) ++ Z) ∘ splitAt n) (splitAt m i) ⟩ + [ X , [ (_ ++ Z) ∘ (_↑ˡ o) ∘ (m ↑ʳ_) , _ ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,-]-cong ([-,]-cong (++-↑ˡ (X ++ Y) Z ∘ (m ↑ʳ_)) ∘ splitAt n) (splitAt m i) ⟩ + [ X , [ (X ++ Y) ∘ (m ↑ʳ_) , _ ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,-]-cong ([-,]-cong (++-↑ʳ X Y) ∘ splitAt n) (splitAt m i) ⟩ + [ X , [ Y , ((X ++ Y) ++ Z) ∘ (m + n ↑ʳ_) ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,-]-cong ([,-]-cong (++-↑ʳ (X ++ Y) Z) ∘ splitAt n) (splitAt m i) ⟩ + [ X , [ Y , Z ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨⟩ + (X ++ (Y ++ Z)) i ∎ diff --git a/DecorationFunctor/Graph.agda b/DecorationFunctor/Graph.agda index 7f05855..8b62430 100644 --- a/DecorationFunctor/Graph.agda +++ b/DecorationFunctor/Graph.agda @@ -4,62 +4,48 @@ module DecorationFunctor.Graph where import Categories.Morphism as Morphism +open import Level using (0ℓ) + open import Categories.Category.BinaryProducts using (module BinaryProducts) open import Categories.Category.Cartesian using (Cartesian) -open import Categories.Category.Cocartesian using (Cocartesian; module BinaryCoproducts) -open import Categories.Category.Core using (Category) -open import Categories.Category.Instance.Nat using (Nat-Cocartesian) open import Categories.Category.Instance.Nat using (Nat) open import Categories.Category.Instance.Setoids using (Setoids) -open import Categories.Category.Instance.SingletonSet using (SingletonSetoid) open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) open import Categories.Category.Product using (_⁂_) open import Categories.Functor using () renaming (_∘F_ to _∘′_) open import Categories.Functor.Core using (Functor) open import Categories.Functor.Monoidal.Symmetric using (module Lax) -open import Categories.NaturalTransformation using (NaturalTransformation) - +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) -open import Category.Instance.Setoids.SymmetricMonoidal using (Setoids-×) open import Category.Instance.Nat.FinitelyCocomplete using (Nat-FinitelyCocomplete) - +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×; ×-symmetric′) open import Data.Empty using (⊥-elim) -open import Data.Fin using (#_) -open import Data.Fin.Base using (Fin; splitAt; join; zero; suc; _↑ˡ_; _↑ʳ_) +open import Data.Fin using (#_; Fin; splitAt; join; zero; suc; _↑ˡ_; _↑ʳ_) open import Data.Fin.Patterns using (0F; 1F) -open import Data.Fin.Properties using (splitAt-join; join-splitAt) +open import Data.Fin.Properties using (splitAt-join; join-splitAt; splitAt-↑ˡ; splitAt⁻¹-↑ˡ) open import Data.Nat using (ℕ; _+_) open import Data.Product.Base using (_,_) open import Data.Product.Relation.Binary.Pointwise.NonDependent using (×-setoid) +open import Data.Setoid.Unit {0ℓ} {0ℓ} using (⊤ₛ) open import Data.Sum.Base using (_⊎_; map; inj₁; inj₂; swap) renaming ([_,_]′ to [_,_]) open import Data.Sum.Properties using (map-map; [,]-map; [,]-∘; [-,]-cong; [,-]-cong; map-cong; swap-involutive) open import Data.Unit using (tt) -open import Data.Unit.Properties using () renaming (≡-setoid to ⊤-setoid) - -open import Function.Base using (_∘_; id; const; case_of_) -open import Function.Bundles using (Func; Inverse; _↔_; mk↔) +open import Function using (_∘_; id; const; Func; Inverse; _↔_; mk↔) open import Function.Construct.Composition using (_↔-∘_) +open import Function.Construct.Constant using () renaming (function to Const) open import Function.Construct.Identity using (↔-id) open import Function.Construct.Symmetry using (↔-sym) - -open import Level using (0ℓ; lift) - -open import Relation.Binary.Bundles using (Setoid) -open import Relation.Binary.PropositionalEquality using (_≗_) -open import Relation.Binary.PropositionalEquality.Core using (_≡_; erefl; refl; sym; trans; cong) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality using (_≗_; _≡_; erefl; refl; sym; trans; cong) open import Relation.Binary.PropositionalEquality.Properties using (isEquivalence; module ≡-Reasoning) open import Relation.Nullary.Negation.Core using (¬_) open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) -open Cocartesian Nat-Cocartesian using (coproducts) open FinitelyCocompleteCategory Nat-FinitelyCocomplete - using () - renaming (symmetricMonoidalCategory to Nat-smc) -open Morphism (Setoids 0ℓ 0ℓ) using (_≅_) + using (-+-) + renaming (symmetricMonoidalCategory to Nat-smc; +-assoc to Nat-+-assoc) open Lax using (SymmetricMonoidalFunctor) - open BinaryProducts products using (-×-) -open BinaryCoproducts coproducts using (-+-) renaming (+-assoc to Nat-+-assoc) record Graph (v : ℕ) : Set where @@ -164,50 +150,33 @@ Graph-Func f = record F-resp-≈ : {f g : Fin n → Fin m} → (∀ (x : Fin n) → f x ≡ g x) - → Graph-same G G′ - → Graph-same (map-nodes f G) (map-nodes g G′) -F-resp-≈ {g = g} f≗g ≡G = record - { ↔e = ↔e - ; ≗s = λ { x → trans (f≗g (s x)) (cong g (≗s x)) } - ; ≗t = λ { x → trans (f≗g (t x)) (cong g (≗t x)) } + → Graph-same (map-nodes f G) (map-nodes g G) +F-resp-≈ {G = G} f≗g = record + { ↔e = ↔-id _ + ; ≗s = f≗g ∘ s + ; ≗t = f≗g ∘ t } where - open Graph-same ≡G + open Graph G F : Functor Nat (Setoids 0ℓ 0ℓ) F = record { F₀ = Graph-setoid ; F₁ = Graph-Func - ; identity = id - ; homomorphism = λ { {_} {_} {_} {f} {g} → homomorphism f g } + ; identity = Graph-same-refl + ; homomorphism = Graph-same-refl ; F-resp-≈ = λ f≗g → F-resp-≈ f≗g } - where - homomorphism - : (f : Fin n → Fin m) - → (g : Fin m → Fin o) - → Graph-same G G′ - → Graph-same (map-nodes (g ∘ f) G) (map-nodes g (map-nodes f G′)) - homomorphism f g ≡G = record - { ↔e = ↔e - ; ≗s = cong (g ∘ f) ∘ ≗s - ; ≗t = cong (g ∘ f) ∘ ≗t - } - where - open Graph-same ≡G -empty-graph : Graph 0 -empty-graph = record +discrete : {n : ℕ} → Graph n +discrete = record { e = 0 ; s = λ () ; t = λ () } -ε : Func (SingletonSetoid {0ℓ} {0ℓ}) (Graph-setoid 0) -ε = record - { to = const empty-graph - ; cong = const Graph-same-refl - } +ε : Func ⊤ₛ (Graph-setoid 0) +ε = Const ⊤ₛ (Graph-setoid 0) discrete together : Graph n → Graph m → Graph (n + m) together {n} {m} G₁ G₂ = record @@ -313,77 +282,67 @@ together-resp-same {n} {m} ≡G₁ ≡G₂ = record commute : ∀ {n n′ m m′} - → {G₁ G₁′ : Graph n} - → {G₂ G₂′ : Graph m} + → {G₁ : Graph n} + → {G₂ : Graph m} → (f : Fin n → Fin n′) → (g : Fin m → Fin m′) - → Graph-same G₁ G₁′ - → Graph-same G₂ G₂′ → Graph-same (together (map-nodes f G₁) (map-nodes g G₂)) - (map-nodes ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n) (together G₁′ G₂′)) -commute {n} {n′} {m} {m′} f g ≡G₁ ≡G₂ = record - { ↔e = +-resp-↔ ≡G₁.↔e ≡G₂.↔e + (map-nodes ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n) (together G₁ G₂)) +commute {n} {n′} {m} {m′} {G₁} {G₂} f g = record + { ↔e = ↔e ; ≗s = source-commute ; ≗t = target-commute } where - ≡fG₁ : Graph-same (map-nodes f _) (map-nodes f _) - ≡fG₁ = F-resp-≈ (erefl ∘ f) ≡G₁ - ≡gG₂ : Graph-same (map-nodes g _) (map-nodes g _) - ≡gG₂ = F-resp-≈ (erefl ∘ g) ≡G₂ - module ≡G₁ = Graph-same ≡G₁ - module ≡G₂ = Graph-same ≡G₂ - module ≡fG₁ = Graph-same ≡fG₁ - module ≡fG₂ = Graph-same ≡gG₂ - module ≡G₁+G₂ = Graph-same (together-resp-same ≡G₁ ≡G₂) - module ≡fG₁+gG₂ = Graph-same (together-resp-same ≡fG₁ ≡gG₂) + open Graph-same (Graph-same-refl {_} {together G₁ G₂}) + module G₁ = Graph G₁ + module G₂ = Graph G₂ + module fG₁ = Graph (map-nodes f G₁) + module gG₂ = Graph (map-nodes g G₂) + module G₁+G₂ = Graph (together G₁ G₂) + module fG₁+gG₂ = Graph (together (map-nodes f G₁) (map-nodes g G₂)) open ≡-Reasoning source-commute - : ≡fG₁+gG₂.s + : fG₁+gG₂.s ≗ [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n - ∘ ≡G₁+G₂.s′ - ∘ ≡fG₁+gG₂.to + ∘ G₁+G₂.s + ∘ to source-commute x = begin - ≡fG₁+gG₂.s x - ≡⟨ ≡fG₁+gG₂.≗s x ⟩ - (≡fG₁+gG₂.s′ ∘ ≡fG₁+gG₂.to) x + fG₁+gG₂.s x ≡⟨⟩ - (join n′ m′ ∘ map ≡fG₁.s′ ≡fG₂.s′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ cong (join n′ m′) (map-map ((splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x)) ⟨ - (join n′ m′ ∘ map f g ∘ map ≡G₁.s′ ≡G₂.s′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ [,]-map ((map ≡G₁.s′ ≡G₂.s′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x) ⟩ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ map ≡G₁.s′ ≡G₂.s′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ cong [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] (splitAt-join n m (map ≡G₁.s′ ≡G₂.s′ (splitAt ≡G₁.e′ (≡fG₁+gG₂.to x)))) ⟨ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ join n m ∘ map ≡G₁.s′ ≡G₂.s′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x + (join n′ m′ ∘ map fG₁.s gG₂.s ∘ splitAt G₁.e ∘ to) x + ≡⟨ cong (join n′ m′) (map-map ((splitAt G₁.e ∘ to) x)) ⟨ + (join n′ m′ ∘ map f g ∘ map G₁.s G₂.s ∘ splitAt fG₁.e ∘ to) x + ≡⟨ [,]-map ((map G₁.s G₂.s ∘ splitAt fG₁.e ∘ to) x) ⟩ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ map G₁.s G₂.s ∘ splitAt fG₁.e ∘ to) x + ≡⟨ cong [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] (splitAt-join n m (map G₁.s G₂.s (splitAt fG₁.e (to x)))) ⟨ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ join n m ∘ map G₁.s G₂.s ∘ splitAt fG₁.e ∘ to) x ≡⟨⟩ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ ≡G₁+G₂.s′ ∘ ≡fG₁+gG₂.to) x ∎ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ G₁+G₂.s ∘ to) x ∎ target-commute - : ≡fG₁+gG₂.t + : fG₁+gG₂.t ≗ [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n - ∘ ≡G₁+G₂.t′ - ∘ ≡fG₁+gG₂.to + ∘ G₁+G₂.t + ∘ to target-commute x = begin - ≡fG₁+gG₂.t x - ≡⟨ ≡fG₁+gG₂.≗t x ⟩ - (≡fG₁+gG₂.t′ ∘ ≡fG₁+gG₂.to) x + fG₁+gG₂.t x ≡⟨⟩ - (join n′ m′ ∘ map ≡fG₁.t′ ≡fG₂.t′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ cong (join n′ m′) (map-map ((splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x)) ⟨ - (join n′ m′ ∘ map f g ∘ map ≡G₁.t′ ≡G₂.t′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ [,]-map ((map ≡G₁.t′ ≡G₂.t′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x) ⟩ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ map ≡G₁.t′ ≡G₂.t′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x - ≡⟨ cong [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] (splitAt-join n m (map ≡G₁.t′ ≡G₂.t′ (splitAt ≡G₁.e′ (≡fG₁+gG₂.to x)))) ⟨ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ join n m ∘ map ≡G₁.t′ ≡G₂.t′ ∘ splitAt ≡G₁.e′ ∘ ≡fG₁+gG₂.to) x + (join n′ m′ ∘ map fG₁.t gG₂.t ∘ splitAt G₁.e ∘ to) x + ≡⟨ cong (join n′ m′) (map-map ((splitAt G₁.e ∘ to) x)) ⟨ + (join n′ m′ ∘ map f g ∘ map G₁.t G₂.t ∘ splitAt fG₁.e ∘ to) x + ≡⟨ [,]-map ((map G₁.t G₂.t ∘ splitAt fG₁.e ∘ to) x) ⟩ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ map G₁.t G₂.t ∘ splitAt fG₁.e ∘ to) x + ≡⟨ cong [ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] (splitAt-join n m (map G₁.t G₂.t (splitAt fG₁.e (to x)))) ⟨ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ join n m ∘ map G₁.t G₂.t ∘ splitAt fG₁.e ∘ to) x ≡⟨⟩ - ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ ≡G₁+G₂.t′ ∘ ≡fG₁+gG₂.to) x ∎ + ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n ∘ G₁+G₂.t ∘ to) x ∎ ⊗-homomorphism : NaturalTransformation (-×- ∘′ (F ⁂ F)) (F ∘′ -+-) -⊗-homomorphism = record +⊗-homomorphism = ntHelper record { η = λ { (n , m) → η {n} {m} } - ; commute = λ { (f , g) (≡G₁ , ≡G₂) → commute f g ≡G₁ ≡G₂ } - ; sym-commute = λ { (f , g) (≡G₁ , ≡G₂) → Graph-same-sym (commute f g (Graph-same-sym ≡G₁) (Graph-same-sym ≡G₂)) } + ; commute = λ { (f , g) {G₁ , G₂} → commute {G₁ = G₁} {G₂ = G₂} f g } } where η : Func (×-setoid (Graph-setoid n) (Graph-setoid m)) (Graph-setoid (n + m)) @@ -406,155 +365,150 @@ commute {n} {n′} {m} {m′} f g ≡G₁ ≡G₂ = record associativity : {X Y Z : ℕ} - → {G₁ G₁′ : Graph X} - → {G₂ G₂′ : Graph Y} - → {G₃ G₃′ : Graph Z} - → Graph-same G₁ G₁′ - → Graph-same G₂ G₂′ - → Graph-same G₃ G₃′ + → (G₁ : Graph X) + → (G₂ : Graph Y) + → (G₃ : Graph Z) → Graph-same (map-nodes (Inverse.to (+-assoc-↔ X Y Z)) (together (together G₁ G₂) G₃)) - (together G₁′ (together G₂′ G₃′)) -associativity {X} {Y} {Z} ≡G₁ ≡G₂ ≡G₃ = record + (together G₁ (together G₂ G₃)) +associativity {X} {Y} {Z} G₁ G₂ G₃ = record { ↔e = ↔e ; ≗s = ≗s ; ≗t = ≗t } where - module ≡G₁ = Graph-same ≡G₁ - module ≡G₂ = Graph-same ≡G₂ - module ≡G₃ = Graph-same ≡G₃ - module ≡G₂+G₃ = Graph-same (together-resp-same ≡G₂ ≡G₃) - module ≡G₁+[G₂+G₃] = Graph-same (together-resp-same ≡G₁ (together-resp-same ≡G₂ ≡G₃)) - module ≡G₁+G₂+G₃ = Graph-same (together-resp-same (together-resp-same ≡G₁ ≡G₂) ≡G₃) - ↔e : Fin (≡G₁.e + ≡G₂.e + ≡G₃.e) ↔ Fin (≡G₁.e′ + (≡G₂.e′ + ≡G₃.e′)) - ↔e = +-resp-↔ ≡G₁.↔e (+-resp-↔ ≡G₂.↔e ≡G₃.↔e) ↔-∘ (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e) + module G₁ = Graph G₁ + module G₂ = Graph G₂ + module G₃ = Graph G₃ + module G₂+G₃ = Graph (together G₂ G₃) + module G₁+[G₂+G₃] = Graph (together G₁ (together G₂ G₃)) + module G₁+G₂+G₃ = Graph (together (together G₁ G₂) G₃) + ↔e : Fin (G₁.e + G₂.e + G₃.e) ↔ Fin (G₁.e + (G₂.e + G₃.e)) + ↔e = +-assoc-↔ G₁.e G₂.e G₃.e open ≡-Reasoning open Inverse - ≗s : to (+-assoc-↔ X Y Z) ∘ ≡G₁+G₂+G₃.s ≗ ≡G₁+[G₂+G₃].s′ ∘ ≡G₁+[G₂+G₃].to ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e) + ≗s : to (+-assoc-↔ X Y Z) ∘ G₁+G₂+G₃.s ≗ G₁+[G₂+G₃].s ∘ to ↔e ≗s x = begin - (to (+-assoc-↔ X Y Z) ∘ ≡G₁+G₂+G₃.s) x ≡⟨⟩ - ([ [ join X (Y + Z) ∘ inj₁ , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt X , _ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.s) x - ≡⟨ [-,]-cong ([,]-∘ (join X (Y + Z)) ∘ splitAt X) (splitAt (X + Y) (≡G₁+G₂+G₃.s x)) ⟨ - ([ join X (Y + Z) ∘ map id _ ∘ splitAt X , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.s) x - ≡⟨ [,]-∘ (join X (Y + Z)) (splitAt (X + Y) (≡G₁+G₂+G₃.s x)) ⟨ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.s) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ join (X + Y) Z ∘ map _ ≡G₃.s ∘ splitAt _) x + (to (+-assoc-↔ X Y Z) ∘ G₁+G₂+G₃.s) x ≡⟨⟩ + ([ [ join X (Y + Z) ∘ inj₁ , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt X , _ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.s) x + ≡⟨ [-,]-cong ([,]-∘ (join X (Y + Z)) ∘ splitAt X) (splitAt (X + Y) (G₁+G₂+G₃.s x)) ⟨ + ([ join X (Y + Z) ∘ map id _ ∘ splitAt X , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.s) x + ≡⟨ [,]-∘ (join X (Y + Z)) (splitAt (X + Y) (G₁+G₂+G₃.s x)) ⟨ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.s) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ join (X + Y) Z ∘ map _ G₃.s ∘ splitAt _) x ≡⟨ cong (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ]) - (splitAt-join (X + Y) Z (map _ ≡G₃.s (splitAt _ x))) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ map _ ≡G₃.s ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([,]-map (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X ∘ join X Y ∘ map ≡G₁.s ≡G₂.s ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.s ] ∘ splitAt _) x + (splitAt-join (X + Y) Z (map _ G₃.s (splitAt _ x))) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ map _ G₃.s ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([,]-map (splitAt (G₁.e + G₂.e) x)) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X ∘ join X Y ∘ map G₁.s G₂.s ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.s ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (cong (map id (_↑ˡ Z)) ∘ splitAt-join X Y ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ map ≡G₁.s ≡G₂.s ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.s ] ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([-,]-cong (map-map ∘ splitAt ≡G₁.e) (splitAt _ x)) ⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s (join Y Z ∘ inj₁ ∘ ≡G₂.s) ∘ splitAt _ , inj₂ ∘ _ ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s (join Y Z ∘ map ≡G₂.s ≡G₃.s ∘ inj₁) ∘ splitAt _ , _ ] ∘ splitAt _) x + (cong (map id (_↑ˡ Z)) ∘ splitAt-join X Y ∘ map G₁.s G₂.s ∘ splitAt G₁.e) + (splitAt (G₁.e + G₂.e) x)) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ map G₁.s G₂.s ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.s ] ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([-,]-cong (map-map ∘ splitAt G₁.e) (splitAt _ x)) ⟩ + (join X (Y + Z) ∘ [ map G₁.s (join Y Z ∘ inj₁ ∘ G₂.s) ∘ splitAt _ , inj₂ ∘ _ ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.s (join Y Z ∘ map G₂.s G₃.s ∘ inj₁) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (map-cong (cong ≡G₁.s ∘ erefl) (cong (join Y Z ∘ map ≡G₂.s ≡G₃.s) ∘ splitAt-join ≡G₂.e ≡G₃.e ∘ inj₁) ∘ splitAt _) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.s (join Y Z ∘ map ≡G₂.s ≡G₃.s ∘ splitAt ≡G₂.e ∘ _) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s (≡G₂+G₃.s ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.s ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s (≡G₂+G₃.s ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map ≡G₂.s ≡G₃.s ∘ inj₂ ] ∘ splitAt _) x + (map-cong (cong G₁.s ∘ erefl) (cong (join Y Z ∘ map G₂.s G₃.s) ∘ splitAt-join G₂.e G₃.e ∘ inj₁) ∘ splitAt _) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.s (join Y Z ∘ map G₂.s G₃.s ∘ splitAt G₂.e ∘ _) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.s (G₂+G₃.s ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.s ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.s (G₂+G₃.s ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map G₂.s G₃.s ∘ inj₂ ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([,-]-cong - (cong (inj₂ ∘ join Y Z ∘ map ≡G₂.s ≡G₃.s) ∘ splitAt-join ≡G₂.e ≡G₃.e ∘ inj₂) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.s _ ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map ≡G₂.s ≡G₃.s ∘ splitAt ≡G₂.e ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s _ ∘ splitAt _ , inj₂ ∘ ≡G₂+G₃.s ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.s _ ∘ splitAt _ , map ≡G₁.s ≡G₂+G₃.s ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x + (cong (inj₂ ∘ join Y Z ∘ map G₂.s G₃.s) ∘ splitAt-join G₂.e G₃.e ∘ inj₂) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.s _ ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map G₂.s G₃.s ∘ splitAt G₂.e ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.s _ ∘ splitAt _ , inj₂ ∘ G₂+G₃.s ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.s _ ∘ splitAt _ , map G₁.s G₂+G₃.s ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (map-map ∘ splitAt ≡G₁.e) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.s ≡G₂+G₃.s ∘ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , map ≡G₁.s ≡G₂+G₃.s ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([,]-∘ (map ≡G₁.s ≡G₂+G₃.s) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ map ≡G₁.s ≡G₂+G₃.s ∘ [ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x + (map-map ∘ splitAt G₁.e) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.s G₂+G₃.s ∘ map id (_↑ˡ G₃.e) ∘ splitAt _ , map G₁.s G₂+G₃.s ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([,]-∘ (map G₁.s G₂+G₃.s) (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ map G₁.s G₂+G₃.s ∘ [ map id (_↑ˡ G₃.e) ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨ cong - (join X (Y + Z) ∘ map ≡G₁.s ≡G₂+G₃.s) - (splitAt-join ≡G₁.e ≡G₂+G₃.e (([ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x)) ⟨ - (join X (Y + Z) ∘ map ≡G₁.s ≡G₂+G₃.s ∘ splitAt ≡G₁.e ∘ join ≡G₁.e ≡G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (≡G₁+[G₂+G₃].s ∘ join ≡G₁.e ≡G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong ≡G₁+[G₂+G₃].s ([,]-∘ (join ≡G₁.e ≡G₂+G₃.e) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (≡G₁+[G₂+G₃].s ∘ [ join ≡G₁.e ≡G₂+G₃.e ∘ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , join ≡G₁.e ≡G₂+G₃.e ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong ≡G₁+[G₂+G₃].s ([-,]-cong ([,]-map ∘ splitAt ≡G₁.e) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (≡G₁+[G₂+G₃].s ∘ [ [ _↑ˡ ≡G₂.e + ≡G₃.e , (≡G₁.e ↑ʳ_) ∘ (_↑ˡ ≡G₃.e) ] ∘ splitAt ≡G₁.e , (≡G₁.e ↑ʳ_) ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (≡G₁+[G₂+G₃].s ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e)) x ≡⟨ ≡G₁+[G₂+G₃].≗s (to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e) x) ⟩ - (≡G₁+[G₂+G₃].s′ ∘ ≡G₁+[G₂+G₃].to ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e)) x ∎ - ≗t : to (+-assoc-↔ X Y Z) ∘ ≡G₁+G₂+G₃.t ≗ ≡G₁+[G₂+G₃].t′ ∘ ≡G₁+[G₂+G₃].to ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e) + (join X (Y + Z) ∘ map G₁.s G₂+G₃.s) + (splitAt-join G₁.e G₂+G₃.e (([ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x)) ⟨ + (join X (Y + Z) ∘ map G₁.s G₂+G₃.s ∘ splitAt G₁.e ∘ join G₁.e G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (G₁+[G₂+G₃].s ∘ join G₁.e G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong G₁+[G₂+G₃].s ([,]-∘ (join G₁.e G₂+G₃.e) (splitAt (G₁.e + G₂.e) x)) ⟩ + (G₁+[G₂+G₃].s ∘ [ join G₁.e G₂+G₃.e ∘ map id (_↑ˡ G₃.e) ∘ splitAt _ , join G₁.e G₂+G₃.e ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong G₁+[G₂+G₃].s ([-,]-cong ([,]-map ∘ splitAt G₁.e) (splitAt (G₁.e + G₂.e) x)) ⟩ + (G₁+[G₂+G₃].s ∘ [ [ _↑ˡ G₂.e + G₃.e , (G₁.e ↑ʳ_) ∘ (_↑ˡ G₃.e) ] ∘ splitAt G₁.e , (G₁.e ↑ʳ_) ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (G₁+[G₂+G₃].s ∘ to (+-assoc-↔ G₁.e G₂.e G₃.e)) x ∎ + ≗t : to (+-assoc-↔ X Y Z) ∘ G₁+G₂+G₃.t ≗ G₁+[G₂+G₃].t ∘ to ↔e ≗t x = begin - (to (+-assoc-↔ X Y Z) ∘ ≡G₁+G₂+G₃.t) x ≡⟨⟩ - ([ [ join X (Y + Z) ∘ inj₁ , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt X , _ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.t) x - ≡⟨ [-,]-cong ([,]-∘ (join X (Y + Z)) ∘ splitAt X) (splitAt (X + Y) (≡G₁+G₂+G₃.t x)) ⟨ - ([ join X (Y + Z) ∘ map id _ ∘ splitAt X , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.t) x - ≡⟨ [,]-∘ (join X (Y + Z)) (splitAt (X + Y) (≡G₁+G₂+G₃.t x)) ⟨ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ ≡G₁+G₂+G₃.t) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ join (X + Y) Z ∘ map _ ≡G₃.t ∘ splitAt _) x + (to (+-assoc-↔ X Y Z) ∘ G₁+G₂+G₃.t) x ≡⟨⟩ + ([ [ join X (Y + Z) ∘ inj₁ , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt X , _ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.t) x + ≡⟨ [-,]-cong ([,]-∘ (join X (Y + Z)) ∘ splitAt X) (splitAt (X + Y) (G₁+G₂+G₃.t x)) ⟨ + ([ join X (Y + Z) ∘ map id _ ∘ splitAt X , join X (Y + Z) ∘ inj₂ ∘ _ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.t) x + ≡⟨ [,]-∘ (join X (Y + Z)) (splitAt (X + Y) (G₁+G₂+G₃.t x)) ⟨ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ G₁+G₂+G₃.t) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ splitAt (X + Y) ∘ join (X + Y) Z ∘ map _ G₃.t ∘ splitAt _) x ≡⟨ cong (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ]) - (splitAt-join (X + Y) Z (map _ ≡G₃.t (splitAt _ x))) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ map _ ≡G₃.t ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([,]-map (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ splitAt X ∘ join X Y ∘ map ≡G₁.t ≡G₂.t ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.t ] ∘ splitAt _) x + (splitAt-join (X + Y) Z (map _ G₃.t (splitAt _ x))) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X , inj₂ ∘ join Y Z ∘ inj₂ ] ∘ map _ G₃.t ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([,]-map (splitAt (G₁.e + G₂.e) x)) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ splitAt X ∘ join X Y ∘ map G₁.t G₂.t ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.t ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (cong (map id (_↑ˡ Z)) ∘ splitAt-join X Y ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (join X (Y + Z) ∘ [ map id _ ∘ map ≡G₁.t ≡G₂.t ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.t ] ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([-,]-cong (map-map ∘ splitAt ≡G₁.e) (splitAt _ x)) ⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t (join Y Z ∘ inj₁ ∘ ≡G₂.t) ∘ splitAt _ , inj₂ ∘ _ ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t (join Y Z ∘ map ≡G₂.t ≡G₃.t ∘ inj₁) ∘ splitAt _ , _ ] ∘ splitAt _) x + (cong (map id (_↑ˡ Z)) ∘ splitAt-join X Y ∘ map G₁.t G₂.t ∘ splitAt G₁.e) + (splitAt (G₁.e + G₂.e) x)) ⟩ + (join X (Y + Z) ∘ [ map id _ ∘ map G₁.t G₂.t ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.t ] ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([-,]-cong (map-map ∘ splitAt G₁.e) (splitAt _ x)) ⟩ + (join X (Y + Z) ∘ [ map G₁.t (join Y Z ∘ inj₁ ∘ G₂.t) ∘ splitAt _ , inj₂ ∘ _ ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.t (join Y Z ∘ map G₂.t G₃.t ∘ inj₁) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (map-cong (cong ≡G₁.t ∘ erefl) (cong (join Y Z ∘ map ≡G₂.t ≡G₃.t) ∘ splitAt-join ≡G₂.e ≡G₃.e ∘ inj₁) ∘ splitAt _) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.t (join Y Z ∘ map ≡G₂.t ≡G₃.t ∘ splitAt ≡G₂.e ∘ _) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t (≡G₂+G₃.t ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ ≡G₃.t ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t (≡G₂+G₃.t ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map ≡G₂.t ≡G₃.t ∘ inj₂ ] ∘ splitAt _) x + (map-cong (cong G₁.t ∘ erefl) (cong (join Y Z ∘ map G₂.t G₃.t) ∘ splitAt-join G₂.e G₃.e ∘ inj₁) ∘ splitAt _) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.t (join Y Z ∘ map G₂.t G₃.t ∘ splitAt G₂.e ∘ _) ∘ splitAt _ , _ ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.t (G₂+G₃.t ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ inj₂ ∘ G₃.t ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.t (G₂+G₃.t ∘ _) ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map G₂.t G₃.t ∘ inj₂ ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([,-]-cong - (cong (inj₂ ∘ join Y Z ∘ map ≡G₂.t ≡G₃.t) ∘ splitAt-join ≡G₂.e ≡G₃.e ∘ inj₂) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.t _ ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map ≡G₂.t ≡G₃.t ∘ splitAt ≡G₂.e ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t _ ∘ splitAt _ , inj₂ ∘ ≡G₂+G₃.t ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (join X (Y + Z) ∘ [ map ≡G₁.t _ ∘ splitAt _ , map ≡G₁.t ≡G₂+G₃.t ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x + (cong (inj₂ ∘ join Y Z ∘ map G₂.t G₃.t) ∘ splitAt-join G₂.e G₃.e ∘ inj₂) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.t _ ∘ splitAt _ , inj₂ ∘ join Y Z ∘ map G₂.t G₃.t ∘ splitAt G₂.e ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.t _ ∘ splitAt _ , inj₂ ∘ G₂+G₃.t ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (join X (Y + Z) ∘ [ map G₁.t _ ∘ splitAt _ , map G₁.t G₂+G₃.t ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨ cong (join X (Y + Z)) ([-,]-cong - (map-map ∘ splitAt ≡G₁.e) - (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ [ map ≡G₁.t ≡G₂+G₃.t ∘ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , map ≡G₁.t ≡G₂+G₃.t ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong (join X (Y + Z)) ([,]-∘ (map ≡G₁.t ≡G₂+G₃.t) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟨ - (join X (Y + Z) ∘ map ≡G₁.t ≡G₂+G₃.t ∘ [ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x + (map-map ∘ splitAt G₁.e) + (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ [ map G₁.t G₂+G₃.t ∘ map id (_↑ˡ G₃.e) ∘ splitAt _ , map G₁.t G₂+G₃.t ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong (join X (Y + Z)) ([,]-∘ (map G₁.t G₂+G₃.t) (splitAt (G₁.e + G₂.e) x)) ⟨ + (join X (Y + Z) ∘ map G₁.t G₂+G₃.t ∘ [ map id (_↑ˡ G₃.e) ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨ cong - (join X (Y + Z) ∘ map ≡G₁.t ≡G₂+G₃.t) - (splitAt-join ≡G₁.e ≡G₂+G₃.e (([ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x)) ⟨ - (join X (Y + Z) ∘ map ≡G₁.t ≡G₂+G₃.t ∘ splitAt ≡G₁.e ∘ join ≡G₁.e ≡G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (≡G₁+[G₂+G₃].t ∘ join ≡G₁.e ≡G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong ≡G₁+[G₂+G₃].t ([,]-∘ (join ≡G₁.e ≡G₂+G₃.e) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (≡G₁+[G₂+G₃].t ∘ [ join ≡G₁.e ≡G₂+G₃.e ∘ map id (_↑ˡ ≡G₃.e) ∘ splitAt _ , join ≡G₁.e ≡G₂+G₃.e ∘ inj₂ ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x - ≡⟨ cong ≡G₁+[G₂+G₃].t ([-,]-cong ([,]-map ∘ splitAt ≡G₁.e) (splitAt (≡G₁.e + ≡G₂.e) x)) ⟩ - (≡G₁+[G₂+G₃].t ∘ [ [ _↑ˡ ≡G₂.e + ≡G₃.e , (≡G₁.e ↑ʳ_) ∘ (_↑ˡ ≡G₃.e) ] ∘ splitAt ≡G₁.e , (≡G₁.e ↑ʳ_) ∘ (≡G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ - (≡G₁+[G₂+G₃].t ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e)) x ≡⟨ ≡G₁+[G₂+G₃].≗t (to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e) x) ⟩ - (≡G₁+[G₂+G₃].t′ ∘ ≡G₁+[G₂+G₃].to ∘ to (+-assoc-↔ ≡G₁.e ≡G₂.e ≡G₃.e)) x ∎ - -unitaryˡ : Graph-same G G′ → Graph-same (together empty-graph G) G′ -unitaryˡ ≡G = ≡G - -n+0↔0 : ∀ n → Fin (n + 0) ↔ Fin n -n+0↔0 n = record + (join X (Y + Z) ∘ map G₁.t G₂+G₃.t) + (splitAt-join G₁.e G₂+G₃.e (([ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x)) ⟨ + (join X (Y + Z) ∘ map G₁.t G₂+G₃.t ∘ splitAt G₁.e ∘ join G₁.e G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (G₁+[G₂+G₃].t ∘ join G₁.e G₂+G₃.e ∘ [ map id _ ∘ splitAt _ , inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong G₁+[G₂+G₃].t ([,]-∘ (join G₁.e G₂+G₃.e) (splitAt (G₁.e + G₂.e) x)) ⟩ + (G₁+[G₂+G₃].t ∘ [ join G₁.e G₂+G₃.e ∘ map id (_↑ˡ G₃.e) ∘ splitAt _ , join G₁.e G₂+G₃.e ∘ inj₂ ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x + ≡⟨ cong G₁+[G₂+G₃].t ([-,]-cong ([,]-map ∘ splitAt G₁.e) (splitAt (G₁.e + G₂.e) x)) ⟩ + (G₁+[G₂+G₃].t ∘ [ [ _↑ˡ G₂.e + G₃.e , (G₁.e ↑ʳ_) ∘ (_↑ˡ G₃.e) ] ∘ splitAt G₁.e , (G₁.e ↑ʳ_) ∘ (G₂.e ↑ʳ_) ] ∘ splitAt _) x ≡⟨⟩ + (G₁+[G₂+G₃].t ∘ to (+-assoc-↔ G₁.e G₂.e G₃.e)) x ∎ + +unitaryˡ : Graph-same (together (discrete {0}) G) G +unitaryˡ = Graph-same-refl + +n+0↔n : ∀ n → Fin (n + 0) ↔ Fin n +n+0↔n n = record { to = to ; from = from ; to-cong = λ { refl → refl } @@ -562,63 +516,33 @@ n+0↔0 n = record ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } } where - to : ∀ {n} → Fin (n + 0) → Fin n - to {ℕ.suc ℕ.zero} n = n - to {ℕ.suc (ℕ.suc n)} zero = zero - to {ℕ.suc (ℕ.suc n)} (suc z) = suc (to z) - from : ∀ {n} → Fin n → Fin (n + 0) - from {ℕ.suc ℕ.zero} n = n - from {ℕ.suc (ℕ.suc n)} zero = zero - from {ℕ.suc (ℕ.suc n)} (suc z) = suc (from z) - from∘to : ∀ {n} → ∀ (x : Fin (n + 0)) → from (to x) ≡ x - from∘to {ℕ.suc ℕ.zero} zero = refl - from∘to {ℕ.suc (ℕ.suc n)} zero = refl - from∘to {ℕ.suc (ℕ.suc n)} (suc x) = cong suc (from∘to x) - to∘from : ∀ {n} → ∀ (x : Fin n) → to (from x) ≡ x - to∘from {ℕ.suc ℕ.zero} zero = refl - to∘from {ℕ.suc (ℕ.suc n)} zero = refl - to∘from {ℕ.suc (ℕ.suc n)} (suc x) = cong suc (to∘from x) + to : Fin (n + 0) → Fin n + to x with inj₁ x₁ ← splitAt n x = x₁ + from : Fin n → Fin (n + 0) + from x = x ↑ˡ 0 + from∘to : (x : Fin (n + 0)) → from (to x) ≡ x + from∘to x with inj₁ x₁ ← splitAt n x in eq = splitAt⁻¹-↑ˡ eq + to∘from : (x : Fin n) → to (from x) ≡ x + to∘from x rewrite splitAt-↑ˡ n x 0 = refl unitaryʳ - : {G G′ : Graph n} - → Graph-same G G′ - → Graph-same (map-nodes ([ (λ x → x) , (λ ()) ] ∘ splitAt n) (together G empty-graph)) G′ -unitaryʳ {n} {G} {G′} ≡G = record - { ↔e = e+0↔e′ + : {G : Graph n} + → Graph-same (map-nodes ([ (λ x → x) , (λ ()) ] ∘ splitAt n) (together G discrete)) G +unitaryʳ {n} {G} = record + { ↔e = e+0↔e ; ≗s = ≗s+0 ; ≗t = ≗t+0 } where - open Graph-same ≡G + open Graph G open ≡-Reasoning - e+0↔e′ : Fin (e + 0) ↔ Fin e′ - e+0↔e′ = ↔e ↔-∘ n+0↔0 e - module e+0↔e′ = Inverse e+0↔e′ - open Inverse - ↑ˡ-0 : ∀ e → (x : Fin e) → x ≡ to (n+0↔0 e) (x ↑ˡ 0) - ↑ˡ-0 (ℕ.suc ℕ.zero) zero = refl - ↑ˡ-0 (ℕ.suc (ℕ.suc e)) zero = refl - ↑ˡ-0 (ℕ.suc (ℕ.suc e)) (suc x) = cong suc (↑ˡ-0 (ℕ.suc e) x) - ≗s+0 : [ id , (λ ()) ] ∘ splitAt n ∘ join n 0 ∘ map s (λ ()) ∘ splitAt e ≗ s′ ∘ e+0↔e′.to - ≗s+0 x+0 with splitAt e x+0 in eq - ... | inj₁ x = begin - [ id , (λ ()) ] (splitAt n (join n 0 (inj₁ (s x)))) ≡⟨ cong [ id , (λ ()) ] (splitAt-join n 0 (inj₁ (s x))) ⟩ - s x ≡⟨ cong s (↑ˡ-0 e x) ⟩ - s (to (n+0↔0 e) (x ↑ˡ 0)) ≡⟨⟩ - s (to (n+0↔0 e) (join e 0 (inj₁ x))) ≡⟨ cong (s ∘ to (n+0↔0 e) ∘ join e 0) eq ⟨ - s (to (n+0↔0 e) (join e 0 (splitAt e x+0))) ≡⟨ cong (s ∘ to (n+0↔0 e)) (join-splitAt e 0 x+0) ⟩ - s (to (n+0↔0 e) x+0) ≡⟨ ≗s (to (n+0↔0 e) x+0) ⟩ - s′ (e+0↔e′.to x+0) ∎ - ≗t+0 : [ id , (λ ()) ] ∘ splitAt n ∘ join n 0 ∘ map t (λ ()) ∘ splitAt e ≗ t′ ∘ e+0↔e′.to - ≗t+0 x+0 with splitAt e x+0 in eq - ... | inj₁ x = begin - [ id , (λ ()) ] (splitAt n (join n 0 (inj₁ (t x)))) ≡⟨ cong [ id , (λ ()) ] (splitAt-join n 0 (inj₁ (t x))) ⟩ - t x ≡⟨ cong t (↑ˡ-0 e x) ⟩ - t (to (n+0↔0 e) (x ↑ˡ 0)) ≡⟨⟩ - t (to (n+0↔0 e) (join e 0 (inj₁ x))) ≡⟨ cong (t ∘ to (n+0↔0 e) ∘ join e 0) eq ⟨ - t (to (n+0↔0 e) (join e 0 (splitAt e x+0))) ≡⟨ cong (t ∘ to (n+0↔0 e)) (join-splitAt e 0 x+0) ⟩ - t (to (n+0↔0 e) x+0) ≡⟨ ≗t (to (n+0↔0 e) x+0) ⟩ - t′ (e+0↔e′.to x+0) ∎ + e+0↔e : Fin (e + 0) ↔ Fin e + e+0↔e = n+0↔n e + module e+0↔e = Inverse e+0↔e + ≗s+0 : [ id , (λ ()) ] ∘ splitAt n ∘ join n 0 ∘ map s (λ ()) ∘ splitAt e ≗ s ∘ e+0↔e.to + ≗s+0 x+0 with inj₁ x ← splitAt e x+0 = cong [ id , (λ ()) ] (splitAt-↑ˡ n (s x) 0) + ≗t+0 : [ id , (λ ()) ] ∘ splitAt n ∘ join n 0 ∘ map t (λ ()) ∘ splitAt e ≗ t ∘ e+0↔e.to + ≗t+0 x+0 with inj₁ x ← splitAt e x+0 = cong [ id , (λ ()) ] (splitAt-↑ˡ n (t x) 0) +-comm-↔ : ∀ (n m : ℕ) → Fin (n + m) ↔ Fin (m + n) +-comm-↔ n m = record @@ -655,85 +579,68 @@ join-swap (inj₁ x) = refl join-swap (inj₂ y) = refl braiding - : {G₁ G₁′ : Graph n} - → {G₂ G₂′ : Graph m} - → Graph-same G₁ G₁′ - → Graph-same G₂ G₂′ - → Graph-same (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together G₁ G₂)) (together G₂′ G₁′) -braiding {n} {m} ≡G₁ ≡G₂ = record - { ↔e = +-comm-↔ ≡G₁.e′ ≡G₂.e′ ↔-∘ +-resp-↔ ≡G₁.↔e ≡G₂.↔e + : {G₁ : Graph n} + → {G₂ : Graph m} + → Graph-same (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together G₁ G₂)) (together G₂ G₁) +braiding {n} {m} {G₁} {G₂} = record + { ↔e = +-comm-↔ G₁.e G₂.e ; ≗s = ≗s ; ≗t = ≗t } where open ≡-Reasoning - module ≡G₁ = Graph-same ≡G₁ - module ≡G₂ = Graph-same ≡G₂ + module G₁ = Graph G₁ + module G₂ = Graph G₂ ≗s : [ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n - ∘ join n m ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e - ≗ join m n ∘ map ≡G₂.s′ ≡G₁.s′ ∘ splitAt ≡G₂.e′ - ∘ join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ - ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e + ∘ join n m ∘ map G₁.s G₂.s ∘ splitAt G₁.e + ≗ join m n ∘ map G₂.s G₁.s ∘ splitAt G₂.e + ∘ Inverse.to (+-comm-↔ G₁.e G₂.e) ≗s x = begin - ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n ∘ join n m ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) x - ≡⟨ (join-swap ∘ splitAt n ∘ join n m ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ swap ∘ map ≡G₁.s ≡G₂.s ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ map-cong ≡G₁.≗s ≡G₂.≗s ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ swap ∘ map (≡G₁.s′ ∘ ≡G₁.to) (≡G₂.s′ ∘ ≡G₂.to) ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ map-map ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ swap ∘ map ≡G₁.s′ ≡G₂.s′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n) ∘ swap-map ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ map ≡G₂.s′ ≡G₁.s′ ∘ swap ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ map ≡G₂.s′ ≡G₁.s′ ∘ swap) ∘ splitAt-join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ map ≡G₂.s′ ≡G₁.s′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ map ≡G₂.s′ ≡G₁.s′) ∘ splitAt-join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ map ≡G₂.s′ ≡G₁.s′ ∘ splitAt ≡G₂.e′ ∘ join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ∎ + ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n ∘ join n m ∘ map G₁.s G₂.s ∘ splitAt G₁.e) x + ≡⟨ (join-swap ∘ splitAt n ∘ join n m ∘ map G₁.s G₂.s ∘ splitAt G₁.e) x ⟨ + (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ map G₁.s G₂.s ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ map G₁.s G₂.s ∘ splitAt G₁.e) x ⟩ + (join m n ∘ swap ∘ map G₁.s G₂.s ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n) ∘ swap-map ∘ splitAt G₁.e) x ⟩ + (join m n ∘ map G₂.s G₁.s ∘ swap ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n ∘ map G₂.s G₁.s) ∘ splitAt-join G₂.e G₁.e ∘ swap ∘ splitAt G₁.e) x ⟨ + (join m n ∘ map G₂.s G₁.s ∘ splitAt G₂.e ∘ join G₂.e G₁.e ∘ swap ∘ splitAt G₁.e) x ∎ ≗t : [ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n - ∘ join n m ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e - ≗ join m n ∘ map ≡G₂.t′ ≡G₁.t′ ∘ splitAt ≡G₂.e′ - ∘ join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ - ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e + ∘ join n m ∘ map G₁.t G₂.t ∘ splitAt G₁.e + ≗ join m n ∘ map G₂.t G₁.t ∘ splitAt G₂.e + ∘ Inverse.to (+-comm-↔ G₁.e G₂.e) ≗t x = begin - ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n ∘ join n m ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) x - ≡⟨ (join-swap ∘ splitAt n ∘ join n m ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ swap ∘ map ≡G₁.t ≡G₂.t ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ map-cong ≡G₁.≗t ≡G₂.≗t ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ swap ∘ map (≡G₁.t′ ∘ ≡G₁.to) (≡G₂.t′ ∘ ≡G₂.to) ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ swap) ∘ map-map ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ swap ∘ map ≡G₁.t′ ≡G₂.t′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n) ∘ swap-map ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟩ - (join m n ∘ map ≡G₂.t′ ≡G₁.t′ ∘ swap ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ map ≡G₂.t′ ≡G₁.t′ ∘ swap) ∘ splitAt-join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ map ≡G₂.t′ ≡G₁.t′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x - ≡⟨ (cong (join m n ∘ map ≡G₂.t′ ≡G₁.t′) ∘ splitAt-join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ⟨ - (join m n ∘ map ≡G₂.t′ ≡G₁.t′ ∘ splitAt ≡G₂.e′ ∘ join ≡G₂.e′ ≡G₁.e′ ∘ swap ∘ splitAt ≡G₁.e′ ∘ join ≡G₁.e′ ≡G₂.e′ ∘ map ≡G₁.to ≡G₂.to ∘ splitAt ≡G₁.e) x ∎ - -graph : SymmetricMonoidalFunctor Nat-smc (Setoids-× {0ℓ}) -graph = record - { F = F - ; isBraidedMonoidal = record - { isMonoidal = record - { ε = ε - ; ⊗-homo = ⊗-homomorphism - ; associativity = λ { ((≡G₁ , ≡G₂) , ≡G₃) → associativity ≡G₁ ≡G₂ ≡G₃ } - ; unitaryˡ = λ { (lift tt , ≡G) → unitaryˡ ≡G } - ; unitaryʳ = λ { (≡G , lift tt) → unitaryʳ ≡G } - } - ; braiding-compat = λ { (≡G₁ , ≡G₂) → braiding ≡G₁ ≡G₂ } - } - } + ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n ∘ join n m ∘ map G₁.t G₂.t ∘ splitAt G₁.e) x + ≡⟨ (join-swap ∘ splitAt n ∘ join n m ∘ map G₁.t G₂.t ∘ splitAt G₁.e) x ⟨ + (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ map G₁.t G₂.t ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ map G₁.t G₂.t ∘ splitAt G₁.e) x ⟩ + (join m n ∘ swap ∘ map G₁.t G₂.t ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n) ∘ swap-map ∘ splitAt G₁.e) x ⟩ + (join m n ∘ map G₂.t G₁.t ∘ swap ∘ splitAt G₁.e) x + ≡⟨ (cong (join m n ∘ map G₂.t G₁.t) ∘ splitAt-join G₂.e G₁.e ∘ swap ∘ splitAt G₁.e) x ⟨ + (join m n ∘ map G₂.t G₁.t ∘ splitAt G₂.e ∘ join G₂.e G₁.e ∘ swap ∘ splitAt G₁.e) x ∎ + +opaque + unfolding ×-symmetric′ + graph : SymmetricMonoidalFunctor Nat-smc Setoids-× + graph = record + { F = F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {x} {y} {z} {(G₁ , G₂) , G₃} → associativity G₁ G₂ G₃ } + ; unitaryˡ = unitaryˡ + ; unitaryʳ = unitaryʳ + } + ; braiding-compat = λ { {x} {y} {G₁ , G₂} → braiding {G₁ = G₁} {G₂ = G₂} } + } + } module F = SymmetricMonoidalFunctor graph -and-gate : Func (SingletonSetoid {0ℓ} {0ℓ}) (F.₀ 3) -and-gate = record - { to = λ { (lift tt) → and-graph } - ; cong = λ { (lift tt) → Graph-same-refl } - } +and-gate : Func ⊤ₛ (Graph-setoid 3) +and-gate = Const ⊤ₛ (Graph-setoid 3) and-graph where and-graph : Graph 3 and-graph = record diff --git a/DecorationFunctor/Hypergraph.agda b/DecorationFunctor/Hypergraph.agda new file mode 100644 index 0000000..2f61bc3 --- /dev/null +++ b/DecorationFunctor/Hypergraph.agda @@ -0,0 +1,653 @@ +{-# OPTIONS --without-K --safe #-} + +module DecorationFunctor.Hypergraph where + +import Categories.Morphism as Morphism +open import Level using (0ℓ) + +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Core using (Category) +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using () renaming (_∘F_ to _∘′_) +open import Categories.Functor.Core using (Functor) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Category.Instance.Nat.FinitelyCocomplete using (Nat-FinitelyCocomplete) +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×; ×-symmetric′) +open import Data.Empty using (⊥-elim) +open import Data.Fin using (#_; Fin; splitAt; join; zero; suc; _↑ˡ_; _↑ʳ_; toℕ; cast) +open import Data.Fin.Patterns using (0F; 1F; 2F) +open import Data.Fin.Properties using (splitAt-join; join-splitAt; cast-is-id; cast-trans; toℕ-cast; subst-is-cast; splitAt-↑ˡ; splitAt-↑ʳ; splitAt⁻¹-↑ˡ; ↑ˡ-injective) +open import Data.Nat using (ℕ; _+_) +open import Data.Product.Base using (_,_; Σ) +open import Data.Setoid using (∣_∣) +open import Data.Setoid.Unit {0ℓ} {0ℓ} using (⊤ₛ) +open import Data.Sum using (_⊎_; map; inj₁; inj₂; swap; map₂) renaming ([_,_]′ to [_,_]) +open import Data.Sum.Properties using (map-map; [,]-map; [,]-∘; [-,]-cong; [,-]-cong; [,]-cong; map-cong; swap-involutive) +open import Data.Unit using (tt) +open import Function using (_∘_; id; const; Func; Inverse; _↔_; mk↔; _⟨$⟩_) +open import Function.Construct.Composition using (_↔-∘_) +open import Function.Construct.Constant using () renaming (function to Const) +open import Function.Construct.Identity using (↔-id) renaming (function to Id) +open import Function.Construct.Symmetry using (↔-sym) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality.Core using (_≗_; _≡_; erefl; refl; sym; trans; cong; cong₂; subst; cong-app) +open import Relation.Binary.PropositionalEquality.Properties using (isEquivalence; module ≡-Reasoning; dcong₂; subst-∘) +open import Relation.Nullary.Negation.Core using (¬_) + +open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) +open FinitelyCocompleteCategory Nat-FinitelyCocomplete + using (-+-; _+₁_) + renaming (symmetricMonoidalCategory to Nat-smc; +-assoc to Nat-+-assoc) +open Morphism (Setoids 0ℓ 0ℓ) using (_≅_) +open Lax using (SymmetricMonoidalFunctor) + +open BinaryProducts products using (-×-) + +record Hypergraph (v : ℕ) : Set where + + field + h : ℕ + a : Fin h → ℕ + + arity : Fin h → ℕ + arity = ℕ.suc ∘ a + + field + j : ∀ (e : Fin h) → Fin (arity e) → Fin v + +record Hypergraph-same {n : ℕ} (H H′ : Hypergraph n) : Set where + + open Hypergraph H public + open Hypergraph H′ renaming (h to h′; a to a′; arity to arity′; j to j′) public + + field + ↔h : Fin h ↔ Fin h′ + + open Inverse ↔h public + + field + ≗a : a ≗ a′ ∘ to + + ≗arity : arity ≗ arity′ ∘ to + ≗arity e = cong ℕ.suc (≗a e) + + field + ≗j : (e : Fin h) + (i : Fin (arity e)) + → j e i + ≡ j′ (to e) (cast (≗arity e) i) + +private + + variable + n n′ m m′ o : ℕ + H H′ H″ H₁ H₁′ : Hypergraph n + H₂ H₂′ : Hypergraph m + H₃ : Hypergraph o + +Hypergraph-same-refl : Hypergraph-same H H +Hypergraph-same-refl {_} {H} = record + { ↔h = ↔-id (Fin h) + ; ≗a = cong a ∘ erefl + ; ≗j = λ e i → cong (j e) (sym (cast-is-id refl i)) + } + where + open Hypergraph H + +Hypergraph-same-sym : Hypergraph-same H H′ → Hypergraph-same H′ H +Hypergraph-same-sym {V} {H} {H′} ≡H = record + { ↔h = ↔-sym ↔h + ; ≗a = sym ∘ a∘from≗a′ + ; ≗j = ≗j′ + } + where + open Hypergraph-same ≡H + open ≡-Reasoning + a∘from≗a′ : a ∘ from ≗ a′ + a∘from≗a′ x = begin + (a ∘ from) x ≡⟨ (≗a ∘ from) x ⟩ + (a′ ∘ to ∘ from) x ≡⟨ (cong a′ ∘ inverseˡ ∘ erefl ∘ from) x ⟩ + a′ x ∎ + id≗to∘from : id ≗ to ∘ from + id≗to∘from e = sym (inverseˡ refl) + ≗arity′ : arity′ ≗ arity ∘ from + ≗arity′ e = cong ℕ.suc (sym (a∘from≗a′ e)) + ≗arity- : arity′ ≗ arity′ ∘ to ∘ from + ≗arity- e = cong arity′ (id≗to∘from e) + ≗j′ : (e : Fin h′) (i : Fin (arity′ e)) → j′ e i ≡ j (from e) (cast (≗arity′ e) i) + ≗j′ e i = begin + j′ e i ≡⟨ dcong₂ j′ (id≗to∘from e) (subst-∘ (id≗to∘from e)) ⟩ + j′ (to (from e)) (subst Fin (cong arity′ (id≗to∘from e)) i) ≡⟨ cong (j′ (to (from e))) (subst-is-cast (cong arity′ (id≗to∘from e)) i) ⟩ + j′ (to (from e)) (cast (cong arity′ (id≗to∘from e)) i) ≡⟨⟩ + j′ (to (from e)) (cast (trans (≗arity′ e) (≗arity (from e))) i) ≡⟨ cong (j′ (to (from e))) (cast-trans (≗arity′ e) (≗arity (from e)) i) ⟨ + j′ (to (from e)) (cast (≗arity (from e)) (cast (≗arity′ e) i)) ≡⟨ ≗j (from e) (cast (≗arity′ e) i) ⟨ + j (from e) (cast (≗arity′ e) i) ∎ + +Hypergraph-same-trans : Hypergraph-same H H′ → Hypergraph-same H′ H″ → Hypergraph-same H H″ +Hypergraph-same-trans ≡H₁ ≡H₂ = record + { ↔h = ↔h ≡H₂ ↔-∘ ↔h ≡H₁ + ; ≗a = λ { x → trans (≗a ≡H₁ x) (≗a ≡H₂ (to (↔h ≡H₁) x)) } + ; ≗j = λ { e i → trans (≗j ≡H₁ e i) (≗j₂ e i) } + } + where + open Hypergraph-same + open Inverse + open ≡-Reasoning + ≗j₂ : (e : Fin (h ≡H₁)) + (i : Fin (arity ≡H₁ e)) + → j ≡H₂ (to (↔h ≡H₁) e) (cast (≗arity ≡H₁ e) i) + ≡ j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (trans (≗arity ≡H₁ e) (≗arity ≡H₂ (to (↔h ≡H₁) e))) i) + ≗j₂ e i = begin + j ≡H₂ (to (↔h ≡H₁) e) (cast (≗arity ≡H₁ e) i) + ≡⟨ ≗j ≡H₂ (to (↔h ≡H₁) e) (cast (≗arity ≡H₁ e) i) ⟩ + j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (≗arity ≡H₂ (to (↔h ≡H₁) e)) (cast (≗arity ≡H₁ e) i)) + ≡⟨ cong (j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e))) (cast-trans (≗arity ≡H₁ e) (≗arity ≡H₂ (to (↔h ≡H₁) e)) i) ⟩ + j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (trans (≗arity ≡H₁ e) (≗arity ≡H₂ (to (↔h ≡H₁) e))) i) ∎ + +Hypergraphₛ : ℕ → Setoid 0ℓ 0ℓ +Hypergraphₛ p = record + { Carrier = Hypergraph p + ; _≈_ = Hypergraph-same + ; isEquivalence = record + { refl = Hypergraph-same-refl + ; sym = Hypergraph-same-sym + ; trans = Hypergraph-same-trans + } + } + +map-nodes : (Fin n → Fin m) → Hypergraph n → Hypergraph m +map-nodes f H = record + { h = h + ; a = a + ; j = λ e i → f (j e i) + } + where + open Hypergraph H + +Hypergraph-same-cong + : (f : Fin n → Fin m) + → Hypergraph-same H H′ + → Hypergraph-same (map-nodes f H) (map-nodes f H′) +Hypergraph-same-cong f ≡H = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ e i → cong f (≗j e i) + } + where + open Hypergraph-same ≡H + +Hypergraph-Func : (Fin n → Fin m) → Func (Hypergraphₛ n) (Hypergraphₛ m) +Hypergraph-Func f = record + { to = map-nodes f + ; cong = Hypergraph-same-cong f + } + +F-resp-≈ + : {f g : Fin n → Fin m} + → f ≗ g + → Hypergraph-same (map-nodes f H) (map-nodes g H) +F-resp-≈ {g = g} f≗g = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ { e i → trans (f≗g (j e i)) (cong g (≗j e i)) } + } + where + open Hypergraph-same Hypergraph-same-refl + +homomorphism + : (f : Fin n → Fin m) + → (g : Fin m → Fin o) + → Hypergraph-same (map-nodes (g ∘ f) H) (map-nodes g (map-nodes f H)) +homomorphism {n} {m} {o} {H} f g = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ e i → cong (g ∘ f) (≗j e i) + } + where + open Hypergraph-same Hypergraph-same-refl + +F : Functor Nat (Setoids 0ℓ 0ℓ) +F = record + { F₀ = Hypergraphₛ + ; F₁ = Hypergraph-Func + ; identity = λ { {n} {H} → Hypergraph-same-refl {H = H} } + ; homomorphism = λ { {f = f} {g = g} → homomorphism f g } + ; F-resp-≈ = λ f≗g → F-resp-≈ f≗g + } + +-- monoidal structure + +discrete : {n : ℕ} → Hypergraph n +discrete {n} = record + { h = 0 + ; a = λ () + ; j = λ () + } + +opaque + unfolding ×-symmetric′ + + ε : Func Setoids-×.unit (Hypergraphₛ 0) + ε = Const ⊤ₛ (Hypergraphₛ 0) discrete + +module _ (H₁ : Hypergraph n) (H₂ : Hypergraph m) where + private + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + j+j : (e : Fin (H₁.h + H₂.h)) + → Fin (ℕ.suc ([ H₁.a , H₂.a ] (splitAt H₁.h e))) + → Fin (n + m) + j+j e i with splitAt H₁.h e + ... | inj₁ e₁ = H₁.j e₁ i ↑ˡ m + ... | inj₂ e₂ = n ↑ʳ H₂.j e₂ i + +together : Hypergraph n → Hypergraph m → Hypergraph (n + m) +together {n} {m} H₁ H₂ = record + { h = h H₁ + h H₂ + ; a = [ a H₁ , a H₂ ] ∘ splitAt (h H₁) + ; j = j+j H₁ H₂ + } + where + open Hypergraph + ++-resp-↔ + : {n n′ m m′ : ℕ} + → Fin n ↔ Fin n′ + → Fin m ↔ Fin m′ + → Fin (n + m) ↔ Fin (n′ + m′) ++-resp-↔ {n} {n′} {m} {m′} ↔n ↔m = record + { to = join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n + ; from = join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ + ; to-cong = cong (join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n) + ; from-cong = cong (join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′) + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + module ↔n = Inverse ↔n + module ↔m = Inverse ↔m + open ≡-Reasoning + to∘from : join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ∘ join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ≗ id + to∘from x = begin + (join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ∘ join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′) x + ≡⟨ cong + (join n′ m′ ∘ map ↔n.to ↔m.to) + (splitAt-join n m (map ↔n.from ↔m.from (splitAt n′ x))) ⟩ + (join n′ m′ ∘ map ↔n.to ↔m.to ∘ map ↔n.from ↔m.from ∘ splitAt n′) x + ≡⟨ cong (join n′ m′) (map-map (splitAt n′ x)) ⟩ + (join n′ m′ ∘ map (↔n.to ∘ ↔n.from) (↔m.to ∘ ↔m.from) ∘ splitAt n′) x + ≡⟨ cong + (join n′ m′) + (map-cong + (λ _ → ↔n.inverseˡ refl) + (λ _ → ↔m.inverseˡ refl) + (splitAt n′ x)) ⟩ + (join n′ m′ ∘ map id id ∘ splitAt n′) x ≡⟨ [,]-map (splitAt n′ x) ⟩ + (join n′ m′ ∘ splitAt n′) x ≡⟨ join-splitAt n′ m′ x ⟩ + x ∎ + from∘to : join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ∘ join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ≗ id + from∘to x = begin + (join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ∘ join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n) x + ≡⟨ cong + (join n m ∘ map ↔n.from ↔m.from) + (splitAt-join n′ m′ (map ↔n.to ↔m.to (splitAt n x))) ⟩ + (join n m ∘ map ↔n.from ↔m.from ∘ map ↔n.to ↔m.to ∘ splitAt n) x + ≡⟨ cong (join n m) (map-map (splitAt n x)) ⟩ + (join n m ∘ map (↔n.from ∘ ↔n.to) (↔m.from ∘ ↔m.to) ∘ splitAt n) x + ≡⟨ cong + (join n m) + (map-cong + (λ _ → ↔n.inverseʳ refl) + (λ _ → ↔m.inverseʳ refl) + (splitAt n x)) ⟩ + (join n m ∘ map id id ∘ splitAt n) x ≡⟨ [,]-map (splitAt n x) ⟩ + (join n m ∘ splitAt n) x ≡⟨ join-splitAt n m x ⟩ + x ∎ + +together-resp-same + : Hypergraph-same H₁ H₁′ + → Hypergraph-same H₂ H₂′ + → Hypergraph-same (together H₁ H₂) (together H₁′ H₂′) +together-resp-same {n} {H₁} {H₁′} {m} {H₂} {H₂′} ≡H₁ ≡H₂ = record + { ↔h = +-resp-↔ ≡H₁.↔h ≡H₂.↔h + ; ≗a = ≗a + ; ≗j = ≗j + } + where + module ≡H₁ = Hypergraph-same ≡H₁ + module ≡H₂ = Hypergraph-same ≡H₂ + module H₁+H₂ = Hypergraph (together H₁ H₂) + module H₁+H₂′ = Hypergraph (together H₁′ H₂′) + open ≡-Reasoning + open Inverse + open Hypergraph + ≗a : [ ≡H₁.a , ≡H₂.a ] ∘ splitAt ≡H₁.h + ≗ [ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt ≡H₁.h′ + ∘ join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h + ≗a e = begin + [ ≡H₁.a , ≡H₂.a ] (splitAt ≡H₁.h e) ≡⟨ [,]-cong ≡H₁.≗a ≡H₂.≗a (splitAt ≡H₁.h e) ⟩ + ([ ≡H₁.a′ ∘ ≡H₁.to , ≡H₂.a′ ∘ ≡H₂.to ] ∘ splitAt ≡H₁.h) e ≡⟨ [,]-map (splitAt ≡H₁.h e) ⟨ + ([ ≡H₁.a′ , ≡H₂.a′ ] ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ≡⟨ (cong [ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt-join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ⟨ + ([ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt ≡H₁.h′ ∘ join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ∎ + ≗arity : H₁+H₂.arity ≗ H₁+H₂′.arity ∘ join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h + ≗arity = cong ℕ.suc ∘ ≗a + ≗j : (e : Fin H₁+H₂.h) + (i : Fin (H₁+H₂.arity e)) + → H₁+H₂.j e i + ≡ H₁+H₂′.j (to (+-resp-↔ ≡H₁.↔h ≡H₂.↔h) e) (cast (cong ℕ.suc (≗a e)) i) + ≗j e i with splitAt ≡H₁.h e + ... | inj₁ e₁ rewrite splitAt-↑ˡ ≡H₁.h′ (≡H₁.to e₁) ≡H₂.h′ = cong (_↑ˡ m) (≡H₁.≗j e₁ i) + ... | inj₂ e₂ rewrite splitAt-↑ʳ ≡H₁.h′ ≡H₂.h′ (≡H₂.to e₂) = cong (n ↑ʳ_) (≡H₂.≗j e₂ i) + +commute + : (f : Fin n → Fin n′) + → (g : Fin m → Fin m′) + → Hypergraph-same + (together (map-nodes f H₁) (map-nodes g H₂)) + (map-nodes (f +₁ g) (together H₁ H₂)) +commute {n} {n′} {m} {m′} {H₁} {H₂} f g = record + { ↔h = ≡H₁+H₂.↔h + ; ≗a = ≡H₁+H₂.≗a + ; ≗j = ≗j + } + where + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + module H₁+H₂ = Hypergraph (together H₁ H₂) + module ≡H₁+H₂ = Hypergraph-same {H = together H₁ H₂} Hypergraph-same-refl + open Hypergraph + open ≡-Reasoning + ≗j : (e : Fin (H₁.h + H₂.h)) + (i : Fin ((ℕ.suc ∘ [ H₁.a , H₂.a ] ∘ splitAt H₁.h) e)) + → j (together (map-nodes f H₁) (map-nodes g H₂)) e i + ≡ j (map-nodes (f +₁ g) (together H₁ H₂)) (≡H₁+H₂.to e) (cast refl i) + ≗j e i with splitAt H₁.h e + ... | inj₁ e₁ rewrite splitAt-↑ˡ n (H₁.j e₁ (cast refl i)) m = cong ((_↑ˡ m′) ∘ f ∘ H₁.j e₁) (sym (cast-is-id refl i)) + ... | inj₂ e₂ rewrite splitAt-↑ʳ n m (H₂.j e₂ (cast refl i)) = cong ((n′ ↑ʳ_) ∘ g ∘ H₂.j e₂) (sym (cast-is-id refl i)) + +open Setoids-× using (_⊗₀_; _⊗₁_) +opaque + unfolding ×-symmetric′ + η : Func (Hypergraphₛ n ⊗₀ Hypergraphₛ m) (Hypergraphₛ (n + m)) + η = record + { to = λ (H₁ , H₂) → together H₁ H₂ + ; cong = λ (≡H₁ , ≡H₂) → together-resp-same ≡H₁ ≡H₂ + } + +opaque + unfolding η + commute′ + : (f : Fin n → Fin n′) + → (g : Fin m → Fin m′) + → {x : ∣ Hypergraphₛ n ⊗₀ Hypergraphₛ m ∣} + → Hypergraph-same + (η ⟨$⟩ (Hypergraph-Func f ⊗₁ Hypergraph-Func g ⟨$⟩ x)) + (map-nodes (f +₁ g) (η ⟨$⟩ x)) + commute′ f g {H₁ , H₂} = commute {H₁ = H₁} {H₂} f g + +⊗-homomorphism : NaturalTransformation (Setoids-×.⊗ ∘′ (F ⁂ F)) (F ∘′ -+-) +⊗-homomorphism = ntHelper record + { η = λ (n , m) → η {n} {m} + ; commute = λ (f , g) → commute′ f g + } + ++-assoc-↔ : ∀ (x y z : ℕ) → Fin (x + y + z) ↔ Fin (x + (y + z)) ++-assoc-↔ x y z = record + { to = to + ; from = from + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → isoˡ _ }) , λ { refl → isoʳ _ } + } + where + module Nat = Morphism Nat + open Nat._≅_ (Nat-+-assoc {x} {y} {z}) + +associativity + : {X Y Z : ℕ} + → (H₁ : Hypergraph X) + → (H₂ : Hypergraph Y) + → (H₃ : Hypergraph Z) + → Hypergraph-same + (map-nodes (Inverse.to (+-assoc-↔ X Y Z)) (together (together H₁ H₂) H₃)) + (together H₁ (together H₂ H₃)) +associativity {X} {Y} {Z} H₁ H₂ H₃ = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = ≗j + } + where + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + module H₃ = Hypergraph H₃ + open ≡-Reasoning + open Hypergraph + ↔h : Fin (H₁.h + H₂.h + H₃.h) ↔ Fin (H₁.h + (H₂.h + H₃.h)) + ↔h = +-assoc-↔ H₁.h H₂.h H₃.h + ≗a : (x : Fin (H₁.h + H₂.h + H₃.h)) + → [ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] (splitAt (H₁.h + H₂.h) x) + ≡ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] (splitAt H₁.h ([ [ _↑ˡ H₂.h + H₃.h , (H₁.h ↑ʳ_) ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] (splitAt (H₁.h + H₂.h) x))) + ≗a x = begin + ([ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ inj₁ ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [-,]-cong ([,-]-cong (cong [ H₂.a , H₃.a ] ∘ (λ i → splitAt-↑ˡ H₂.h i H₃.h)) ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [-,]-cong ([,]-∘ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₂.a , H₃.a ] ∘ inj₂ ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [,-]-cong (cong [ H₂.a , H₃.a ] ∘ splitAt-↑ʳ H₂.h H₃.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ inj₂ ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [,]-∘ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] (splitAt (H₁.h + H₂.h) x) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , inj₂ ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([,-]-cong (splitAt-↑ʳ H₁.h (H₂.h + H₃.h) ∘ (H₂.h ↑ʳ_)) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , splitAt H₁.h ∘ (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([-,]-cong (splitAt-join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ splitAt H₁.h ∘ join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , splitAt H₁.h ∘ (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([,]-∘ (splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ [ _↑ˡ H₂.h + H₃.h , H₁.h ↑ʳ_ ] ∘ [ inj₁ , inj₂ ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h) ([-,]-cong ([,]-∘ [ _↑ˡ H₂.h + H₃.h , H₁.h ↑ʳ_ ] ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟩ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ [ _↑ˡ H₂.h + H₃.h , (H₁.h ↑ʳ_) ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x ∎ + + ≗j : (e : Fin (H₁.h + H₂.h + H₃.h)) + (i : Fin (ℕ.suc ([ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] (splitAt (H₁.h + H₂.h) e)))) + → Inverse.to (+-assoc-↔ X Y Z) (j+j (together H₁ H₂) H₃ e i) + ≡ j+j H₁ (together H₂ H₃) (Inverse.to ↔h e) (cast (cong ℕ.suc (≗a e)) i) + ≗j e i with splitAt (H₁.h + H₂.h) e + ≗j e i | inj₁ e₁₂ with splitAt H₁.h e₁₂ + ≗j e i | inj₁ e₁₂ | inj₁ e₁ + rewrite splitAt-↑ˡ H₁.h e₁ (H₂.h + H₃.h) + rewrite splitAt-↑ˡ (X + Y) (H₁.j e₁ i ↑ˡ Y) Z + rewrite splitAt-↑ˡ X (H₁.j e₁ i) Y = cong ((_↑ˡ Y + Z) ∘ H₁.j e₁) (sym (cast-is-id refl i)) + ≗j e i | inj₁ e₁₂ | inj₂ e₂ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (e₂ ↑ˡ H₃.h) + rewrite splitAt-↑ˡ H₂.h e₂ H₃.h + rewrite splitAt-↑ˡ (X + Y) (X ↑ʳ H₂.j e₂ i) Z + rewrite splitAt-↑ʳ X Y (H₂.j e₂ i) = cong ((X ↑ʳ_) ∘ (_↑ˡ Z) ∘ H₂.j e₂) (sym (cast-is-id refl i)) + ≗j e i | inj₂ e₃ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (H₂.h ↑ʳ e₃) + rewrite splitAt-↑ʳ H₂.h H₃.h e₃ + rewrite splitAt-↑ʳ (X + Y) Z (H₃.j e₃ i) = cong ((X ↑ʳ_) ∘ (Y ↑ʳ_) ∘ H₃.j e₃) (sym (cast-is-id refl i)) + +n+0↔n : ∀ n → Fin (n + 0) ↔ Fin n +n+0↔n n = record + { to = to + ; from = from + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + to : Fin (n + 0) → Fin n + to x with inj₁ x₁ ← splitAt n x = x₁ + from : Fin n → Fin (n + 0) + from x = x ↑ˡ 0 + from∘to : (x : Fin (n + 0)) → from (to x) ≡ x + from∘to x with inj₁ x₁ ← splitAt n x in eq = splitAt⁻¹-↑ˡ eq + to∘from : (x : Fin n) → to (from x) ≡ x + to∘from x rewrite splitAt-↑ˡ n x 0 = refl + +unitaryʳ : Hypergraph-same (map-nodes ([ id , (λ ()) ] ∘ splitAt n) (together H discrete)) H +unitaryʳ {n} {H} = record + { ↔h = h+0↔h + ; ≗a = ≗a + ; ≗j = ≗j + } + where + module H = Hypergraph H + module H+0 = Hypergraph (together {n} {0} H discrete) + h+0↔h : Fin H+0.h ↔ Fin H.h + h+0↔h = n+0↔n H.h + ≗a : (e : Fin (H.h + 0)) → [ H.a , (λ ()) ] (splitAt H.h e) ≡ H.a (Inverse.to h+0↔h e) + ≗a e with inj₁ e₁ ← splitAt H.h e in eq = refl + ≗j : (e : Fin (H.h + 0)) + (i : Fin (ℕ.suc ([ H.a , (λ ()) ] (splitAt H.h e)))) + → [ (λ x → x) , (λ ()) ] (splitAt n (j+j H discrete e i)) + ≡ H.j (Inverse.to h+0↔h e) (cast (cong ℕ.suc (≗a e)) i) + ≗j e i = ≗j-aux (splitAt H.h e) refl (j+j H discrete e) refl (≗a e) i + where + ≗j-aux + : (w : Fin H.h ⊎ Fin 0) + → (eq₁ : splitAt H.h e ≡ w) + → (w₁ : Fin (ℕ.suc ([ H.a , (λ ()) ] w)) → Fin (n + 0)) + → j+j H discrete e ≡ subst (λ hole → Fin (ℕ.suc ([ H.a , (λ ()) ] hole)) → Fin (n + 0)) (sym eq₁) w₁ + → (w₂ : [ H.a , (λ ()) ] w ≡ H.a (Inverse.to h+0↔h e)) + (i : Fin (ℕ.suc ([ H.a , (λ ()) ] w))) + → [ (λ x → x) , (λ ()) ] (splitAt n (w₁ i)) + ≡ H.j (Inverse.to h+0↔h e) (cast (cong ℕ.suc w₂) i) + ≗j-aux (inj₁ e₁) eq w₁ eq₁ w₂ i + with (inj₁ x) ← splitAt n (w₁ i) in eq₂ + rewrite eq = trans + (↑ˡ-injective 0 x (H.j e₁ i) (trans (splitAt⁻¹-↑ˡ eq₂) (sym (cong-app eq₁ i)))) + (cong (H.j e₁) (sym (cast-is-id refl i))) + ++-comm-↔ : ∀ (n m : ℕ) → Fin (n + m) ↔ Fin (m + n) ++-comm-↔ n m = record + { to = join m n ∘ swap ∘ splitAt n + ; from = join n m ∘ swap ∘ splitAt m + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + open ≡-Reasoning + to∘from : join m n ∘ swap ∘ splitAt n ∘ join n m ∘ swap ∘ splitAt m ≗ id + to∘from x = begin + (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ swap ∘ splitAt m) x ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ swap ∘ splitAt m) x ⟩ + (join m n ∘ swap ∘ swap ∘ splitAt m) x ≡⟨ (cong (join m n) ∘ swap-involutive ∘ splitAt m) x ⟩ + (join m n ∘ splitAt m) x ≡⟨ join-splitAt m n x ⟩ + x ∎ + from∘to : join n m ∘ swap ∘ splitAt m ∘ join m n ∘ swap ∘ splitAt n ≗ id + from∘to x = begin + (join n m ∘ swap ∘ splitAt m ∘ join m n ∘ swap ∘ splitAt n) x ≡⟨ (cong (join n m ∘ swap) ∘ splitAt-join m n ∘ swap ∘ splitAt n) x ⟩ + (join n m ∘ swap ∘ swap ∘ splitAt n) x ≡⟨ (cong (join n m) ∘ swap-involutive ∘ splitAt n) x ⟩ + (join n m ∘ splitAt n) x ≡⟨ join-splitAt n m x ⟩ + x ∎ + +[,]∘swap : {A B C : Set} {f : A → C} {g : B → C} → [ f , g ] ∘ swap ≗ [ g , f ] +[,]∘swap (inj₁ x) = refl +[,]∘swap (inj₂ y) = refl + +braiding + : {n m : ℕ} + {H₁ : Hypergraph n} + {H₂ : Hypergraph m} + → Hypergraph-same (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) (together H₂ H₁) +braiding {n} {m} {H₁} {H₂} = record + { ↔h = +-comm-↔ H₁.h H₂.h + ; ≗a = ≗a + ; ≗j = ≗j + } + where + open ≡-Reasoning + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + ≗a : (e : Fin (H₁.h + H₂.h)) + → [ H₁.a , H₂.a ] (splitAt H₁.h e) + ≡ [ H₂.a , H₁.a ] (splitAt H₂.h (join H₂.h H₁.h (swap (splitAt H₁.h e)))) + ≗a e = begin + [ H₁.a , H₂.a ] (splitAt H₁.h e) ≡⟨ [,]∘swap (splitAt H₁.h e) ⟨ + [ H₂.a , H₁.a ] (swap (splitAt H₁.h e)) ≡⟨ cong [ H₂.a , H₁.a ] (splitAt-join H₂.h H₁.h (swap (splitAt H₁.h e))) ⟨ + [ H₂.a , H₁.a ] (splitAt H₂.h (join H₂.h H₁.h (swap (splitAt H₁.h e)))) ∎ + ≗j : (e : Fin (Hypergraph.h (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)))) + (i : Fin (Hypergraph.arity (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) e)) + → Hypergraph.j (map-nodes ([ _↑ʳ_ m , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) e i + ≡ Hypergraph.j (together H₂ H₁) (Inverse.to (+-comm-↔ H₁.h H₂.h) e) (cast (cong ℕ.suc (≗a e)) i) + ≗j e i with splitAt H₁.h e + ≗j e i | inj₁ e₁ + rewrite splitAt-↑ˡ n (H₁.j e₁ i) m + rewrite splitAt-↑ʳ H₂.h H₁.h e₁ = cong ((m ↑ʳ_) ∘ H₁.j e₁) (sym (cast-is-id refl i)) + ≗j e i | inj₂ e₂ + rewrite splitAt-↑ʳ n m (H₂.j e₂ i) + rewrite splitAt-↑ˡ H₂.h e₂ H₁.h = cong ((_↑ˡ n) ∘ H₂.j e₂) (sym (cast-is-id refl i)) + +opaque + unfolding η ε + + associativity′ + : {n m o : ℕ} + → {x : ∣ (Hypergraphₛ n ⊗₀ Hypergraphₛ m) ⊗₀ Hypergraphₛ o ∣} + → Hypergraph-same + (map-nodes (Inverse.to (+-assoc-↔ n m o)) (η {n + m} {o} ⟨$⟩ ((η {n} {m} ⊗₁ (Id _)) ⟨$⟩ x))) + (η {n} {m + o} ⟨$⟩ ((Id _ ⊗₁ η {m} {o}) ⟨$⟩ (Setoids-×.associator.from ⟨$⟩ x))) + associativity′ {n} {m} {o} {(x , y) , z} = associativity x y z + + unitaryˡ′ + : {X : ∣ Setoids-×.unit ⊗₀ Hypergraphₛ n ∣} + → Hypergraph-same (η {0} {n} ⟨$⟩ ((ε ⊗₁ Id _) ⟨$⟩ X)) (Setoids-×.unitorˡ.from ⟨$⟩ X) + unitaryˡ′ = Hypergraph-same-refl + + unitaryʳ′ + : {X : ∣ Hypergraphₛ n ⊗₀ Setoids-×.unit ∣} + → Hypergraph-same (map-nodes ([ id , (λ ()) ] ∘ splitAt n) (η {n} {0} ⟨$⟩ ((Id _ ⊗₁ ε) ⟨$⟩ X))) (Setoids-×.unitorʳ.from ⟨$⟩ X) + unitaryʳ′ = unitaryʳ + + braiding-compat + : {n m : ℕ} + → {X : ∣ Hypergraphₛ n ⊗₀ Hypergraphₛ m ∣} + → Hypergraph-same + (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (η {n} {m} ⟨$⟩ X)) + (η {m} {n} ⟨$⟩ (Setoids-×.braiding.⇒.η (Hypergraphₛ n , Hypergraphₛ m) ⟨$⟩ X)) + braiding-compat {n} {m} {H₁ , H₂} = braiding {n} {m} {H₁} {H₂} + +hypergraph : SymmetricMonoidalFunctor Nat-smc Setoids-× +hypergraph = record + { F = F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = associativity′ + ; unitaryˡ = unitaryˡ′ + ; unitaryʳ = unitaryʳ′ + } + ; braiding-compat = braiding-compat + } + } + +module F = SymmetricMonoidalFunctor hypergraph + +and-gate : Func ⊤ₛ (F.₀ 3) +and-gate = Const ⊤ₛ (Hypergraphₛ 3) and-graph + where + and-graph : Hypergraph 3 + and-graph = record + { h = 1 + ; a = λ { 0F → 2 } + ; j = λ { 0F → id } + } + where + edge-0-nodes : Fin 3 → Fin 3 + edge-0-nodes 0F = # 0 + edge-0-nodes 1F = # 1 + edge-0-nodes 2F = # 2 diff --git a/DecorationFunctor/Hypergraph/Labeled.agda b/DecorationFunctor/Hypergraph/Labeled.agda new file mode 100644 index 0000000..31402b1 --- /dev/null +++ b/DecorationFunctor/Hypergraph/Labeled.agda @@ -0,0 +1,689 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module DecorationFunctor.Hypergraph.Labeled {c ℓ : Level} where + +import Categories.Morphism as Morphism + +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Core using (Category) +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using () renaming (_∘F_ to _∘′_) +open import Categories.Functor.Core using (Functor) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Category.Instance.Nat.FinitelyCocomplete using (Nat-FinitelyCocomplete) +open import Category.Instance.Setoids.SymmetricMonoidal {c} {ℓ} using (Setoids-×; ×-symmetric′) +open import Data.Fin using (#_; Fin; splitAt; join; zero; suc; _↑ˡ_; _↑ʳ_; cast) +open import Data.Fin.Patterns using (0F; 1F; 2F) +open import Data.Fin.Properties + using + ( splitAt-join ; join-splitAt + ; cast-is-id ; cast-trans ; subst-is-cast + ; splitAt-↑ˡ ; splitAt-↑ʳ + ; splitAt⁻¹-↑ˡ ; ↑ˡ-injective + ) +open import Data.Nat using (ℕ; _+_) +open import Data.Product.Base using (_,_; Σ) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid.Unit {c} {ℓ} using (⊤ₛ) +open import Data.Sum.Base using (_⊎_; map; inj₁; inj₂; swap; map₂) renaming ([_,_]′ to [_,_]) +open import Data.Sum.Properties using (map-map; [,]-map; [,]-∘; [-,]-cong; [,-]-cong; [,]-cong; map-cong; swap-involutive) +open import Data.Unit.Properties using () renaming (≡-setoid to ⊤-setoid) +open import Function using (_∘_; id; const; Func; Inverse; _↔_; mk↔; _⟶ₛ_) +open import Function.Construct.Composition using (_↔-∘_) +open import Function.Construct.Constant using () renaming (function to Const) +open import Function.Construct.Identity using (↔-id) +open import Function.Construct.Symmetry using (↔-sym) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality using (_≗_; _≡_; erefl; refl; sym; trans; cong; cong₂; subst; cong-app) +open import Relation.Binary.PropositionalEquality.Properties using (isEquivalence; module ≡-Reasoning; dcong; dcong₂; subst-∘; subst-subst; sym-cong; subst-subst-sym; trans-cong; cong-∘; trans-reflʳ) +open import Relation.Nullary.Negation.Core using (¬_) + +open Cartesian (Setoids-Cartesian {c} {ℓ}) using (products) +open FinitelyCocompleteCategory Nat-FinitelyCocomplete + using (-+-) + renaming (symmetricMonoidalCategory to Nat-smc; +-assoc to Nat-+-assoc) +open import Category.Monoidal.Instance.Nat using (Nat,+,0) +open Morphism (Setoids c ℓ) using (_≅_) +open Lax using (SymmetricMonoidalFunctor) + +open BinaryProducts products using (-×-) + +open import Data.Circuit.Gate using (Gate; cast-gate; cast-gate-trans; cast-gate-is-id; subst-is-cast-gate) + +record Hypergraph (v : ℕ) : Set c where + + field + h : ℕ + a : Fin h → ℕ + j : (e : Fin h) → Fin (a e) → Fin v + l : (e : Fin h) → Gate (a e) + +record Hypergraph-same {n : ℕ} (H H′ : Hypergraph n) : Set ℓ where + + open Hypergraph H public + open Hypergraph H′ renaming (h to h′; a to a′; j to j′; l to l′) public + + field + ↔h : Fin h ↔ Fin h′ + + open Inverse ↔h public + + field + ≗a : a ≗ a′ ∘ to + ≗j : (e : Fin h) + (i : Fin (a e)) + → j e i + ≡ j′ (to e) (cast (≗a e) i) + ≗l : (e : Fin h) + → l e + ≡ cast-gate (sym (≗a e)) (l′ (to e)) + +private + + variable + n n′ m m′ o : ℕ + H H′ H″ H₁ H₁′ : Hypergraph n + H₂ H₂′ : Hypergraph m + H₃ : Hypergraph o + +Hypergraph-same-refl : Hypergraph-same H H +Hypergraph-same-refl {_} {H} = record + { ↔h = ↔-id (Fin h) + ; ≗a = cong a ∘ erefl + ; ≗j = λ e i → cong (j e) (sym (cast-is-id refl i)) + ; ≗l = λ { e → sym (cast-gate-is-id refl (l e)) } + } + where + open Hypergraph H + +sym-sym : {A : Set} {x y : A} (p : x ≡ y) → sym (sym p) ≡ p +sym-sym refl = refl + +Hypergraph-same-sym : Hypergraph-same H H′ → Hypergraph-same H′ H +Hypergraph-same-sym {V} {H} {H′} ≡H = record + { ↔h = ↔-sym ↔h + ; ≗a = sym ∘ a∘from≗a′ + ; ≗j = ≗j′ + ; ≗l = ≗l′ + } + where + open Hypergraph-same ≡H + open ≡-Reasoning + a∘from≗a′ : a ∘ from ≗ a′ + a∘from≗a′ x = begin + (a ∘ from) x ≡⟨ (≗a ∘ from) x ⟩ + (a′ ∘ to ∘ from) x ≡⟨ (cong a′ ∘ inverseˡ ∘ erefl ∘ from) x ⟩ + a′ x ∎ + id≗to∘from : id ≗ to ∘ from + id≗to∘from e = sym (inverseˡ refl) + ≗arity′ : a′ ≗ a ∘ from + ≗arity′ e = sym (a∘from≗a′ e) + ≗arity- : a′ ≗ a′ ∘ to ∘ from + ≗arity- e = cong a′ (id≗to∘from e) + + ≗j′ : (e : Fin h′) (i : Fin (a′ e)) → j′ e i ≡ j (from e) (cast (≗arity′ e) i) + ≗j′ e i = begin + j′ e i ≡⟨ dcong₂ j′ (id≗to∘from e) (subst-∘ (id≗to∘from e)) ⟩ + j′ (to (from e)) (subst Fin (cong a′ (id≗to∘from e)) i) ≡⟨ cong (j′ (to (from e))) (subst-is-cast (cong a′ (id≗to∘from e)) i) ⟩ + j′ (to (from e)) (cast (cong a′ (id≗to∘from e)) i) ≡⟨⟩ + j′ (to (from e)) (cast (trans (≗arity′ e) (≗a (from e))) i) ≡⟨ cong (j′ (to (from e))) (cast-trans (≗arity′ e) (≗a (from e)) i) ⟨ + j′ (to (from e)) (cast (≗a (from e)) (cast (≗arity′ e) i)) ≡⟨ ≗j (from e) (cast (≗arity′ e) i) ⟨ + j (from e) (cast (≗arity′ e) i) ∎ + + ≗l′ : (e : Fin h′) → l′ e ≡ cast-gate (sym (sym (a∘from≗a′ e))) (l (from e)) + ≗l′ e = begin + l′ e ≡⟨ dcong l′ (sym (id≗to∘from e)) ⟨ + subst (Gate ∘ a′) (sym (id≗to∘from e)) (l′ (to (from e))) ≡⟨ subst-∘ (sym (id≗to∘from e)) ⟩ + subst Gate (cong a′ (sym (id≗to∘from e))) (l′ (to (from e))) ≡⟨ subst-is-cast-gate (cong a′ (sym (id≗to∘from e))) (l′ (to (from e))) ⟩ + cast-gate _ (l′ (to (from e))) ≡⟨ cast-gate-trans _ (sym (sym (a∘from≗a′ e))) (l′ (to (from e))) ⟨ + cast-gate (sym (sym (a∘from≗a′ e))) (cast-gate _ (l′ (to (from e)))) ≡⟨ cong (cast-gate (sym (sym (a∘from≗a′ e)))) (≗l (from e)) ⟨ + cast-gate (sym (sym (a∘from≗a′ e))) (l (from e)) ∎ + +Hypergraph-same-trans : Hypergraph-same H H′ → Hypergraph-same H′ H″ → Hypergraph-same H H″ +Hypergraph-same-trans ≡H₁ ≡H₂ = record + { ↔h = ↔h ≡H₂ ↔-∘ ↔h ≡H₁ + ; ≗a = λ { x → trans (≗a ≡H₁ x) (≗a ≡H₂ (to (↔h ≡H₁) x)) } + ; ≗j = λ { e i → trans (≗j ≡H₁ e i) (≗j₂ e i) } + ; ≗l = λ { e → trans (≗l ≡H₁ e) (≗l₂ e) } + } + where + open Hypergraph-same + open Inverse + open ≡-Reasoning + ≗j₂ : (e : Fin (h ≡H₁)) + (i : Fin (a ≡H₁ e)) + → j ≡H₂ (to (↔h ≡H₁) e) (cast (≗a ≡H₁ e) i) + ≡ j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (trans (≗a ≡H₁ e) (≗a ≡H₂ (to (↔h ≡H₁) e))) i) + ≗j₂ e i = begin + j ≡H₂ (to (↔h ≡H₁) e) (cast (≗a ≡H₁ e) i) + ≡⟨ ≗j ≡H₂ (to (↔h ≡H₁) e) (cast (≗a ≡H₁ e) i) ⟩ + j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (≗a ≡H₂ (to (↔h ≡H₁) e)) (cast (≗a ≡H₁ e) i)) + ≡⟨ cong (j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e))) (cast-trans (≗a ≡H₁ e) (≗a ≡H₂ (to (↔h ≡H₁) e)) i) ⟩ + j′ ≡H₂ (to (↔h ≡H₂) (to (↔h ≡H₁) e)) (cast (trans (≗a ≡H₁ e) (≗a ≡H₂ (to (↔h ≡H₁) e))) i) ∎ + ≗l₂ : (e : Fin (h ≡H₁)) → cast-gate _ (l′ ≡H₁ (to ≡H₁ e)) ≡ cast-gate _ (l′ ≡H₂ (to ≡H₂ (to ≡H₁ e))) + ≗l₂ e = trans (cong (cast-gate _) (≗l ≡H₂ (to ≡H₁ e))) (cast-gate-trans _ (sym (≗a ≡H₁ e)) (l′ ≡H₂ (to ≡H₂ (to ≡H₁ e)))) + +Hypergraph-setoid : ℕ → Setoid c ℓ +Hypergraph-setoid p = record + { Carrier = Hypergraph p + ; _≈_ = Hypergraph-same + ; isEquivalence = record + { refl = Hypergraph-same-refl + ; sym = Hypergraph-same-sym + ; trans = Hypergraph-same-trans + } + } + +map-nodes : (Fin n → Fin m) → Hypergraph n → Hypergraph m +map-nodes f H = record + { h = h + ; a = a + ; j = λ e i → f (j e i) + ; l = l + } + where + open Hypergraph H + +Hypergraph-same-cong + : (f : Fin n → Fin m) + → Hypergraph-same H H′ + → Hypergraph-same (map-nodes f H) (map-nodes f H′) +Hypergraph-same-cong f ≡H = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ { e i → cong f (≗j e i) } + ; ≗l = ≗l + } + where + open Hypergraph-same ≡H + +Hypergraph-Func : (Fin n → Fin m) → Hypergraph-setoid n ⟶ₛ Hypergraph-setoid m +Hypergraph-Func f = record + { to = map-nodes f + ; cong = Hypergraph-same-cong f + } + +F-resp-≈ + : {f g : Fin n → Fin m} + → f ≗ g + → Hypergraph-same (map-nodes f H) (map-nodes g H) +F-resp-≈ {g = g} f≗g = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ { e i → trans (f≗g (j e i)) (cong g (≗j e i)) } + ; ≗l = ≗l + } + where + open Hypergraph-same Hypergraph-same-refl + +homomorphism + : (f : Fin n → Fin m) + → (g : Fin m → Fin o) + → Hypergraph-same (map-nodes (g ∘ f) H) (map-nodes g (map-nodes f H)) +homomorphism {n} {m} {o} {H} f g = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = λ e i → cong (g ∘ f) (≗j e i) + ; ≗l = ≗l + } + where + open Hypergraph-same Hypergraph-same-refl + +F : Functor Nat (Setoids c ℓ) +F = record + { F₀ = λ n → Hypergraph-setoid n + ; F₁ = Hypergraph-Func + ; identity = λ { {n} {H} → Hypergraph-same-refl {H = H} } + ; homomorphism = λ { {f = f} {g = g} → homomorphism f g } + ; F-resp-≈ = λ f≗g → F-resp-≈ f≗g + } + +-- monoidal structure + +discrete : {n : ℕ} → Hypergraph n +discrete = record + { h = 0 + ; a = λ () + ; j = λ () + ; l = λ () + } + +ε : ⊤ₛ ⟶ₛ Hypergraph-setoid 0 +ε = Const ⊤ₛ (Hypergraph-setoid 0) discrete + +module _ (H₁ : Hypergraph n) (H₂ : Hypergraph m) where + private + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + j+j : (e : Fin (H₁.h + H₂.h)) + → Fin ([ H₁.a , H₂.a ] (splitAt H₁.h e)) + → Fin (n + m) + j+j e i with splitAt H₁.h e + ... | inj₁ e₁ = H₁.j e₁ i ↑ˡ m + ... | inj₂ e₂ = n ↑ʳ H₂.j e₂ i + l+l : (e : Fin (H₁.h + H₂.h)) → Gate ([ H₁.a , H₂.a ] (splitAt H₁.h e)) + l+l e with splitAt H₁.h e + ... | inj₁ e₁ = H₁.l e₁ + ... | inj₂ e₂ = H₂.l e₂ + +together : Hypergraph n → Hypergraph m → Hypergraph (n + m) +together {n} {m} H₁ H₂ = record + { h = h H₁ + h H₂ + ; a = [ a H₁ , a H₂ ] ∘ splitAt (h H₁) + ; j = j+j H₁ H₂ + ; l = l+l H₁ H₂ + } + where + open Hypergraph + ++-resp-↔ + : {n n′ m m′ : ℕ} + → Fin n ↔ Fin n′ + → Fin m ↔ Fin m′ + → Fin (n + m) ↔ Fin (n′ + m′) ++-resp-↔ {n} {n′} {m} {m′} ↔n ↔m = record + { to = join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n + ; from = join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ + ; to-cong = cong (join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n) + ; from-cong = cong (join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′) + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + module ↔n = Inverse ↔n + module ↔m = Inverse ↔m + open ≡-Reasoning + to∘from : join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ∘ join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ≗ id + to∘from x = begin + (join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ∘ join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′) x + ≡⟨ cong + (join n′ m′ ∘ map ↔n.to ↔m.to) + (splitAt-join n m (map ↔n.from ↔m.from (splitAt n′ x))) ⟩ + (join n′ m′ ∘ map ↔n.to ↔m.to ∘ map ↔n.from ↔m.from ∘ splitAt n′) x + ≡⟨ cong (join n′ m′) (map-map (splitAt n′ x)) ⟩ + (join n′ m′ ∘ map (↔n.to ∘ ↔n.from) (↔m.to ∘ ↔m.from) ∘ splitAt n′) x + ≡⟨ cong + (join n′ m′) + (map-cong + (λ _ → ↔n.inverseˡ refl) + (λ _ → ↔m.inverseˡ refl) + (splitAt n′ x)) ⟩ + (join n′ m′ ∘ map id id ∘ splitAt n′) x ≡⟨ [,]-map (splitAt n′ x) ⟩ + (join n′ m′ ∘ splitAt n′) x ≡⟨ join-splitAt n′ m′ x ⟩ + x ∎ + from∘to : join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ∘ join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n ≗ id + from∘to x = begin + (join n m ∘ map ↔n.from ↔m.from ∘ splitAt n′ ∘ join n′ m′ ∘ map ↔n.to ↔m.to ∘ splitAt n) x + ≡⟨ cong + (join n m ∘ map ↔n.from ↔m.from) + (splitAt-join n′ m′ (map ↔n.to ↔m.to (splitAt n x))) ⟩ + (join n m ∘ map ↔n.from ↔m.from ∘ map ↔n.to ↔m.to ∘ splitAt n) x + ≡⟨ cong (join n m) (map-map (splitAt n x)) ⟩ + (join n m ∘ map (↔n.from ∘ ↔n.to) (↔m.from ∘ ↔m.to) ∘ splitAt n) x + ≡⟨ cong + (join n m) + (map-cong + (λ _ → ↔n.inverseʳ refl) + (λ _ → ↔m.inverseʳ refl) + (splitAt n x)) ⟩ + (join n m ∘ map id id ∘ splitAt n) x ≡⟨ [,]-map (splitAt n x) ⟩ + (join n m ∘ splitAt n) x ≡⟨ join-splitAt n m x ⟩ + x ∎ + +together-resp-same + : Hypergraph-same H₁ H₁′ + → Hypergraph-same H₂ H₂′ + → Hypergraph-same (together H₁ H₂) (together H₁′ H₂′) +together-resp-same {n} {H₁} {H₁′} {m} {H₂} {H₂′} ≡H₁ ≡H₂ = record + { ↔h = +-resp-↔ ≡H₁.↔h ≡H₂.↔h + ; ≗a = ≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + module ≡H₁ = Hypergraph-same ≡H₁ + module ≡H₂ = Hypergraph-same ≡H₂ + module H₁+H₂ = Hypergraph (together H₁ H₂) + module H₁+H₂′ = Hypergraph (together H₁′ H₂′) + open ≡-Reasoning + open Inverse + open Hypergraph + ≗a : [ ≡H₁.a , ≡H₂.a ] ∘ splitAt ≡H₁.h + ≗ [ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt ≡H₁.h′ + ∘ join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h + ≗a e = begin + [ ≡H₁.a , ≡H₂.a ] (splitAt ≡H₁.h e) ≡⟨ [,]-cong ≡H₁.≗a ≡H₂.≗a (splitAt ≡H₁.h e) ⟩ + ([ ≡H₁.a′ ∘ ≡H₁.to , ≡H₂.a′ ∘ ≡H₂.to ] ∘ splitAt ≡H₁.h) e ≡⟨ [,]-map (splitAt ≡H₁.h e) ⟨ + ([ ≡H₁.a′ , ≡H₂.a′ ] ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ≡⟨ (cong [ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt-join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ⟨ + ([ ≡H₁.a′ , ≡H₂.a′ ] ∘ splitAt ≡H₁.h′ ∘ join ≡H₁.h′ ≡H₂.h′ ∘ map ≡H₁.to ≡H₂.to ∘ splitAt ≡H₁.h) e ∎ + ≗j : (e : Fin H₁+H₂.h) + (i : Fin (H₁+H₂.a e)) + → H₁+H₂.j e i + ≡ H₁+H₂′.j (to (+-resp-↔ ≡H₁.↔h ≡H₂.↔h) e) (cast (≗a e) i) + ≗j e i with splitAt ≡H₁.h e + ... | inj₁ e₁ rewrite splitAt-↑ˡ ≡H₁.h′ (≡H₁.to e₁) ≡H₂.h′ = cong (_↑ˡ m) (≡H₁.≗j e₁ i) + ... | inj₂ e₂ rewrite splitAt-↑ʳ ≡H₁.h′ ≡H₂.h′ (≡H₂.to e₂) = cong (n ↑ʳ_) (≡H₂.≗j e₂ i) + ≗l : (e : Fin H₁+H₂.h) → l+l H₁ H₂ e ≡ cast-gate (sym (≗a e)) (l+l H₁′ H₂′ (to (+-resp-↔ ≡H₁.↔h ≡H₂.↔h) e)) + ≗l e with splitAt ≡H₁.h e | .{≗a e} + ... | inj₁ e₁ rewrite splitAt-↑ˡ ≡H₁.h′ (≡H₁.to e₁) ≡H₂.h′ = ≡H₁.≗l e₁ + ... | inj₂ e₂ rewrite splitAt-↑ʳ ≡H₁.h′ ≡H₂.h′ (≡H₂.to e₂) = ≡H₂.≗l e₂ + +commute + : (f : Fin n → Fin n′) + → (g : Fin m → Fin m′) + → Hypergraph-same + (together (map-nodes f H₁) (map-nodes g H₂)) + (map-nodes ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n) (together H₁ H₂)) +commute {n} {n′} {m} {m′} {H₁} {H₂} f g = record + { ↔h = ≡H₁+H₂.↔h + ; ≗a = ≡H₁+H₂.≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + module H₁+H₂ = Hypergraph (together H₁ H₂) + module ≡H₁+H₂ = Hypergraph-same {H = together H₁ H₂} Hypergraph-same-refl + open Hypergraph + open ≡-Reasoning + ≗j : (e : Fin (H₁.h + H₂.h)) + (i : Fin (([ H₁.a , H₂.a ] ∘ splitAt H₁.h) e)) + → j (together (map-nodes f H₁) (map-nodes g H₂)) e i + ≡ j (map-nodes ([ (_↑ˡ m′) ∘ f , (n′ ↑ʳ_) ∘ g ] ∘ splitAt n) (together H₁ H₂)) (≡H₁+H₂.to e) (cast refl i) + ≗j e i rewrite (cast-is-id refl i) with splitAt H₁.h e + ... | inj₁ e₁ rewrite splitAt-↑ˡ n (H₁.j e₁ i) m = refl + ... | inj₂ e₂ rewrite splitAt-↑ʳ n m (H₂.j e₂ i) = refl + ≗l : (e : Fin (H₁.h + H₂.h)) + → l+l (map-nodes f H₁) (map-nodes g H₂) e + ≡ cast-gate refl (l+l H₁ H₂ (≡H₁+H₂.to e)) + ≗l e rewrite cast-gate-is-id refl (l+l H₁ H₂ (≡H₁+H₂.to e)) with splitAt H₁.h e + ... | inj₁ e₁ = refl + ... | inj₂ e₁ = refl + +⊗-homomorphism : NaturalTransformation (-×- ∘′ (F ⁂ F)) (F ∘′ -+-) +⊗-homomorphism = record + { η = λ { (m , n) → η } + ; commute = λ { (f , g) {H₁ , H₂} → commute {H₁ = H₁} {H₂ = H₂} f g } + ; sym-commute = λ { (f , g) {H₁ , H₂} → Hypergraph-same-sym (commute {H₁ = H₁} {H₂ = H₂} f g) } + } + where + η : Hypergraph-setoid n ×ₛ Hypergraph-setoid m ⟶ₛ Hypergraph-setoid (n + m) + η = record + { to = λ { (H₁ , H₂) → together H₁ H₂ } + ; cong = λ { (≡H₁ , ≡H₂) → together-resp-same ≡H₁ ≡H₂ } + } + ++-assoc-↔ : ∀ (x y z : ℕ) → Fin (x + y + z) ↔ Fin (x + (y + z)) ++-assoc-↔ x y z = record + { to = to + ; from = from + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → isoˡ _ }) , λ { refl → isoʳ _ } + } + where + module Nat = Morphism Nat + open Nat._≅_ (Nat-+-assoc {x} {y} {z}) + +associativity + : {X Y Z : ℕ} + → {H₁ : Hypergraph X} + → {H₂ : Hypergraph Y} + → {H₃ : Hypergraph Z} + → Hypergraph-same + (map-nodes (Inverse.to (+-assoc-↔ X Y Z)) (together (together H₁ H₂) H₃)) + (together H₁ (together H₂ H₃)) +associativity {X} {Y} {Z} {H₁} {H₂} {H₃} = record + { ↔h = ↔h + ; ≗a = ≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + module H₃ = Hypergraph H₃ + open ≡-Reasoning + open Hypergraph + ↔h : Fin (H₁.h + H₂.h + H₃.h) ↔ Fin (H₁.h + (H₂.h + H₃.h)) + ↔h = +-assoc-↔ H₁.h H₂.h H₃.h + ≗a : (x : Fin (H₁.h + H₂.h + H₃.h)) + → [ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] (splitAt (H₁.h + H₂.h) x) + ≡ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] (splitAt H₁.h ([ [ _↑ˡ H₂.h + H₃.h , (H₁.h ↑ʳ_) ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] (splitAt (H₁.h + H₂.h) x))) + ≗a x = begin + ([ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ inj₁ ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [-,]-cong ([,-]-cong (cong [ H₂.a , H₃.a ] ∘ (λ i → splitAt-↑ˡ H₂.h i H₃.h)) ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [-,]-cong ([,]-∘ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , H₃.a ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₂.a , H₃.a ] ∘ inj₂ ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [,-]-cong (cong [ H₂.a , H₃.a ] ∘ splitAt-↑ʳ H₂.h H₃.h) (splitAt (H₁.h + H₂.h) x) ⟨ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ inj₂ ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ [,]-∘ [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] (splitAt (H₁.h + H₂.h) x) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , inj₂ ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([,-]-cong (splitAt-↑ʳ H₁.h (H₂.h + H₃.h) ∘ (H₂.h ↑ʳ_)) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , splitAt H₁.h ∘ (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([-,]-cong (splitAt-join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ [ splitAt H₁.h ∘ join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , splitAt H₁.h ∘ (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong [ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ([,]-∘ (splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟨ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ join H₁.h (H₂.h + H₃.h) ∘ map₂ (_↑ˡ H₃.h) ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨⟩ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ [ _↑ˡ H₂.h + H₃.h , H₁.h ↑ʳ_ ] ∘ [ inj₁ , inj₂ ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x + ≡⟨ cong ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h) ([-,]-cong ([,]-∘ [ _↑ˡ H₂.h + H₃.h , H₁.h ↑ʳ_ ] ∘ splitAt H₁.h) (splitAt (H₁.h + H₂.h) x)) ⟩ + ([ H₁.a , [ H₂.a , H₃.a ] ∘ splitAt H₂.h ] ∘ splitAt H₁.h ∘ [ [ _↑ˡ H₂.h + H₃.h , (H₁.h ↑ʳ_) ∘ (_↑ˡ H₃.h) ] ∘ splitAt H₁.h , (H₁.h ↑ʳ_) ∘ (H₂.h ↑ʳ_) ] ∘ splitAt (H₁.h + H₂.h)) x ∎ + ≗j : (e : Fin (H₁.h + H₂.h + H₃.h)) + (i : Fin ([ [ H₁.a , H₂.a ] ∘ splitAt H₁.h , H₃.a ] (splitAt (H₁.h + H₂.h) e))) + → Inverse.to (+-assoc-↔ X Y Z) (j+j (together H₁ H₂) H₃ e i) + ≡ j+j H₁ (together H₂ H₃) (Inverse.to ↔h e) (cast (≗a e) i) + ≗j e i with splitAt (H₁.h + H₂.h) e + ≗j e i | inj₁ e₁₂ with splitAt H₁.h e₁₂ + ≗j e i | inj₁ e₁₂ | inj₁ e₁ + rewrite splitAt-↑ˡ H₁.h e₁ (H₂.h + H₃.h) + rewrite splitAt-↑ˡ (X + Y) (H₁.j e₁ i ↑ˡ Y) Z + rewrite splitAt-↑ˡ X (H₁.j e₁ i) Y = cong ((_↑ˡ Y + Z) ∘ H₁.j e₁) (sym (cast-is-id refl i)) + ≗j e i | inj₁ e₁₂ | inj₂ e₂ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (e₂ ↑ˡ H₃.h) + rewrite splitAt-↑ˡ H₂.h e₂ H₃.h + rewrite splitAt-↑ˡ (X + Y) (X ↑ʳ H₂.j e₂ i) Z + rewrite splitAt-↑ʳ X Y (H₂.j e₂ i) = cong ((X ↑ʳ_) ∘ (_↑ˡ Z) ∘ H₂.j e₂) (sym (cast-is-id refl i)) + ≗j e i | inj₂ e₃ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (H₂.h ↑ʳ e₃) + rewrite splitAt-↑ʳ H₂.h H₃.h e₃ + rewrite splitAt-↑ʳ (X + Y) Z (H₃.j e₃ i) = cong ((X ↑ʳ_) ∘ (Y ↑ʳ_) ∘ H₃.j e₃) (sym (cast-is-id refl i)) + ≗l : (e : Fin (H₁.h + H₂.h + H₃.h)) + → l (map-nodes (Inverse.to (+-assoc-↔ X Y Z)) (together (together H₁ H₂) H₃)) e + ≡ cast-gate (sym (≗a e)) (l (together H₁ (together H₂ H₃)) (Inverse.to ↔h e)) + ≗l e with splitAt (H₁.h + H₂.h) e + ≗l e | inj₁ e₁₂ with splitAt H₁.h e₁₂ + ≗l e | inj₁ e₁₂ | inj₁ e₁ + rewrite splitAt-↑ˡ H₁.h e₁ (H₂.h + H₃.h) = sym (cast-gate-is-id refl (H₁.l e₁)) + ≗l e | inj₁ e₁₂ | inj₂ e₂ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (e₂ ↑ˡ H₃.h) + rewrite splitAt-↑ˡ H₂.h e₂ H₃.h = sym (cast-gate-is-id refl (H₂.l e₂)) + ≗l e | inj₂ e₃ + rewrite splitAt-↑ʳ H₁.h (H₂.h + H₃.h) (H₂.h ↑ʳ e₃) + rewrite splitAt-↑ʳ H₂.h H₃.h e₃ = sym (cast-gate-is-id refl (H₃.l e₃)) + +n+0↔n : ∀ n → Fin (n + 0) ↔ Fin n +n+0↔n n = record + { to = to + ; from = from + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + to : Fin (n + 0) → Fin n + to x with inj₁ x₁ ← splitAt n x = x₁ + from : Fin n → Fin (n + 0) + from x = x ↑ˡ 0 + from∘to : (x : Fin (n + 0)) → from (to x) ≡ x + from∘to x with inj₁ x₁ ← splitAt n x in eq = splitAt⁻¹-↑ˡ eq + to∘from : (x : Fin n) → to (from x) ≡ x + to∘from x rewrite splitAt-↑ˡ n x 0 = refl + +unitaryʳ : Hypergraph-same (map-nodes ([ id , (λ ()) ] ∘ splitAt n) (together H (discrete {0}))) H +unitaryʳ {n} {H} = record + { ↔h = h+0↔h + ; ≗a = ≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + module H = Hypergraph H + module H+0 = Hypergraph (together H (discrete {0})) + h+0↔h : Fin H+0.h ↔ Fin H.h + h+0↔h = n+0↔n H.h + ≗a : (e : Fin (H.h + 0)) → [ H.a , (λ ()) ] (splitAt H.h e) ≡ H.a (Inverse.to h+0↔h e) + ≗a e with inj₁ e₁ ← splitAt H.h e in eq = refl + ≗j : (e : Fin (H.h + 0)) + (i : Fin ([ H.a , (λ ()) ] (splitAt H.h e))) + → [ (λ x → x) , (λ ()) ] (splitAt n (j+j H discrete e i)) + ≡ H.j (Inverse.to h+0↔h e) (cast (≗a e) i) + ≗j e i = ≗j-aux (splitAt H.h e) refl (j+j H discrete e) refl (≗a e) i + where + ≗j-aux + : (w : Fin H.h ⊎ Fin 0) + → (eq₁ : splitAt H.h e ≡ w) + → (w₁ : Fin ([ H.a , (λ ()) ] w) → Fin (n + 0)) + → j+j H discrete e ≡ subst (λ hole → Fin ([ H.a , (λ ()) ] hole) → Fin (n + 0)) (sym eq₁) w₁ + → (w₂ : [ H.a , (λ ()) ] w ≡ H.a (Inverse.to h+0↔h e)) + (i : Fin ([ H.a , (λ ()) ] w)) + → [ (λ x → x) , (λ ()) ] (splitAt n (w₁ i)) + ≡ H.j (Inverse.to h+0↔h e) (cast w₂ i) + ≗j-aux (inj₁ e₁) eq w₁ eq₁ w₂ i + with (inj₁ x) ← splitAt n (w₁ i) in eq₂ + rewrite eq = trans + (↑ˡ-injective 0 x (H.j e₁ i) (trans (splitAt⁻¹-↑ˡ eq₂) (sym (cong-app eq₁ i)))) + (cong (H.j e₁) (sym (cast-is-id refl i))) + ≗l : (e : Fin (H.h + 0)) + → l+l H discrete e + ≡ cast-gate (sym (≗a e)) (H.l (Inverse.to h+0↔h e)) + ≗l e with splitAt H.h e | {(≗a e)} + ... | inj₁ e₁ = sym (cast-gate-is-id refl (H.l e₁)) + ++-comm-↔ : ∀ (n m : ℕ) → Fin (n + m) ↔ Fin (m + n) ++-comm-↔ n m = record + { to = join m n ∘ swap ∘ splitAt n + ; from = join n m ∘ swap ∘ splitAt m + ; to-cong = λ { refl → refl } + ; from-cong = λ { refl → refl } + ; inverse = (λ { refl → to∘from _ }) , λ { refl → from∘to _ } + } + where + open ≡-Reasoning + to∘from : join m n ∘ swap ∘ splitAt n ∘ join n m ∘ swap ∘ splitAt m ≗ id + to∘from x = begin + (join m n ∘ swap ∘ splitAt n ∘ join n m ∘ swap ∘ splitAt m) x ≡⟨ (cong (join m n ∘ swap) ∘ splitAt-join n m ∘ swap ∘ splitAt m) x ⟩ + (join m n ∘ swap ∘ swap ∘ splitAt m) x ≡⟨ (cong (join m n) ∘ swap-involutive ∘ splitAt m) x ⟩ + (join m n ∘ splitAt m) x ≡⟨ join-splitAt m n x ⟩ + x ∎ + from∘to : join n m ∘ swap ∘ splitAt m ∘ join m n ∘ swap ∘ splitAt n ≗ id + from∘to x = begin + (join n m ∘ swap ∘ splitAt m ∘ join m n ∘ swap ∘ splitAt n) x ≡⟨ (cong (join n m ∘ swap) ∘ splitAt-join m n ∘ swap ∘ splitAt n) x ⟩ + (join n m ∘ swap ∘ swap ∘ splitAt n) x ≡⟨ (cong (join n m) ∘ swap-involutive ∘ splitAt n) x ⟩ + (join n m ∘ splitAt n) x ≡⟨ join-splitAt n m x ⟩ + x ∎ + +[,]∘swap : {A B C : Set} {f : A → C} {g : B → C} → [ f , g ] ∘ swap ≗ [ g , f ] +[,]∘swap (inj₁ x) = refl +[,]∘swap (inj₂ y) = refl + +braiding + : {n m : ℕ} + {H₁ : Hypergraph n} + {H₂ : Hypergraph m} + → Hypergraph-same (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) (together H₂ H₁) +braiding {n} {m} {H₁} {H₂} = record + { ↔h = +-comm-↔ H₁.h H₂.h + ; ≗a = ≗a + ; ≗j = ≗j + ; ≗l = ≗l + } + where + open ≡-Reasoning + module H₁ = Hypergraph H₁ + module H₂ = Hypergraph H₂ + ≗a : (e : Fin (H₁.h + H₂.h)) + → [ H₁.a , H₂.a ] (splitAt H₁.h e) + ≡ [ H₂.a , H₁.a ] (splitAt H₂.h (join H₂.h H₁.h (swap (splitAt H₁.h e)))) + ≗a e = begin + [ H₁.a , H₂.a ] (splitAt H₁.h e) ≡⟨ [,]∘swap (splitAt H₁.h e) ⟨ + [ H₂.a , H₁.a ] (swap (splitAt H₁.h e)) ≡⟨ cong [ H₂.a , H₁.a ] (splitAt-join H₂.h H₁.h (swap (splitAt H₁.h e))) ⟨ + [ H₂.a , H₁.a ] (splitAt H₂.h (join H₂.h H₁.h (swap (splitAt H₁.h e)))) ∎ + ≗j : (e : Fin (Hypergraph.h (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)))) + (i : Fin (Hypergraph.a (map-nodes ([ m ↑ʳ_ , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) e)) + → Hypergraph.j (map-nodes ([ _↑ʳ_ m , _↑ˡ n ] ∘ splitAt n) (together H₁ H₂)) e i + ≡ Hypergraph.j (together H₂ H₁) (Inverse.to (+-comm-↔ H₁.h H₂.h) e) (cast (≗a e) i) + ≗j e i with splitAt H₁.h e + ≗j e i | inj₁ e₁ + rewrite splitAt-↑ˡ n (H₁.j e₁ i) m + rewrite splitAt-↑ʳ H₂.h H₁.h e₁ = cong ((m ↑ʳ_) ∘ H₁.j e₁) (sym (cast-is-id refl i)) + ≗j e i | inj₂ e₂ + rewrite splitAt-↑ʳ n m (H₂.j e₂ i) + rewrite splitAt-↑ˡ H₂.h e₂ H₁.h = cong ((_↑ˡ n) ∘ H₂.j e₂) (sym (cast-is-id refl i)) + ≗l : (e : Fin (H₁.h + H₂.h)) + → l+l H₁ H₂ e + ≡ cast-gate (sym (≗a e)) (l+l H₂ H₁ (Inverse.to (+-comm-↔ H₁.h H₂.h) e)) + ≗l e with splitAt H₁.h e | .{≗a e} + ≗l e | inj₁ e₁ rewrite splitAt-↑ʳ H₂.h H₁.h e₁ = sym (cast-gate-is-id refl (H₁.l e₁)) + ≗l e | inj₂ e₂ rewrite splitAt-↑ˡ H₂.h e₂ H₁.h = sym (cast-gate-is-id refl (H₂.l e₂)) + +opaque + unfolding ×-symmetric′ + Circ : SymmetricMonoidalFunctor Nat,+,0 Setoids-× + Circ = record + { F = F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = ε + ; ⊗-homo = ntHelper record + { η = λ { (m , n) → η } + ; commute = λ { (f , g) {H₁ , H₂} → commute {H₁ = H₁} {H₂ = H₂} f g } + } + ; associativity = λ { {X} {Y} {Z} {(H₁ , H₂) , H₃} → associativity {X} {Y} {Z} {H₁} {H₂} {H₃} } + ; unitaryˡ = Hypergraph-same-refl + ; unitaryʳ = unitaryʳ + } + ; braiding-compat = λ { {X} {Y} {H₁ , H₂} → braiding {X} {Y} {H₁} {H₂} } + } + } + where + η : Hypergraph-setoid n ×ₛ Hypergraph-setoid m ⟶ₛ Hypergraph-setoid (n + m) + η = record + { to = λ { (H₁ , H₂) → together H₁ H₂ } + ; cong = λ { (≡H₁ , ≡H₂) → together-resp-same ≡H₁ ≡H₂ } + } + +module F = SymmetricMonoidalFunctor Circ + +open Gate + +and-gate : ⊤ₛ ⟶ₛ Hypergraph-setoid 3 +and-gate = Const ⊤ₛ (Hypergraph-setoid 3) and-graph + where + and-graph : Hypergraph 3 + and-graph = record + { h = 1 + ; a = λ { 0F → 3 } + ; j = λ { 0F → edge-0-nodes } + ; l = λ { 0F → AND } + } + where + edge-0-nodes : Fin 3 → Fin 3 + edge-0-nodes 0F = # 0 + edge-0-nodes 1F = # 1 + edge-0-nodes 2F = # 2 diff --git a/DecorationFunctor/Trivial.agda b/DecorationFunctor/Trivial.agda index dee7c2e..6278d05 100644 --- a/DecorationFunctor/Trivial.agda +++ b/DecorationFunctor/Trivial.agda @@ -2,86 +2,70 @@ module DecorationFunctor.Trivial where -import Categories.Morphism as Morphism +open import Level using (0ℓ) open import Categories.Category.BinaryProducts using (module BinaryProducts) open import Categories.Category.Cartesian using (Cartesian) -open import Categories.Category.Cocartesian using (Cocartesian; module BinaryCoproducts) -open import Categories.Category.Core using (Category) -open import Categories.Category.Instance.Nat using (Nat-Cocartesian) open import Categories.Category.Instance.Nat using (Nat) open import Categories.Category.Instance.Setoids using (Setoids) -open import Categories.Category.Instance.SingletonSet using (SingletonSetoid) open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) open import Categories.Category.Product using (_⁂_) -open import Categories.Functor using () renaming (_∘F_ to _∘_) -open import Categories.Functor.Core using (Functor) +open import Categories.Functor using (Functor) renaming (_∘F_ to _∘_) open import Categories.Functor.Monoidal.Symmetric using (module Lax) -open import Categories.NaturalTransformation using (NaturalTransformation) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) -open import Category.Instance.Setoids.SymmetricMonoidal using (Setoids-×) open import Category.Instance.Nat.FinitelyCocomplete using (Nat-FinitelyCocomplete) -open import Data.Nat using (ℕ) -open import Data.Product.Base using (_,_) +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×; ×-symmetric′) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid.Unit {0ℓ} {0ℓ} using (⊤ₛ) open import Data.Unit using (tt) -open import Data.Unit.Properties using () renaming (≡-setoid to ⊤-setoid) -open import Function.Base using (const) -open import Function.Bundles using (Func) -open import Level using (0ℓ; suc; lift) -open import Relation.Binary.Bundles using (Setoid) +open import Data.Unit.Properties using () renaming (≡-setoid to ⊤-≡ₛ) +open import Function using (Func; const) +open import Function.Construct.Constant using () renaming (function to Const) open import Relation.Binary.PropositionalEquality.Core using (refl) open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) -open Cocartesian Nat-Cocartesian using (coproducts) +open BinaryProducts products using (-×-) open FinitelyCocompleteCategory Nat-FinitelyCocomplete - using () - renaming (symmetricMonoidalCategory to Nat-smc) -open Morphism (Setoids 0ℓ 0ℓ) using (_≅_) + using (-+-) + renaming (symmetricMonoidalCategory to Nat-smc) open Lax using (SymmetricMonoidalFunctor) -open BinaryProducts products using (-×-) -open BinaryCoproducts coproducts using (-+-) - - F : Functor Nat (Setoids 0ℓ 0ℓ) F = record - { F₀ = const (⊤-setoid) - ; F₁ = const (record { to = const tt ; cong = const refl }) - ; identity = const refl - ; homomorphism = const refl - ; F-resp-≈ = const (const refl) + { F₀ = const ⊤-≡ₛ + ; F₁ = const (Const ⊤-≡ₛ ⊤-≡ₛ tt) + ; identity = refl + ; homomorphism = refl + ; F-resp-≈ = const refl } -ε : Func (SingletonSetoid {0ℓ} {0ℓ}) ⊤-setoid -ε = record - { to = const tt - ; cong = const refl - } +ε : Func ⊤ₛ ⊤-≡ₛ +ε = Const ⊤ₛ ⊤-≡ₛ tt ⊗-homomorphism : NaturalTransformation (-×- ∘ (F ⁂ F)) (F ∘ -+-) -⊗-homomorphism = record - { η = const (record { to = const tt ; cong = const refl }) - ; commute = const (const refl) - ; sym-commute = const (const refl) +⊗-homomorphism = ntHelper record + { η = const (Const (⊤-≡ₛ ×ₛ ⊤-≡ₛ) ⊤-≡ₛ tt) + ; commute = const refl } -trivial : SymmetricMonoidalFunctor Nat-smc (Setoids-× {0ℓ}) -trivial = record - { F = F - ; isBraidedMonoidal = record - { isMonoidal = record - { ε = ε - ; ⊗-homo = ⊗-homomorphism - ; associativity = const refl - ; unitaryˡ = const refl - ; unitaryʳ = const refl - } - ; braiding-compat = const refl - } - } +opaque + unfolding ×-symmetric′ -and-gate : Func (SingletonSetoid {0ℓ} {0ℓ}) ⊤-setoid -and-gate = record - { to = const tt - ; cong = const refl - } + trivial : SymmetricMonoidalFunctor Nat-smc Setoids-× + trivial = record + { F = F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = refl + ; unitaryˡ = refl + ; unitaryʳ = refl + } + ; braiding-compat = refl + } + } + +and-gate : Func ⊤ₛ ⊤-≡ₛ +and-gate = Const ⊤ₛ ⊤-≡ₛ tt diff --git a/FinMerge.agda b/FinMerge.agda index d7b3f0b..2210e2f 100644 --- a/FinMerge.agda +++ b/FinMerge.agda @@ -13,8 +13,9 @@ open import Relation.Binary.PropositionalEquality.Properties using (module ≡-R open import Function using (id ; _∘_ ; _$_) open import Data.Maybe.Base using (Maybe; just; nothing; fromMaybe; map) -open import Util using (_<_<_; _<_≤_; toℕ<; Ordering; less; equal; greater; compare) +open import FinMerge.Util using (_<_≤_; Ordering; compare) +open Ordering private variable diff --git a/FinMerge/Properties.agda b/FinMerge/Properties.agda index 5f92905..6e9d53a 100644 --- a/FinMerge/Properties.agda +++ b/FinMerge/Properties.agda @@ -13,10 +13,11 @@ open import Data.Maybe.Base using (Maybe; map; nothing; just; fromMaybe) open import Function using (id; _∘_) open import Relation.Binary.Definitions using (tri<; tri≈; tri>) -open import Util using (_<_<_; _<_≤_; toℕ<; Ordering; less; equal; greater; compare) +open import FinMerge.Util using (_<_≤_; Ordering; compare) open import FinMerge using (merge; unmerge; pluck; glue-once; glue-unglue-once; glue-iter; unpluck) +open Ordering private variable diff --git a/Util.agda b/FinMerge/Util.agda index d7053c5..6a7f7b4 100644 --- a/Util.agda +++ b/FinMerge/Util.agda @@ -1,5 +1,5 @@ {-# OPTIONS --without-K --safe #-} -module Util where +module FinMerge.Util where open import Data.Fin using (Fin; toℕ) open import Data.Nat using (ℕ; _≤_; _<_ ; z<s; s≤s) diff --git a/Functor/Cartesian/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda b/Functor/Cartesian/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda new file mode 100644 index 0000000..346999b --- /dev/null +++ b/Functor/Cartesian/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda @@ -0,0 +1,169 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --lossy-unification #-} + +open import Level using (Level) + +module Functor.Cartesian.Instance.Underlying.SymmetricMonoidal.FinitelyCocomplete {o ℓ e : Level} where + +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Category.Monoidal.Utilities as ⊗-Util + +open import Categories.Category.Cartesian.Bundle using (CartesianCategory) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Category.Product using (_※_) +open import Categories.Category.Product.Properties using () renaming (project₁ to p₁; project₂ to p₂; unique to u) +open import Categories.Functor.Core using (Functor) +open import Categories.Functor.Cartesian using (CartesianF) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.NaturalTransformation.Core using (ntHelper) +open import Categories.NaturalTransformation.NaturalIsomorphism.Monoidal.Symmetric using (module Lax) +open import Categories.Object.Product using (IsProduct) +open import Categories.Object.Terminal using (IsTerminal) +open import Data.Product.Base using (_,_; <_,_>; proj₁; proj₂) + +open import Category.Cartesian.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes-CC) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Category.Cartesian.Instance.SymMonCat {o} {ℓ} {e} using (SymMonCat-CC) +open import Functor.Instance.Underlying.SymmetricMonoidal.FinitelyCocomplete {o} {ℓ} {e} using () renaming (Underlying to U) + +module CartesianCategory′ {o ℓ e : Level} (C : CartesianCategory o ℓ e) where + module CC = CartesianCategory C + open import Categories.Object.Terminal using (Terminal) + open Terminal CC.terminal public + open import Categories.Category.BinaryProducts using (BinaryProducts) + open BinaryProducts CC.products public + open CC public + +module FC = CartesianCategory′ FinitelyCocompletes-CC +module SMC = CartesianCategory′ SymMonCat-CC +module U = Functor U + +F-resp-⊤ : IsTerminal SMC.U (U.₀ FC.⊤) +F-resp-⊤ = _ + +F-resp-× : {A B : FC.Obj} → IsProduct SMC.U (U.₁ (FC.π₁ {A} {B})) (U.₁ (FC.π₂ {A} {B})) +F-resp-× {A} {B} = record + { ⟨_,_⟩ = pairing + ; project₁ = λ { {C} {F} {G} → project₁ {C} F G } + ; project₂ = λ { {C} {F} {G} → project₂ {C} F G } + ; unique = λ { {C} {H} {F} {G} ≃₁ ≃₂ → unique {C} F G H ≃₁ ≃₂ } + } + where + module _ {C : SMC.Obj} (F : Lax.SymmetricMonoidalFunctor C (U.₀ A)) (G : Lax.SymmetricMonoidalFunctor C (U.₀ B)) where + module F = Lax.SymmetricMonoidalFunctor F + module G = Lax.SymmetricMonoidalFunctor G + pairing : Lax.SymmetricMonoidalFunctor C (U.₀ (A FC.× B)) + pairing = record + { F = F.F ※ G.F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = F.ε , G.ε + ; ⊗-homo = ntHelper record + { η = < F.⊗-homo.η , G.⊗-homo.η > + ; commute = < F.⊗-homo.commute , G.⊗-homo.commute > + } + ; associativity = F.associativity , G.associativity + ; unitaryˡ = F.unitaryˡ , G.unitaryˡ + ; unitaryʳ = F.unitaryʳ , G.unitaryʳ + } + ; braiding-compat = F.braiding-compat , G.braiding-compat + } + } + module pairing = Lax.SymmetricMonoidalFunctor pairing + module A = FinitelyCocompleteCategory A + module B = FinitelyCocompleteCategory B + module C = SymmetricMonoidalCategory C + module A′ = SymmetricMonoidalCategory (U.₀ A) + module B′ = SymmetricMonoidalCategory (U.₀ B) + project₁ : Lax.SymmetricMonoidalNaturalIsomorphism ((U.₁ (FC.π₁ {A} {B})) SMC.∘ pairing) F + project₁ = record + { U = p₁ {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {C.U} {F.F} {G.F} + ; F⇒G-isMonoidal = record + { ε-compat = ¡-unique₂ (id ∘ F.ε ∘ ¡) F.ε + ; ⊗-homo-compat = λ { {X} {Y} → identityˡ ○ refl⟩∘⟨ sym ([]-cong₂ identityʳ identityʳ) } + } + } + where + open FinitelyCocompleteCategory A + open HomReasoning + open Equiv + open ⇒-Reasoning A.U + project₂ : Lax.SymmetricMonoidalNaturalIsomorphism (U.₁ {A FC.× B} {B} FC.π₂ SMC.∘ pairing) G + project₂ = record + { U = p₂ {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {C.U} {F.F} {G.F} + ; F⇒G-isMonoidal = record + { ε-compat = ¡-unique₂ (id ∘ G.ε ∘ ¡) G.ε + ; ⊗-homo-compat = λ { {X} {Y} → identityˡ ○ refl⟩∘⟨ sym ([]-cong₂ identityʳ identityʳ) } + } + } + where + open FinitelyCocompleteCategory B + open HomReasoning + open Equiv + open ⇒-Reasoning B.U + unique + : (H : Lax.SymmetricMonoidalFunctor C (U.F₀ (A FC.× B))) + → Lax.SymmetricMonoidalNaturalIsomorphism (U.₁ {A FC.× B} {A} FC.π₁ SMC.∘ H) F + → Lax.SymmetricMonoidalNaturalIsomorphism (U.₁ {A FC.× B} {B} FC.π₂ SMC.∘ H) G + → Lax.SymmetricMonoidalNaturalIsomorphism pairing H + unique H ≃₁ ≃₂ = record + { U = u {o} {ℓ} {e} {o} {ℓ} {e} {o} {ℓ} {e} {A.U} {B.U} {C.U} {F.F} {G.F} {H.F} ≃₁.U ≃₂.U + ; F⇒G-isMonoidal = record + { ε-compat = ε-compat₁ , ε-compat₂ + ; ⊗-homo-compat = + λ { {X} {Y} → ⊗-homo-compat₁ {X} {Y} , ⊗-homo-compat₂ {X} {Y} } + } + } + where + module H = Lax.SymmetricMonoidalFunctor H + module ≃₁ = Lax.SymmetricMonoidalNaturalIsomorphism ≃₁ + module ≃₂ = Lax.SymmetricMonoidalNaturalIsomorphism ≃₂ + ε-compat₁ : ≃₁.⇐.η C.unit A.∘ F.ε A.≈ proj₁ H.ε + ε-compat₁ = refl⟩∘⟨ sym ≃₁.ε-compat ○ cancelˡ (≃₁.iso.isoˡ C.unit) ○ ¡-unique₂ (proj₁ H.ε ∘ ¡) (proj₁ H.ε) + where + open A + open HomReasoning + open ⇒-Reasoning A.U + open Equiv + ε-compat₂ : ≃₂.⇐.η C.unit B.∘ G.ε B.≈ proj₂ H.ε + ε-compat₂ = refl⟩∘⟨ sym ≃₂.ε-compat ○ cancelˡ (≃₂.iso.isoˡ C.unit) ○ ¡-unique₂ (proj₂ H.ε ∘ ¡) (proj₂ H.ε) + where + open B + open HomReasoning + open ⇒-Reasoning B.U + open Equiv + ⊗-homo-compat₁ + : {X Y : C.Obj} + → ≃₁.⇐.η (C.⊗.₀ (X , Y)) + A.∘ F.⊗-homo.η (X , Y) + A.≈ proj₁ (H.⊗-homo.η (X , Y)) + A.∘ (≃₁.⇐.η X A.+₁ ≃₁.⇐.η Y) + ⊗-homo-compat₁ {X} {Y} = switch-fromtoʳ (≃₁.FX≅GX ⊗ᵢ ≃₁.FX≅GX) (assoc ○ sym (switch-fromtoˡ ≃₁.FX≅GX (refl⟩∘⟨ introʳ A.+-η ○ ≃₁.⊗-homo-compat))) + where + open A + open HomReasoning + open Equiv + open ⇒-Reasoning A.U + open ⊗-Util A′.monoidal + ⊗-homo-compat₂ + : {X Y : C.Obj} + → ≃₂.⇐.η (C.⊗.₀ (X , Y)) + B.∘ G.⊗-homo.η (X , Y) + B.≈ proj₂ (H.⊗-homo.η (X , Y)) + B.∘ (≃₂.⇐.η X B.+₁ ≃₂.⇐.η Y) + ⊗-homo-compat₂ = switch-fromtoʳ (≃₂.FX≅GX ⊗ᵢ ≃₂.FX≅GX) (assoc ○ sym (switch-fromtoˡ ≃₂.FX≅GX (refl⟩∘⟨ introʳ B.+-η ○ ≃₂.⊗-homo-compat))) + where + open B + open HomReasoning + open Equiv + open ⇒-Reasoning B.U + open ⊗-Util B′.monoidal + +Underlying : CartesianF FinitelyCocompletes-CC SymMonCat-CC +Underlying = record + { F = U + ; isCartesian = record + { F-resp-⊤ = F-resp-⊤ + ; F-resp-× = F-resp-× + } + } diff --git a/Functor/Exact/Instance/Swap.agda b/Functor/Exact/Instance/Swap.agda new file mode 100644 index 0000000..99a27c5 --- /dev/null +++ b/Functor/Exact/Instance/Swap.agda @@ -0,0 +1,79 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + +module Functor.Exact.Instance.Swap {o ℓ e : Level} (𝒞 𝒟 : FinitelyCocompleteCategory o ℓ e) where + +open import Categories.Category using (_[_,_]) +open import Categories.Category.BinaryProducts using (BinaryProducts) +open import Categories.Category.Product using (Product) renaming (Swap to Swap′) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Diagram.Coequalizer using (IsCoequalizer) +open import Categories.Object.Initial using (IsInitial) +open import Categories.Object.Coproduct using (IsCoproduct) +open import Data.Product.Base using (_,_; proj₁; proj₂; swap) + +open import Category.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes-Cartesian) +open import Functor.Exact using (RightExactFunctor) + +module FCC = Cartesian FinitelyCocompletes-Cartesian +open BinaryProducts (FCC.products) using (_×_) -- ; π₁; π₂; _⁂_; assocˡ) + + +module 𝒞 = FinitelyCocompleteCategory 𝒞 +module 𝒟 = FinitelyCocompleteCategory 𝒟 + +swap-resp-⊥ : {A : 𝒞.Obj} {B : 𝒟.Obj} → IsInitial (Product 𝒞.U 𝒟.U) (A , B) → IsInitial (Product 𝒟.U 𝒞.U) (B , A) +swap-resp-⊥ {A} {B} isInitial = record + { ! = swap ! + ; !-unique = λ { (f , g) → swap (!-unique (g , f)) } + } + where + open IsInitial isInitial + +swap-resp-+ + : {A₁ B₁ A+B₁ : 𝒞.Obj} + → {A₂ B₂ A+B₂ : 𝒟.Obj} + → {i₁₁ : 𝒞.U [ A₁ , A+B₁ ]} + → {i₂₁ : 𝒞.U [ B₁ , A+B₁ ]} + → {i₁₂ : 𝒟.U [ A₂ , A+B₂ ]} + → {i₂₂ : 𝒟.U [ B₂ , A+B₂ ]} + → IsCoproduct (Product 𝒞.U 𝒟.U) (i₁₁ , i₁₂) (i₂₁ , i₂₂) + → IsCoproduct (Product 𝒟.U 𝒞.U) (i₁₂ , i₁₁) (i₂₂ , i₂₁) +swap-resp-+ {A₁} {B₁} {A+B₁} {A₂} {B₂} {A+B₂} {i₁₁} {i₂₁} {i₁₂} {i₂₂} isCoproduct = record + { [_,_] = λ { X Y → swap ([ swap X , swap Y ]) } + ; inject₁ = swap inject₁ + ; inject₂ = swap inject₂ + ; unique = λ { ≈₁ ≈₂ → swap (unique (swap ≈₁) (swap ≈₂)) } + } + where + open IsCoproduct isCoproduct + +swap-resp-coeq + : {A₁ B₁ C₁ : 𝒞.Obj} + {A₂ B₂ C₂ : 𝒟.Obj} + {f₁ g₁ : 𝒞.U [ A₁ , B₁ ]} + {h₁ : 𝒞.U [ B₁ , C₁ ]} + {f₂ g₂ : 𝒟.U [ A₂ , B₂ ]} + {h₂ : 𝒟.U [ B₂ , C₂ ]} + → IsCoequalizer (Product 𝒞.U 𝒟.U) (f₁ , f₂) (g₁ , g₂) (h₁ , h₂) + → IsCoequalizer (Product 𝒟.U 𝒞.U) (f₂ , f₁) (g₂ , g₁) (h₂ , h₁) +swap-resp-coeq isCoequalizer = record + { equality = swap equality + ; coequalize = λ { x → swap (coequalize (swap x)) } + ; universal = swap universal + ; unique = λ { x → swap (unique (swap x)) } + } + where + open IsCoequalizer isCoequalizer + +Swap : RightExactFunctor (𝒞 × 𝒟) (𝒟 × 𝒞) +Swap = record + { F = Swap′ + ; isRightExact = record + { F-resp-⊥ = swap-resp-⊥ + ; F-resp-+ = swap-resp-+ + ; F-resp-coeq = swap-resp-coeq + } + } diff --git a/Functor/Forgetful/Instance/CMonoid.agda b/Functor/Forgetful/Instance/CMonoid.agda new file mode 100644 index 0000000..fd8ecc1 --- /dev/null +++ b/Functor/Forgetful/Instance/CMonoid.agda @@ -0,0 +1,36 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category using (Category) +open import Categories.Category.Monoidal using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) +open import Level using (Level) + +module Functor.Forgetful.Instance.CMonoid + {o ℓ e : Level} + {S : Category o ℓ e} + {monoidal : Monoidal S} + (symmetric : Symmetric monoidal) + where + +open import Categories.Category.Construction.Monoids monoidal using (Monoids) +open import Categories.Functor using (Functor) +open import Category.Construction.CMonoids symmetric using (CMonoids) +open import Function using (id) +open import Object.Monoid.Commutative using (CommutativeMonoid; CommutativeMonoid⇒) + +private + module S = Category S + +open CommutativeMonoid +open CommutativeMonoid⇒ +open Functor +open S.Equiv using (refl) + +Forget : Functor CMonoids Monoids +Forget .F₀ = monoid +Forget .F₁ = monoid⇒ +Forget .identity = refl +Forget .homomorphism = refl +Forget .F-resp-≈ = id + +module Forget = Functor Forget diff --git a/Functor/Forgetful/Instance/Monoid.agda b/Functor/Forgetful/Instance/Monoid.agda new file mode 100644 index 0000000..2f9e4d8 --- /dev/null +++ b/Functor/Forgetful/Instance/Monoid.agda @@ -0,0 +1,29 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category using (Category) +open import Categories.Category.Monoidal using (Monoidal) +open import Level using (Level) + +module Functor.Forgetful.Instance.Monoid {o ℓ e : Level} {S : Category o ℓ e} (monoidal : Monoidal S) where + +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Functor using (Functor) +open import Categories.Object.Monoid using (Monoid; Monoid⇒) +open import Function using (id) + +private + module S = Category S + +open Monoid +open Monoid⇒ +open S.Equiv using (refl) +open Functor + +Forget : Functor (Monoids monoidal) S +Forget .F₀ = Carrier +Forget .F₁ = arr +Forget .identity = refl +Forget .homomorphism = refl +Forget .F-resp-≈ = id + +module Forget = Functor Forget diff --git a/Functor/Free/Instance/CMonoid.agda b/Functor/Free/Instance/CMonoid.agda new file mode 100644 index 0000000..be9cb94 --- /dev/null +++ b/Functor/Free/Instance/CMonoid.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module Functor.Free.Instance.CMonoid {c ℓ : Level} where + +import Categories.Object.Monoid as MonoidObject +import Object.Monoid.Commutative as CMonoidObject + +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor using (Functor) +open import Categories.NaturalTransformation using (NaturalTransformation) +open import Category.Construction.CMonoids using (CMonoids) +open import Category.Instance.Setoids.SymmetricMonoidal {c} {c ⊔ ℓ} using (Setoids-×; ×-symmetric′) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (++-assoc; ++-identityˡ; ++-identityʳ; ++-comm) +open import Data.Product using (_,_) +open import Data.Setoid using (∣_∣) +open import Data.Opaque.Multiset using ([]ₛ; Multisetₛ; ++ₛ; mapₛ) +open import Function using (_⟶ₛ_; _⟨$⟩_) +open import Functor.Instance.Multiset {c} {ℓ} using (Multiset) +open import NaturalTransformation.Instance.EmptyMultiset {c} {ℓ} using (⊤⇒[]) +open import NaturalTransformation.Instance.MultisetAppend {c} {ℓ} using (++) +open import Relation.Binary using (Setoid) + +module ++ = NaturalTransformation ++ +module ⊤⇒[] = NaturalTransformation ⊤⇒[] + +open Functor +open MonoidObject Setoids-×.monoidal using (Monoid; IsMonoid; Monoid⇒) +open CMonoidObject Setoids-×.symmetric using (CommutativeMonoid; IsCommutativeMonoid; CommutativeMonoid⇒) +open IsCommutativeMonoid +open CommutativeMonoid using () renaming (μ to μ′; η to η′) +open IsMonoid +open CommutativeMonoid⇒ +open Monoid⇒ + +module _ (X : Setoid c ℓ) where + + open Setoid (Multiset.₀ X) + + opaque + + unfolding Multisetₛ + + ++ₛ-assoc + : (x y z : ∣ Multisetₛ X ∣) + → ++ₛ ⟨$⟩ (++ₛ ⟨$⟩ (x , y) , z) + ≈ ++ₛ ⟨$⟩ (x , ++ₛ ⟨$⟩ (y , z)) + ++ₛ-assoc x y z = ++-assoc X x y z + + ++ₛ-identityˡ + : (x : ∣ Multisetₛ X ∣) + → x ≈ ++ₛ ⟨$⟩ ([]ₛ ⟨$⟩ _ , x) + ++ₛ-identityˡ x = ++-identityˡ X x + + ++ₛ-identityʳ + : (x : ∣ Multisetₛ X ∣) + → x ≈ ++ₛ ⟨$⟩ (x , []ₛ ⟨$⟩ _) + ++ₛ-identityʳ x = sym (++-identityʳ X x) + + ++ₛ-comm + : (x y : ∣ Multisetₛ X ∣) + → ++ₛ ⟨$⟩ (x , y) ≈ ++ₛ ⟨$⟩ (y , x) + ++ₛ-comm x y = ++-comm X x y + + opaque + unfolding ×-symmetric′ + MultisetCMonoid : IsCommutativeMonoid (Multiset.₀ X) + MultisetCMonoid .isMonoid .μ = ++.η X + MultisetCMonoid .isMonoid .η = ⊤⇒[].η X + MultisetCMonoid .isMonoid .assoc {(x , y) , z} = ++ₛ-assoc x y z + MultisetCMonoid .isMonoid .identityˡ {_ , x} = ++ₛ-identityˡ x + MultisetCMonoid .isMonoid .identityʳ {x , _} = ++ₛ-identityʳ x + MultisetCMonoid .commutative {x , y} = ++ₛ-comm x y + +Multisetₘ : (X : Setoid c ℓ) → CommutativeMonoid +Multisetₘ X = record { isCommutativeMonoid = MultisetCMonoid X } + +open Setoids-× using (_⊗₀_; _⊗₁_) +opaque + unfolding MultisetCMonoid + mapₛ-++ₛ + : {A B : Setoid c ℓ} + → (f : A ⟶ₛ B) + → {xy : ∣ Multisetₛ A ⊗₀ Multisetₛ A ∣} + → (open Setoid (Multisetₛ B)) + → mapₛ f ⟨$⟩ (μ′ (Multisetₘ A) ⟨$⟩ xy) + ≈ μ′ (Multisetₘ B) ⟨$⟩ (mapₛ f ⊗₁ mapₛ f ⟨$⟩ xy) + mapₛ-++ₛ = ++.sym-commute + +opaque + unfolding MultisetCMonoid mapₛ + mapₛ-[]ₛ + : {A B : Setoid c ℓ} + → (f : A ⟶ₛ B) + → {x : ∣ Setoids-×.unit ∣} + → (open Setoid (Multisetₛ B)) + → mapₛ f ⟨$⟩ (η′ (Multisetₘ A) ⟨$⟩ x) + ≈ η′ (Multisetₘ B) ⟨$⟩ x + mapₛ-[]ₛ = ⊤⇒[].commute + +mapₘ + : {A B : Setoid c ℓ} + (f : A ⟶ₛ B) + → CommutativeMonoid⇒ (Multisetₘ A) (Multisetₘ B) +mapₘ f .monoid⇒ .arr = Multiset.₁ f +mapₘ f .monoid⇒ .preserves-μ = mapₛ-++ₛ f +mapₘ f .monoid⇒ .preserves-η = mapₛ-[]ₛ f + +Free : Functor (Setoids c ℓ) (CMonoids Setoids-×.symmetric) +Free .F₀ = Multisetₘ +Free .F₁ = mapₘ +Free .identity {X} = Multiset.identity {X} +Free .homomorphism {X} {Y} {Z} {f} {g} = Multiset.homomorphism {X} {Y} {Z} {f} {g} +Free .F-resp-≈ {A} {B} {f} {g} = Multiset.F-resp-≈ {A} {B} {f} {g} diff --git a/Functor/Free/Instance/Monoid.agda b/Functor/Free/Instance/Monoid.agda new file mode 100644 index 0000000..c8450b9 --- /dev/null +++ b/Functor/Free/Instance/Monoid.agda @@ -0,0 +1,116 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_; suc) + +module Functor.Free.Instance.Monoid {c ℓ : Level} where + +import Categories.Object.Monoid as MonoidObject + +open import Categories.Category using (Category) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor using (Functor) +open import Categories.NaturalTransformation using (NaturalTransformation) +open import Category.Instance.Setoids.SymmetricMonoidal {c} {c ⊔ ℓ} using (Setoids-×; ×-monoidal′) +open import Data.List.Properties using (++-assoc; ++-identityˡ; ++-identityʳ) +open import Data.Opaque.List using ([]ₛ; Listₛ; ++ₛ; mapₛ) +open import Data.Product using (_,_) +open import Data.Setoid using (∣_∣) +open import Function using (_⟶ₛ_; _⟨$⟩_) +open import Functor.Instance.List {c} {ℓ} using (List) +open import NaturalTransformation.Instance.EmptyList {c} {ℓ} using (⊤⇒[]) +open import NaturalTransformation.Instance.ListAppend {c} {ℓ} using (++) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) + +module ++ = NaturalTransformation ++ +module ⊤⇒[] = NaturalTransformation ⊤⇒[] + +open Functor +open MonoidObject Setoids-×.monoidal using (Monoid; IsMonoid; Monoid⇒) + +open IsMonoid + +-- the functor sending a setoid A to the monoid List A + +module _ (X : Setoid c ℓ) where + + open Setoid (List.₀ X) + + opaque + + unfolding []ₛ + + ++ₛ-assoc + : (x y z : ∣ Listₛ X ∣) + → ++ₛ ⟨$⟩ (++ₛ ⟨$⟩ (x , y) , z) + ≈ ++ₛ ⟨$⟩ (x , ++ₛ ⟨$⟩ (y , z)) + ++ₛ-assoc x y z = reflexive (++-assoc x y z) + + ++ₛ-identityˡ + : (x : ∣ Listₛ X ∣) + → x ≈ ++ₛ ⟨$⟩ ([]ₛ ⟨$⟩ _ , x) + ++ₛ-identityˡ x = reflexive (++-identityˡ x) + + ++ₛ-identityʳ + : (x : ∣ Listₛ X ∣) + → x ≈ ++ₛ ⟨$⟩ (x , []ₛ ⟨$⟩ _) + ++ₛ-identityʳ x = sym (reflexive (++-identityʳ x)) + + opaque + + unfolding ×-monoidal′ + + ListMonoid : IsMonoid (List.₀ X) + ListMonoid = record + { μ = ++.η X + ; η = ⊤⇒[].η X + ; assoc = λ { {(x , y) , z} → ++ₛ-assoc x y z } + ; identityˡ = λ { {_ , x} → ++ₛ-identityˡ x } + ; identityʳ = λ { {x , _} → ++ₛ-identityʳ x } + } + +Listₘ : Setoid c ℓ → Monoid +Listₘ X = record { isMonoid = ListMonoid X } + +opaque + + unfolding ListMonoid + + mapₘ + : {Aₛ Bₛ : Setoid c ℓ} + (f : Aₛ ⟶ₛ Bₛ) + → Monoid⇒ (Listₘ Aₛ) (Listₘ Bₛ) + mapₘ f = record + { arr = List.₁ f + ; preserves-μ = λ {x,y} → ++.sym-commute f {x,y} + ; preserves-η = ⊤⇒[].sym-commute f + } + +module U = Category Setoids-×.U +open Monoid⇒ using (arr) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Composition using () renaming (function to compose) +opaque + unfolding mapₘ + Free-identity : {X : Setoid c ℓ} → arr (mapₘ (Id X)) U.≈ U.id + Free-identity = List.identity + + Free-homomorphism : {X Y Z : Setoid c ℓ} {f : X ⟶ₛ Y} {g : Y ⟶ₛ Z} → arr (mapₘ (compose f g)) U.≈ arr (mapₘ g) U.∘ arr (mapₘ f) + Free-homomorphism = List.homomorphism + + Free-resp-≈ + : {X Y : Setoid c ℓ} + {f g : X ⟶ₛ Y} + (let module Y = Setoid Y) + → (∀ {x} → f ⟨$⟩ x Y.≈ g ⟨$⟩ x) + → arr (mapₘ f) U.≈ arr (mapₘ g) + Free-resp-≈ = List.F-resp-≈ + +Free : Functor (Setoids c ℓ) (Monoids Setoids-×.monoidal) +Free .F₀ = Listₘ +Free .F₁ = mapₘ +Free .identity = Free-identity +Free .homomorphism = Free-homomorphism +Free .F-resp-≈ = Free-resp-≈ diff --git a/Functor/Instance/Cospan/Embed.agda b/Functor/Instance/Cospan/Embed.agda index 77f0361..6dbc04a 100644 --- a/Functor/Instance/Cospan/Embed.agda +++ b/Functor/Instance/Cospan/Embed.agda @@ -1,4 +1,5 @@ {-# OPTIONS --without-K --safe #-} +{-# OPTIONS --hidden-argument-puns #-} open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) @@ -14,9 +15,10 @@ 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 Category.Diagram.Cospan 𝒞 using (cospan) open import Data.Product.Base using (_,_) open import Function.Base using (id) -open import Functor.Instance.Cospan.Stack using (⊗) +open import Functor.Instance.Cospan.Stack 𝒞 using (⊗) module 𝒞 = FinitelyCocompleteCategory 𝒞 module Cospans = Category Cospans @@ -28,24 +30,26 @@ 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 - } +private + variable + A B C : 𝒞.Obj + W X Y Z : 𝒞.Obj + +L₁ : U [ A , B ] → Cospans [ A , B ] +L₁ f = cospan f 𝒞.id -L-identity : {A : 𝒞.Obj} → L₁ 𝒞.id ≈ Cospans.id {A} +L-identity : L₁ 𝒞.id ≈ Cospans.id {A} L-identity = record { ≅N = ≅.refl - ; from∘f₁≈f₁′ = 𝒞.identity² - ; from∘f₂≈f₂′ = 𝒞.identity² + ; 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 : {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ʳ + ; 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 @@ -57,11 +61,11 @@ L-homomorphism {X} {Y} {Z} {f} {g} = record 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-≈ : {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² + ; from∘f₁≈f₁ = 𝒞.identityˡ ○ f≈g + ; from∘f₂≈f₂ = 𝒞.identity² } where open 𝒞.HomReasoning @@ -75,24 +79,21 @@ L = record ; F-resp-≈ = L-resp-≈ } -R₁ : {A B : 𝒞.Obj} → U [ B , A ] → Cospans [ A , B ] -R₁ g = record - { f₁ = 𝒞.id - ; f₂ = g - } +R₁ : U [ B , A ] → Cospans [ A , B ] +R₁ g = cospan 𝒞.id g -R-identity : {A : 𝒞.Obj} → R₁ 𝒞.id ≈ Cospans.id {A} +R-identity : R₁ 𝒞.id ≈ Cospans.id {A} R-identity = record { ≅N = ≅.refl - ; from∘f₁≈f₁′ = 𝒞.identity² - ; from∘f₂≈f₂′ = 𝒞.identity² + ; 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 +R-homomorphism : {f : U [ Y , X ]} {g : U [ Z , Y ]} → R₁ (U [ f ∘ g ]) ≈ Cospans [ R₁ g ∘ R₁ f ] +R-homomorphism {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}) + ; 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 @@ -104,11 +105,11 @@ R-homomorphism {X} {Y} {Z} {f} {g} = record 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 +R-resp-≈ : {f g : U [ A , B ]} → U [ f ≈ g ] → Cospans [ R₁ f ≈ R₁ g ] +R-resp-≈ {f} {g} f≈g = record { ≅N = ≅.refl - ; from∘f₁≈f₁′ = 𝒞.identity² - ; from∘f₂≈f₂′ = 𝒞.identityˡ ○ f≈g + ; from∘f₁≈f₁ = 𝒞.identity² + ; from∘f₂≈f₂ = 𝒞.identityˡ ○ f≈g } where open 𝒞.HomReasoning @@ -122,17 +123,11 @@ R = record ; 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 +B∘L : {f : U [ W , X ]} {g : U [ X , Y ]} {h : U [ Z , Y ]} → Cospans [ cospan g h ∘ L₁ f ] ≈ cospan (U [ g ∘ f ]) h +B∘L {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ˡ + ; 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 @@ -144,11 +139,11 @@ B∘L {W} {X} {Y} {Z} {f} {g} {h} = record 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 +R∘B : {f : U [ W , X ]} {g : U [ Y , X ]} {h : U [ Z , Y ]} → Cospans [ R₁ h ∘ cospan f g ] ≈ cospan f (U [ g ∘ h ]) +R∘B {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}) + ; 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 @@ -164,20 +159,18 @@ 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 + ≅-L-R : (X≅Y : X ≅ Y) → L₁ (to X≅Y) ≈ R₁ (from X≅Y) + ≅-L-R X≅Y = record { ≅N = X≅Y - ; from∘f₁≈f₁′ = isoʳ X≅Y - ; from∘f₂≈f₂′ = 𝒞.identityʳ + ; 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 +L-resp-⊗ : {a : U [ W , X ]} {b : U [ Y , Z ]} → L₁ (a +₁ b) ≈ ⊗.₁ (L₁ a , L₁ b) +L-resp-⊗ {a} {b} = record { ≅N = ≅.refl - ; from∘f₁≈f₁′ = 𝒞.identityˡ - ; from∘f₂≈f₂′ = 𝒞.identityˡ ○ sym +-η ○ sym ([]-cong₂ identityʳ identityʳ) + ; from∘f₁≈f₁ = 𝒞.identityˡ + ; from∘f₂≈f₂ = 𝒞.identityˡ ○ sym +-η ○ sym ([]-cong₂ identityʳ identityʳ) } where open 𝒞.HomReasoning diff --git a/Functor/Instance/Cospan/Stack.agda b/Functor/Instance/Cospan/Stack.agda index b7664dc..b72219b 100644 --- a/Functor/Instance/Cospan/Stack.agda +++ b/Functor/Instance/Cospan/Stack.agda @@ -9,11 +9,13 @@ 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.Category using (Category) +open import Categories.Functor using (Functor) open import Categories.Functor.Bifunctor using (Bifunctor) -open import Category.Instance.Cospans 𝒞 using (Cospan; Cospans; Same; id-Cospan; compose) +open import Category.Instance.Cospans 𝒞 using (Cospans) +open import Category.Diagram.Cospan 𝒞 as Cospan using (Cospan; identity; compose; _⊗_) open import Category.Instance.FinitelyCocompletes {o} {ℓ} {e} using () renaming (_×_ to _×′_) -open import Category.Instance.Properties.FinitelyCocompletes {o} {ℓ} {e} using (-+-; FinitelyCocompletes-CC) +open import Category.Cartesian.Instance.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) @@ -32,27 +34,19 @@ 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 : Obj} → identity {A} ⊗ identity {B} Cospan.≈ identity {A + B} id⊗id≈id {A} {B} = record { ≅N = ≅.refl - ; from∘f₁≈f₁′ = from∘f≈f′ - ; from∘f₂≈f₂′ = from∘f≈f′ + ; 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 + 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₂ ] ≈⟨ +-η ⟩ @@ -64,14 +58,14 @@ homomorphism → (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′) ) + → compose A⇒B B⇒C ⊗ compose A⇒B′ B⇒C′ Cospan.≈ compose (A⇒B ⊗ A⇒B′) (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₂′ + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where - open Cospan + open Cospan.Cospan open Pushout open HomReasoning open ⇒-Reasoning U @@ -89,56 +83,62 @@ homomorphism A⇒B B⇒C A⇒B′ B⇒C′ = record 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∘f₁≈f₁ : from ≅N ∘ (f₁ (compose A⇒B B⇒C) +₁ f₁ (compose A⇒B′ B⇒C′)) ≈ f₁ (compose (A⇒B ⊗ A⇒B′) (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 ∘ i₁ P₃′ ∘ f₁ (A⇒B ⊗ A⇒B′) ≈⟨ pullˡ (universal∘i₁≈h₁ P₃′) ⟩ + i₁ P₃ ∘ f₁ (A⇒B ⊗ A⇒B′) ∎ + from∘f₂≈f₂ : from ≅N ∘ (f₂ (compose A⇒B B⇒C) +₁ f₂ (compose A⇒B′ B⇒C′)) ≈ f₂ (compose (A⇒B ⊗ A⇒B′) (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′) ∎ + from ≅N ∘ i₂ P₃′ ∘ f₂ (B⇒C ⊗ B⇒C′) ≈⟨ pullˡ (universal∘i₂≈h₂ P₃′) ⟩ + i₂ P₃ ∘ f₂ (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′) + → f Cospan.≈ f′ + → g Cospan.≈ g′ + → f ⊗ g Cospan.≈ 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₂′ + ; from∘f₁≈f₁ = from∘f₁≈f₁ + ; from∘f₂≈f₂ = from∘f₂≈f₂ } where open 𝒞 using (-+-) - module ≈f = Same ≈f - module ≈g = Same ≈g + module ≈f = Cospan._≈_ ≈f + module ≈g = Cospan._≈_ ≈g open HomReasoning - open Cospan + open Cospan.Cospan open 𝒞 using (+₁-cong₂; +₁∘+₁) - from∘f₁≈f₁′ : (≈f.from +₁ ≈g.from) ∘ (f₁ f +₁ f₁ g) ≈ f₁ f′ +₁ f₁ g′ - from∘f₁≈f₁′ = begin + 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.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 + 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.from ∘ f₂ f) +₁ (≈g.from ∘ f₂ g) ≈⟨ +₁-cong₂ ≈f.from∘f₂≈f₂ ≈g.from∘f₂≈f₂ ⟩ f₂ f′ +₁ f₂ g′ ∎ +private + ⊗′ : Bifunctor Cospans Cospans Cospans + ⊗′ = record + { F₀ = λ (A , A′) → A + A′ + ; F₁ = λ (f , g) → 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 + } + ⊗ : 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 } - } +⊗ = ⊗′ + +module ⊗ = Functor ⊗ diff --git a/Functor/Instance/Decorate.agda b/Functor/Instance/Decorate.agda new file mode 100644 index 0000000..fedddba --- /dev/null +++ b/Functor/Instance/Decorate.agda @@ -0,0 +1,165 @@ +{-# 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) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Data.Product.Base using (_,_) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module Functor.Instance.Decorate + {o o′ ℓ ℓ′ e e′} + (𝒞 : FinitelyCocompleteCategory o ℓ e) + {𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′} + (F : SymmetricMonoidalFunctor (smc 𝒞) 𝒟) where + +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) +open import Categories.Category.Monoidal.Properties using (coherence-inv₃) +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 (Cospans) +open import Category.Instance.DecoratedCospans 𝒞 F using (DecoratedCospans) +open import Functor.Instance.Cospan.Stack 𝒞 using (module ⊗) +open import Functor.Instance.DecoratedCospan.Stack 𝒞 F using () renaming (module ⊗ to ⊗′) + +module 𝒞 = FinitelyCocompleteCategory 𝒞 +module 𝒟 = SymmetricMonoidalCategory 𝒟 +module F = SymmetricMonoidalFunctor F +module Cospans = Category Cospans +module DecoratedCospans = Category DecoratedCospans +module mc𝒞 = CocartesianMonoidal 𝒞.U 𝒞.cocartesian + +-- For every cospan there exists a free decorated cospan +-- i.e. the original cospan with the discrete decoration + +private + variable + A A′ B B′ C C′ D : 𝒞.Obj + f : Cospans [ A , B ] + g : Cospans [ C , D ] + +decorate : Cospans [ A , B ] → DecoratedCospans [ A , B ] +decorate f = record + { cospan = f + ; decoration = F₁ ¡ ∘ ε + } + where + open 𝒞 using (¡) + open 𝒟 using (_∘_) + open F using (ε; F₁) + +identity : DecoratedCospans [ decorate (Cospans.id {A}) ≈ DecoratedCospans.id ] +identity = record + { cospans-≈ = Cospans.Equiv.refl + ; same-deco = elimˡ F.identity + } + where + open ⇒-Reasoning 𝒟.U + +homomorphism : DecoratedCospans [ decorate (Cospans [ g ∘ f ]) ≈ DecoratedCospans [ decorate g ∘ decorate f ] ] +homomorphism {g} {f} = record + { cospans-≈ = Cospans.Equiv.refl + ; same-deco = same-deco + } + where + + 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 _+₁_ ) + open Category U + open Equiv + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + open F using (F₁; ε) + open Shorthands monoidal + + open DiagramPushout 𝒞.U using (Pushout) + open Pushout (pushout f₂ g₁) using (i₁; i₂) + open mc𝒞 using (unitorˡ) + open unitorˡ using () renaming (to to λ⇐′) + + same-deco : F₁ 𝒞.id ∘ F₁ ¡ ∘ F.ε ≈ F₁ [ i₁ , i₂ ]′ ∘ φ (N , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ + same-deco = begin + F₁ 𝒞.id ∘ F₁ ¡ ∘ ε ≈⟨ elimˡ F.identity ⟩ + F₁ ¡ ∘ ε ≈⟨ F.F-resp-≈ (¡-unique _) ⟩∘⟨refl ⟩ + F₁ ([ i₁ , i₂ ]′ 𝒞.∘ ¡ +₁ ¡ 𝒞.∘ λ⇐′) ∘ ε ≈⟨ refl⟩∘⟨ introʳ λ-.isoʳ ⟩ + 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 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 (¡ , ¡)) ⟨ + F₁ [ i₁ , i₂ ]′ ∘ φ (N , M) ∘ F₁ ¡ ⊗₁ F₁ ¡ ∘ ε ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ (sym ⊗-distrib-over-∘) ⟩ + F₁ [ i₁ , i₂ ]′ ∘ φ (N , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ∎ + +F-resp-≈ : Cospans [ f ≈ g ] → DecoratedCospans [ decorate f ≈ decorate g ] +F-resp-≈ f≈g = record + { cospans-≈ = f≈g + ; same-deco = pullˡ (sym F.homomorphism) ○ sym (F.F-resp-≈ (¡-unique _)) ⟩∘⟨refl + } + where + open ⇒-Reasoning 𝒟.U + open 𝒟.Equiv + open 𝒟.HomReasoning + open 𝒞 using (¡-unique) + +Decorate : Functor Cospans DecoratedCospans +Decorate = record + { F₀ = idf + ; F₁ = decorate + ; identity = identity + ; homomorphism = homomorphism + ; F-resp-≈ = F-resp-≈ + } + +Decorate-resp-⊗ : DecoratedCospans [ decorate (⊗.₁ (f , g)) ≈ ⊗′.₁ (decorate f , decorate g) ] +Decorate-resp-⊗ {f} {g} = record + { cospans-≈ = Cospan.≈-refl + ; same-deco = same-deco + } + where + + 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 _+₁_ ) + open Category U + open Equiv + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + open F using (F₁; ε) + open Shorthands monoidal + open mc𝒞 using (unitorˡ) + open unitorˡ using () renaming (to to λ⇐′) + + same-deco : F₁ 𝒞.id ∘ F₁ ¡ ∘ ε ≈ φ (N , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ + same-deco = begin + F₁ 𝒞.id ∘ F₁ ¡ ∘ ε ≈⟨ elimˡ F.identity ⟩ + F₁ ¡ ∘ ε ≈⟨ F.F-resp-≈ (¡-unique _) ⟩∘⟨refl ⟩ + F₁ (¡ +₁ ¡ 𝒞.∘ λ⇐′) ∘ ε ≈⟨ refl⟩∘⟨ introʳ λ-.isoʳ ⟩ + F₁ (¡ +₁ ¡ 𝒞.∘ λ⇐′) ∘ ε ∘ λ⇒ ∘ λ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ coherence-inv₃ monoidal ⟩ + F₁ (¡ +₁ ¡ 𝒞.∘ λ⇐′) ∘ ε ∘ λ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ unitorˡ-commute-from ⟨ + F₁ (¡ +₁ ¡ 𝒞.∘ λ⇐′) ∘ λ⇒ ∘ id ⊗₁ ε ∘ ρ⇐ ≈⟨ pushˡ F.homomorphism ⟩ + F₁ (¡ +₁ ¡) ∘ F₁ λ⇐′ ∘ λ⇒ ∘ id ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ unitorˡ) F.unitaryˡ) ⟨ + F₁ (¡ +₁ ¡) ∘ φ (⊥ , ⊥) ∘ ε ⊗₁ id ∘ id ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pullˡ (sym serialize₁₂) ⟩ + F₁ (¡ +₁ ¡) ∘ φ (⊥ , ⊥) ∘ ε ⊗₁ ε ∘ ρ⇐ ≈⟨ extendʳ (φ-commute (¡ , ¡)) ⟨ + φ (N , M) ∘ F₁ ¡ ⊗₁ F₁ ¡ ∘ ε ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ pullˡ (sym ⊗-distrib-over-∘) ⟩ + φ (N , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ∎ diff --git a/Functor/Instance/DecoratedCospan/Embed.agda b/Functor/Instance/DecoratedCospan/Embed.agda new file mode 100644 index 0000000..77b16fa --- /dev/null +++ b/Functor/Instance/DecoratedCospan/Embed.agda @@ -0,0 +1,275 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory; SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module Functor.Instance.DecoratedCospan.Embed + {o o′ ℓ ℓ′ e e′} + (𝒞 : FinitelyCocompleteCategory o ℓ e) + {𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′} + (F : SymmetricMonoidalFunctor (smc 𝒞) 𝒟) where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Diagram.Pushout.Properties as PushoutProperties +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Category.Diagram.Pushout as Pushout′ +import Category.Diagram.Cospan as Cospan +import Functor.Instance.Cospan.Embed 𝒞 as Embed + +open import Categories.Category using (Category; _[_,_]; _[_≈_]; _[_∘_]) +open import Categories.Category.Monoidal.Properties using (coherence-inv₃) +open import Categories.Functor.Properties using ([_]-resp-≅) +open import Category.Instance.Cospans 𝒞 using (Cospans) +open import Category.Instance.DecoratedCospans 𝒞 F using (DecoratedCospans) + +import Categories.Diagram.Pushout as DiagramPushout +import Categories.Diagram.Pushout.Properties as PushoutProperties +import Categories.Morphism as Morphism + +open import Categories.Category.Cocartesian using (module CocartesianMonoidal) +open import Categories.Category.Monoidal.Utilities using (module Shorthands) +open import Categories.Functor using (Functor; _∘F_) +open import Data.Product using (_,_) +open import Function using () renaming (id to idf) +open import Functor.Instance.DecoratedCospan.Stack 𝒞 F using (⊗) + +module 𝒞 = FinitelyCocompleteCategory 𝒞 +module 𝒟 = SymmetricMonoidalCategory 𝒟 +module F = SymmetricMonoidalFunctor F +module Cospans = Category Cospans +module DecoratedCospans = Category DecoratedCospans +module mc𝒞 = CocartesianMonoidal 𝒞.U 𝒞.cocartesian + +open import Functor.Instance.Decorate 𝒞 F using (Decorate; Decorate-resp-⊗) + +private + variable + A B C D E H : 𝒞.Obj + f : 𝒞.U [ A , B ] + g : 𝒞.U [ C , D ] + h : 𝒞.U [ E , H ] + +L : Functor 𝒞.U DecoratedCospans +L = Decorate ∘F Embed.L + +R : Functor 𝒞.op DecoratedCospans +R = Decorate ∘F Embed.R + +B₁ : 𝒞.U [ A , C ] → 𝒞.U [ B , C ] → 𝒟.U [ 𝒟.unit , F.F₀ C ] → DecoratedCospans [ A , B ] +B₁ f g s = record + { cospan = Cospan.cospan f g + ; decoration = s + } + +module _ where + + module L = Functor L + module R = Functor R + + module Codiagonal where + + open mc𝒞 using (unitorˡ; unitorʳ; +-monoidal) public + open unitorˡ using () renaming (to to λ⇐′) public + open unitorʳ using () renaming (to to ρ⇐′) public + open 𝒞 using (U; _+_; []-cong₂; []∘+₁; ∘-distribˡ-[]; inject₁; inject₂; ¡) + renaming ([_,_] to [_,_]′; _+₁_ to infixr 10 _+₁_ ) + open Category U + open Equiv + open HomReasoning + open ⇒-Reasoning 𝒞.U + + μ : {X : Obj} → X + X ⇒ X + μ = [ id , id ]′ + + μ∘+ : {X Y Z : Obj} {f : X ⇒ Z} {g : Y ⇒ Z} → [ f , g ]′ ≈ μ ∘ f +₁ g + μ∘+ = []-cong₂ (sym identityˡ) (sym identityˡ) ○ sym []∘+₁ + + μ∘¡ˡ : {X : Obj} → μ ∘ ¡ +₁ id ∘ λ⇐′ ≈ id {X} + μ∘¡ˡ = begin + μ ∘ ¡ +₁ id ∘ λ⇐′ ≈⟨ pullˡ (sym μ∘+) ⟩ + [ ¡ , id ]′ ∘ λ⇐′ ≈⟨ inject₂ ⟩ + id ∎ + + μ∘¡ʳ : {X : Obj} → μ ∘ id +₁ ¡ ∘ ρ⇐′ ≈ id {X} + μ∘¡ʳ = begin + μ ∘ id +₁ ¡ ∘ ρ⇐′ ≈⟨ pullˡ (sym μ∘+) ⟩ + [ id , ¡ ]′ ∘ ρ⇐′ ≈⟨ inject₁ ⟩ + id ∎ + + + μ-natural : {X Y : Obj} {f : X ⇒ Y} → f ∘ μ ≈ μ ∘ f +₁ f + μ-natural = ∘-distribˡ-[] ○ []-cong₂ (identityʳ ○ sym identityˡ) (identityʳ ○ sym identityˡ) ○ sym []∘+₁ + + B∘L : {A B M C : 𝒞.Obj} + → {f : 𝒞.U [ A , B ]} + → {g : 𝒞.U [ B , M ]} + → {h : 𝒞.U [ C , M ]} + → {s : 𝒟.U [ 𝒟.unit , F.₀ M ]} + → DecoratedCospans [ DecoratedCospans [ B₁ g h s ∘ L.₁ f ] ≈ B₁ (𝒞.U [ g ∘ f ]) h s ] + B∘L {A} {B} {M} {C} {f} {g} {h} {s} = record + { cospans-≈ = Embed.B∘L + ; same-deco = same-deco + } + where + + module _ where + open 𝒞 using (¡; ⊥; ¡-unique; pushout) renaming ([_,_] to [_,_]′; _+₁_ to infixr 10 _+₁_ ) + open 𝒞 using (U) + open Category U + open mc𝒞 using (unitorˡ; unitorˡ-commute-to; +-monoidal) public + open unitorˡ using () renaming (to to λ⇐′) public + open ⊗-Reasoning +-monoidal + open ⇒-Reasoning 𝒞.U + open Equiv + + open Pushout′ 𝒞.U using (pushout-id-g) + open PushoutProperties 𝒞.U using (up-to-iso) + open DiagramPushout 𝒞.U using (Pushout) + P P′ : Pushout 𝒞.id g + P = pushout 𝒞.id g + P′ = pushout-id-g + module P = Pushout P + module P′ = Pushout P′ + open Morphism 𝒞.U using (_≅_) + open _≅_ using (from) + open P using (i₁ ; i₂; universal∘i₂≈h₂) public + + open Codiagonal using (μ; μ-natural; μ∘+; μ∘¡ˡ) + + ≅ : P.Q ⇒ M + ≅ = up-to-iso P P′ .from + + ≅∘[]∘¡+id : ((≅ ∘ [ i₁ , i₂ ]′) ∘ ¡ +₁ id) ∘ λ⇐′ ≈ id + ≅∘[]∘¡+id = begin + ((≅ ∘ [ i₁ , i₂ ]′) ∘ ¡ +₁ id) ∘ λ⇐′ ≈⟨ assoc²αε ⟩ + ≅ ∘ [ i₁ , i₂ ]′ ∘ ¡ +₁ id ∘ λ⇐′ ≈⟨ refl⟩∘⟨ pushˡ μ∘+ ⟩ + ≅ ∘ μ ∘ i₁ +₁ i₂ ∘ ¡ +₁ id ∘ λ⇐′ ≈⟨ refl⟩∘⟨ pull-center (sym split₁ʳ) ⟩ + ≅ ∘ μ ∘ (i₁ ∘ ¡) +₁ i₂ ∘ λ⇐′ ≈⟨ extendʳ μ-natural ⟩ + μ ∘ ≅ +₁ ≅ ∘ (i₁ ∘ ¡) +₁ i₂ ∘ λ⇐′ ≈⟨ pull-center (sym ⊗-distrib-over-∘) ⟩ + μ ∘ (≅ ∘ i₁ ∘ ¡) +₁ (≅ ∘ i₂) ∘ λ⇐′ ≈⟨ refl⟩∘⟨ sym (¡-unique (≅ ∘ i₁ ∘ ¡)) ⟩⊗⟨ universal∘i₂≈h₂ ⟩∘⟨refl ⟩ + μ ∘ ¡ +₁ id ∘ λ⇐′ ≈⟨ μ∘¡ˡ ⟩ + id ∎ + + open 𝒟 using (U; monoidal; _⊗₁_; unitorˡ-commute-from) renaming (module unitorˡ to λ-) + open 𝒞 using (¡; ⊥; ¡-unique; pushout) renaming ([_,_] to [_,_]′; _+₁_ to infixr 10 _+₁_ ) + open Category U + open Equiv + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + open F using (F₁; ε) + open Shorthands monoidal + + same-deco : F₁ ≅ ∘ F₁ [ i₁ , i₂ ]′ ∘ φ (B , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈ s + same-deco = begin + F₁ ≅ ∘ F₁ [ i₁ , i₂ ]′ ∘ φ (B , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (B , M) ∘ (F₁ ¡ ∘ ε) ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₁ˡ ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (B , M) ∘ F₁ ¡ ⊗₁ id ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ sym F.identity ⟩∘⟨refl ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (B , M) ∘ F₁ ¡ ⊗₁ F₁ 𝒞.id ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (¡ , 𝒞.id)) ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ F₁ (¡ +₁ 𝒞.id) ∘ φ (⊥ , M) ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) ∘ φ (⊥ , M) ∘ ε ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₁₂ ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) ∘ φ (⊥ , M) ∘ ε ⊗₁ id ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ unitorˡ) F.unitaryˡ) ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) ∘ F₁ λ⇐′ ∘ λ⇒ ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) 𝒞.∘ λ⇐′) ∘ λ⇒ ∘ id ⊗₁ s ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ unitorˡ-commute-from ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) 𝒞.∘ λ⇐′) ∘ s ∘ λ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ coherence-inv₃ monoidal ⟨ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) 𝒞.∘ λ⇐′) ∘ s ∘ λ⇒ ∘ λ⇐ ≈⟨ refl⟩∘⟨ (sym-assoc ○ cancelʳ λ-.isoʳ) ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ ¡ +₁ 𝒞.id) 𝒞.∘ λ⇐′) ∘ s ≈⟨ elimˡ (F.F-resp-≈ ≅∘[]∘¡+id ○ F.identity) ⟩ + s ∎ + + R∘B : {A N B C : 𝒞.Obj} + → {f : 𝒞.U [ A , N ]} + → {g : 𝒞.U [ B , N ]} + → {h : 𝒞.U [ C , B ]} + → {s : 𝒟.U [ 𝒟.unit , F.₀ N ]} + → DecoratedCospans [ DecoratedCospans [ R.₁ h ∘ B₁ f g s ] ≈ B₁ f (𝒞.U [ g ∘ h ]) s ] + R∘B {A} {N} {B} {C} {f} {g} {h} {s} = record + { cospans-≈ = Embed.R∘B + ; same-deco = same-deco + } + where + + module _ where + open 𝒞 using (¡; ⊥; ¡-unique; pushout) renaming ([_,_] to [_,_]′; _+₁_ to infixr 10 _+₁_ ) + open 𝒞 using (U) + open Category U + open mc𝒞 using (unitorʳ; unitorˡ; unitorˡ-commute-to; +-monoidal) public + open unitorˡ using () renaming (to to λ⇐′) public + open unitorʳ using () renaming (to to ρ⇐′) public + open ⊗-Reasoning +-monoidal + open ⇒-Reasoning 𝒞.U + open Equiv + + open Pushout′ 𝒞.U using (pushout-f-id) + open PushoutProperties 𝒞.U using (up-to-iso) + open DiagramPushout 𝒞.U using (Pushout) + P P′ : Pushout g 𝒞.id + P = pushout g 𝒞.id + P′ = pushout-f-id + module P = Pushout P + module P′ = Pushout P′ + open Morphism 𝒞.U using (_≅_) + open _≅_ using (from) + open P using (i₁ ; i₂; universal∘i₁≈h₁) public + + open Codiagonal using (μ; μ-natural; μ∘+; μ∘¡ʳ) + + ≅ : P.Q ⇒ N + ≅ = up-to-iso P P′ .from + + ≅∘[]∘id+¡ : ((≅ ∘ [ i₁ , i₂ ]′) ∘ id +₁ ¡) ∘ ρ⇐′ ≈ id + ≅∘[]∘id+¡ = begin + ((≅ ∘ [ i₁ , i₂ ]′) ∘ id +₁ ¡) ∘ ρ⇐′ ≈⟨ assoc²αε ⟩ + ≅ ∘ [ i₁ , i₂ ]′ ∘ id +₁ ¡ ∘ ρ⇐′ ≈⟨ refl⟩∘⟨ pushˡ μ∘+ ⟩ + ≅ ∘ μ ∘ i₁ +₁ i₂ ∘ id +₁ ¡ ∘ ρ⇐′ ≈⟨ refl⟩∘⟨ pull-center merge₂ʳ ⟩ + ≅ ∘ μ ∘ i₁ +₁ (i₂ ∘ ¡) ∘ ρ⇐′ ≈⟨ extendʳ μ-natural ⟩ + μ ∘ ≅ +₁ ≅ ∘ i₁ +₁ (i₂ ∘ ¡) ∘ ρ⇐′ ≈⟨ pull-center (sym ⊗-distrib-over-∘) ⟩ + μ ∘ (≅ ∘ i₁) +₁ (≅ ∘ i₂ ∘ ¡) ∘ ρ⇐′ ≈⟨ refl⟩∘⟨ universal∘i₁≈h₁ ⟩⊗⟨ sym (¡-unique (≅ ∘ i₂ ∘ ¡)) ⟩∘⟨refl ⟩ + μ ∘ id +₁ ¡ ∘ ρ⇐′ ≈⟨ μ∘¡ʳ ⟩ + id ∎ + + open 𝒟 using (U; monoidal; _⊗₁_; unitorʳ-commute-from) renaming (module unitorˡ to λ-; module unitorʳ to ρ) + open 𝒞 using (¡; ⊥; ¡-unique; pushout) renaming ([_,_] to [_,_]′; _+₁_ to infixr 10 _+₁_ ) + open Category U + open Equiv + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + open F using (F₁; ε) + open Shorthands monoidal + + same-deco : F₁ ≅ ∘ F₁ [ i₁ , i₂ ]′ ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈ s + same-deco = begin + F₁ ≅ ∘ F₁ [ i₁ , i₂ ]′ ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (N , B) ∘ s ⊗₁ (F₁ ¡ ∘ ε) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ split₂ˡ ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (N , B) ∘ id ⊗₁ F₁ ¡ ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym F.identity ⟩⊗⟨refl ⟩∘⟨refl ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ φ (N , B) ∘ F₁ 𝒞.id ⊗₁ F₁ ¡ ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (φ-commute (𝒞.id , ¡)) ⟩ + F₁ (≅ 𝒞.∘ [ i₁ , i₂ ]′) ∘ F₁ (𝒞.id +₁ ¡) ∘ φ (N , ⊥) ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) ∘ φ (N , ⊥) ∘ s ⊗₁ ε ∘ ρ⇐ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ pushˡ serialize₂₁ ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) ∘ φ (N , ⊥) ∘ id ⊗₁ ε ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ unitorʳ) F.unitaryʳ) ⟩ + F₁ ((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) ∘ F₁ ρ⇐′ ∘ ρ⇒ ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ pullˡ (sym F.homomorphism) ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) 𝒞.∘ ρ⇐′) ∘ ρ⇒ ∘ s ⊗₁ id ∘ ρ⇐ ≈⟨ refl⟩∘⟨ extendʳ unitorʳ-commute-from ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) 𝒞.∘ ρ⇐′) ∘ s ∘ ρ⇒ ∘ ρ⇐ ≈⟨ refl⟩∘⟨ (sym-assoc ○ cancelʳ ρ.isoʳ) ⟩ + F₁ (((≅ 𝒞.∘ [ i₁ , i₂ ]′) 𝒞.∘ 𝒞.id +₁ ¡) 𝒞.∘ ρ⇐′) ∘ s ≈⟨ elimˡ (F.F-resp-≈ ≅∘[]∘id+¡ ○ F.identity) ⟩ + s ∎ + + open Morphism 𝒞.U using (_≅_) + open _≅_ + + ≅-L-R : (X≅Y : A ≅ B) → DecoratedCospans [ L.₁ (to X≅Y) ≈ R.₁ (from X≅Y) ] + ≅-L-R X≅Y = Decorate.F-resp-≈ (Embed.≅-L-R X≅Y) + where + module Decorate = Functor Decorate + + open 𝒞 using (_+₁_) + + L-resp-⊗ : DecoratedCospans [ L.₁ (f +₁ g) ≈ ⊗.₁ (L.₁ f , L.₁ g) ] + L-resp-⊗ = Decorate.F-resp-≈ Embed.L-resp-⊗ ○ Decorate-resp-⊗ + where + module Decorate = Functor Decorate + open DecoratedCospans.HomReasoning diff --git a/Functor/Instance/DecoratedCospan/Stack.agda b/Functor/Instance/DecoratedCospan/Stack.agda new file mode 100644 index 0000000..381ee06 --- /dev/null +++ b/Functor/Instance/DecoratedCospan/Stack.agda @@ -0,0 +1,430 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) + +open Lax using (SymmetricMonoidalFunctor) +open FinitelyCocompleteCategory + using () + renaming (symmetricMonoidalCategory to smc) + +module Functor.Instance.DecoratedCospan.Stack + {o o′ ℓ ℓ′ e e′} + (𝒞 : FinitelyCocompleteCategory o ℓ e) + {𝒟 : SymmetricMonoidalCategory o′ ℓ′ e′} + (F : SymmetricMonoidalFunctor (smc 𝒞) 𝒟) where + +import Categories.Diagram.Pushout as DiagramPushout +import Categories.Morphism as Morphism +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning + +import Functor.Instance.Cospan.Stack 𝒞 as Stack + +open import Categories.Category using (Category; _[_,_]; _[_≈_]; _[_∘_]) +open import Categories.Category.BinaryProducts using (BinaryProducts) +open import Categories.Category.Monoidal.Utilities using (module Shorthands) +open import Categories.Category.Monoidal.Properties using (coherence-inv₃) +open import Categories.Category.Monoidal.Braided.Properties using (braiding-coherence-inv) +open import Categories.Functor using (Functor) +open import Categories.Functor.Bifunctor using (Bifunctor) +open import Categories.Functor.Properties using ([_]-resp-≅) +open import Categories.Category.Cocartesian using (module CocartesianMonoidal; module CocartesianSymmetricMonoidal) +open import Categories.Object.Initial using (Initial) +open import Categories.Object.Duality using (Coproduct⇒coProduct) +open import Category.Instance.DecoratedCospans 𝒞 F using () renaming (DecoratedCospans to Cospans; _≈_ to _≈_′) + +import Category.Diagram.Cospan 𝒞 as Cospan + +open import Cospan.Decorated 𝒞 F using (DecoratedCospan) +open import Data.Product.Base using (_,_) + +module 𝒞 = FinitelyCocompleteCategory 𝒞 +module 𝒟 = SymmetricMonoidalCategory 𝒟 +module F = SymmetricMonoidalFunctor F +module Cospans = Category Cospans + +open 𝒞 using (Obj; _+_; cocartesian) + +module mc𝒞 = CocartesianMonoidal 𝒞.U cocartesian +module smc𝒞 = CocartesianSymmetricMonoidal 𝒞.U cocartesian + +open DiagramPushout 𝒞.U using (Pushout) + +private + variable + A A′ B B′ C C′ : Obj + +together : Cospans [ A , B ] → Cospans [ A′ , B′ ] → Cospans [ A + A′ , B + B′ ] +together A⇒B A⇒B′ = record + { cospan = A⇒B.cospan Cospan.⊗ A⇒B′.cospan + ; decoration = ⊗-homo.η (A⇒B.N , A⇒B′.N) ∘ A⇒B.decoration ⊗₁ A⇒B′.decoration ∘ unitorʳ.to + } + where + module A⇒B = DecoratedCospan A⇒B + module A⇒B′ = DecoratedCospan A⇒B′ + open 𝒟 using (_∘_; _⊗₁_; module unitorʳ) + open F using (module ⊗-homo) + +id⊗id≈id : Cospans [ together (Cospans.id {A}) (Cospans.id {B}) ≈ Cospans.id ] +id⊗id≈id {A} {B} = record + { cospans-≈ = Stack.id⊗id≈id + ; same-deco = F.identity ⟩∘⟨refl + ○ identityˡ + ○ refl⟩∘⟨ ⊗-distrib-over-∘ ⟩∘⟨refl + ○ extendʳ (extendʳ (⊗-homo.commute (! , !))) + ○ refl⟩∘⟨ pullʳ (pushˡ serialize₂₁ ○ refl⟩∘⟨ sym unitorʳ-commute-to) + ○ pushˡ (F-resp-≈ !+!≈! ○ homomorphism) + ○ refl⟩∘⟨ (refl⟩∘⟨ sym-assoc ○ pullˡ unitaryʳ ○ cancelˡ unitorʳ.isoʳ) + } + where + open 𝒞 using (_+₁_; ¡-unique) + open 𝒟 using (identityˡ; U; monoidal; module unitorʳ; unitorʳ-commute-to; assoc; sym-assoc) + open 𝒟.Equiv + open ⇒-Reasoning U + open ⇒-Reasoning 𝒞.U using () renaming (flip-iso to flip-iso′) + open ⊗-Reasoning monoidal + open F using (module ⊗-homo; F-resp-≈; homomorphism; unitaryʳ) + open 𝒞 using (initial) + open Initial initial using (!; !-unique₂) + open Morphism using (_≅_; module ≅) + open mc𝒞 using (A+⊥≅A) + module A+⊥≅A = _≅_ A+⊥≅A + !+!≈! : 𝒞.U [ (! {A} +₁ ! {B}) ≈ ! {A + B} 𝒞.∘ A+⊥≅A.from ] + !+!≈! = 𝒞.Equiv.sym (flip-iso′ (≅.sym 𝒞.U A+⊥≅A) (¡-unique ((! +₁ !) 𝒞.∘ A+⊥≅A.to))) + +homomorphism + : (A⇒B : Cospans [ A , B ]) + → (B⇒C : Cospans [ B , C ]) + → (A⇒B′ : Cospans [ A′ , B′ ]) + → (B⇒C′ : Cospans [ B′ , C′ ]) + → Cospans + [ together (Cospans [ B⇒C ∘ A⇒B ]) (Cospans [ B⇒C′ ∘ A⇒B′ ]) + ≈ Cospans [ together B⇒C B⇒C′ ∘ together A⇒B A⇒B′ ] + ] +homomorphism {A} {B} {C} {A′} {B′} {C′} f g f′ g′ = record + { cospans-≈ = cospans-≈ + ; same-deco = same-deco + } + where + + module _ where + open DecoratedCospan using (cospan) + cospans-≈ : _ Cospan.⊗ _ Cospan.≈ Cospan.compose (_ Cospan.⊗ _) (_ Cospan.⊗ _) + cospans-≈ = Stack.homomorphism (f .cospan) (g .cospan) (f′ .cospan) (g′ .cospan) + open Cospan._≈_ cospans-≈ using () renaming (≅N to Q+Q′≅Q″) public + + module DecorationNames where + open DecoratedCospan f using (N) renaming (decoration to s) public + open DecoratedCospan g using () renaming (decoration to t; N to M) public + open DecoratedCospan f′ using () renaming (decoration to s′; N to N′) public + open DecoratedCospan g′ using () renaming (decoration to t′; N to M′) public + + module PushoutNames where + open DecoratedCospan using (f₁; f₂) + open 𝒞 using (pushout) + open Pushout (pushout (f .f₂) (g .f₁)) using (i₁; i₂; Q) public + open Pushout (pushout (f′ .f₂) (g′ .f₁)) using () renaming (i₁ to i₁′; i₂ to i₂′; Q to Q′) public + open Pushout (pushout (together f f′ .f₂) (together g g′ .f₁)) + using (universal∘i₁≈h₁; universal∘i₂≈h₂) + renaming (i₁ to i₁″; i₂ to i₂″; Q to Q″) public + + module _ where + + open DecorationNames + open PushoutNames + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + + open 𝒞 using () renaming ([_,_] to [_,_]′) + + module _ where + + open 𝒞 + using (U; +-swap; inject₁; inject₂; +-η) + renaming (i₁ to ι₁; i₂ to ι₂; _+₁_ to infixr 10 _+₁_) + open Category U hiding (Obj) + open Equiv + open Shorthands mc𝒞.+-monoidal + open ⊗-Reasoning mc𝒞.+-monoidal + open ⇒-Reasoning U + open mc𝒞 using (assoc-commute-from; assoc-commute-to; module ⊗; associator) + open smc𝒞 using () renaming (module braiding to σ) + + module Codiagonal where + + open 𝒞 using (coproduct; +-unique; []-cong₂; []∘+₁; ∘-distribˡ-[]) + μ : {X : Obj} → X + X ⇒ X + μ = [ id , id ]′ + + μ∘+ : {X Y Z : Obj} {f : X ⇒ Z} {g : Y ⇒ Z} → [ f , g ]′ ≈ μ ∘ f +₁ g + μ∘+ = []-cong₂ (sym identityˡ) (sym identityˡ) ○ sym []∘+₁ + + μ∘σ : {X : Obj} → μ ∘ +-swap ≈ μ {X} + μ∘σ = sym (+-unique (pullʳ inject₁ ○ inject₂) (pullʳ inject₂ ○ inject₁) ) + + op-binaryProducts : BinaryProducts op + op-binaryProducts = record { product = Coproduct⇒coProduct U coproduct } + + module op-binaryProducts = BinaryProducts op-binaryProducts + open op-binaryProducts using () renaming (assocʳ∘⟨⟩ to []∘assocˡ) + + μ-assoc : {X : Obj} → μ {X} ∘ μ +₁ (id {X}) ≈ μ ∘ (id {X}) +₁ μ ∘ α⇒ + μ-assoc = begin + μ ∘ μ +₁ id ≈⟨ μ∘+ ⟨ + [ [ id , id ]′ , id ]′ ≈⟨ []∘assocˡ ⟨ + [ id , [ id , id ]′ ]′ ∘ α⇒ ≈⟨ pushˡ μ∘+ ⟩ + μ ∘ id +₁ μ ∘ α⇒ ∎ + + μ-natural : {X Y : Obj} {f : X ⇒ Y} → f ∘ μ ≈ μ ∘ f +₁ f + μ-natural = ∘-distribˡ-[] ○ []-cong₂ (identityʳ ○ sym identityˡ) (identityʳ ○ sym identityˡ) ○ sym []∘+₁ + + open Codiagonal + + ≅ : Q + Q′ ⇒ Q″ + ≅ = Q+Q′≅Q″.from + + ≅∘[]+[]≈μ∘μ+μ : ≅ ∘ [ i₁ , i₂ ]′ +₁ [ i₁′ , i₂′ ]′ ≈ (μ ∘ (μ +₁ μ)) ∘ ((i₁″ ∘ ι₁) +₁ (i₂″ ∘ ι₁)) +₁ ((i₁″ ∘ ι₂) +₁ (i₂″ ∘ ι₂)) + ≅∘[]+[]≈μ∘μ+μ = begin + ≅ ∘ [ i₁ , i₂ ]′ +₁ [ i₁′ , i₂′ ]′ ≈⟨ refl⟩∘⟨ μ∘+ ⟩⊗⟨ μ∘+ ⟩ + ≅ ∘ (μ ∘ i₁ +₁ i₂) +₁ (μ ∘ i₁′ +₁ i₂′) ≈⟨ refl⟩∘⟨ introˡ +-η ⟩ + ≅ ∘ [ ι₁ , ι₂ ]′ ∘ (μ ∘ i₁ +₁ i₂) +₁ (μ ∘ i₁′ +₁ i₂′) ≈⟨ push-center μ∘+ ⟩ + ≅ ∘ μ ∘ (ι₁ +₁ ι₂) ∘ (μ ∘ i₁ +₁ i₂) +₁ (μ ∘ i₁′ +₁ i₂′) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ sym ⊗-distrib-over-∘ ⟩ + ≅ ∘ μ ∘ (ι₁ ∘ μ ∘ i₁ +₁ i₂) +₁ (ι₂ ∘ μ ∘ i₁′ +₁ i₂′) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (extendʳ μ-natural) ⟩⊗⟨ (extendʳ μ-natural) ⟩ + ≅ ∘ μ ∘ (μ ∘ ι₁ +₁ ι₁ ∘ i₁ +₁ i₂) +₁ (μ ∘ ι₂ +₁ ι₂ ∘ i₁′ +₁ i₂′) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ (refl⟩∘⟨ sym ⊗-distrib-over-∘) ⟩⊗⟨ (refl⟩∘⟨ sym ⊗-distrib-over-∘) ⟩ + ≅ ∘ μ ∘ (μ ∘ (ι₁ ∘ i₁) +₁ (ι₁ ∘ i₂)) +₁ (μ ∘ (ι₂ ∘ i₁′) +₁ (ι₂ ∘ i₂′)) ≈⟨ extendʳ μ-natural ⟩ + μ ∘ ≅ +₁ ≅ ∘ (μ ∘ _) +₁ (μ ∘ _) ≈⟨ refl⟩∘⟨ sym ⊗-distrib-over-∘ ⟩ + μ ∘ (≅ ∘ μ ∘ _) +₁ (≅ ∘ μ ∘ _) ≈⟨ refl⟩∘⟨ extendʳ μ-natural ⟩⊗⟨ extendʳ μ-natural ⟩ + μ ∘ (μ ∘ ≅ +₁ ≅ ∘ _) +₁ (μ ∘ ≅ +₁ ≅ ∘ _) ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ sym ⊗-distrib-over-∘) ⟩⊗⟨ (refl⟩∘⟨ sym ⊗-distrib-over-∘) ⟩ + μ ∘ (μ ∘ (≅ ∘ ι₁ ∘ i₁) +₁ (≅ ∘ ι₁ ∘ i₂)) +₁ (μ ∘ (≅ ∘ ι₂ ∘ i₁′) +₁ (≅ ∘ ι₂ ∘ i₂′)) ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ eq₁ ⟩⊗⟨ eq₂ ) ⟩⊗⟨ (refl⟩∘⟨ eq₃ ⟩⊗⟨ eq₄ ) ⟩ + μ ∘ (μ ∘ (i₁″ ∘ ι₁) +₁ (i₂″ ∘ ι₁)) +₁ (μ ∘ (i₁″ ∘ ι₂) +₁ (i₂″ ∘ ι₂)) ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟩ + μ ∘ (μ +₁ μ) ∘ ((i₁″ ∘ ι₁) +₁ (i₂″ ∘ ι₁)) +₁ ((i₁″ ∘ ι₂) +₁ (i₂″ ∘ ι₂)) ≈⟨ sym-assoc ⟩ + (μ ∘ (μ +₁ μ)) ∘ ((i₁″ ∘ ι₁) +₁ (i₂″ ∘ ι₁)) +₁ ((i₁″ ∘ ι₂) +₁ (i₂″ ∘ ι₂)) ∎ + where + eq₁ : ≅ ∘ ι₁ ∘ i₁ ≈ i₁″ ∘ ι₁ + eq₁ = refl⟩∘⟨ sym inject₁ ○ pullˡ (sym (switch-tofromˡ Q+Q′≅Q″ universal∘i₁≈h₁)) + eq₂ : ≅ ∘ ι₁ ∘ i₂ ≈ i₂″ ∘ ι₁ + eq₂ = refl⟩∘⟨ sym inject₁ ○ pullˡ (sym (switch-tofromˡ Q+Q′≅Q″ universal∘i₂≈h₂)) + eq₃ : ≅ ∘ ι₂ ∘ i₁′ ≈ i₁″ ∘ ι₂ + eq₃ = refl⟩∘⟨ sym inject₂ ○ pullˡ (sym (switch-tofromˡ Q+Q′≅Q″ universal∘i₁≈h₁)) + eq₄ : ≅ ∘ ι₂ ∘ i₂′ ≈ i₂″ ∘ ι₂ + eq₄ = refl⟩∘⟨ sym inject₂ ○ pullˡ (sym (switch-tofromˡ Q+Q′≅Q″ universal∘i₂≈h₂)) + + swap-inner : {W X Y Z : Obj} → (W + X) + (Y + Z) ⇒ (W + Y) + (X + Z) + swap-inner = α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ + + swap-inner-natural + : {W X Y Z W′ X′ Y′ Z′ : Obj} + {w : W ⇒ W′} {x : X ⇒ X′} {y : Y ⇒ Y′} {z : Z ⇒ Z′} + → (w +₁ x) +₁ (y +₁ z) ∘ swap-inner + ≈ swap-inner ∘ (w +₁ y) +₁ (x +₁ z) + swap-inner-natural {w = w} {x = x} {y = y} {z = z} = begin + (w +₁ x) +₁ (y +₁ z) ∘ α⇐ ∘ _ ≈⟨ extendʳ assoc-commute-to ⟨ + α⇐ ∘ w +₁ (x +₁ _) ∘ id +₁ _ ∘ α⇒ ≈⟨ pull-center merge₂ʳ ⟩ + α⇐ ∘ w +₁ (x +₁ _ ∘ α⇒ ∘ _) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ extendʳ assoc-commute-from ⟩∘⟨refl ⟨ + α⇐ ∘ w +₁ (α⇒ ∘ (x +₁ y) +₁ z ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ pushˡ split₁ʳ) ⟩∘⟨refl ⟨ + α⇐ ∘ w +₁ (α⇒ ∘ (x +₁ y ∘ +-swap) +₁ z ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ σ.⇒.sym-commute _ ⟩⊗⟨refl ⟩∘⟨refl) ⟩∘⟨refl ⟩ + α⇐ ∘ w +₁ (α⇒ ∘ (+-swap ∘ y +₁ x) +₁ z ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ push-center split₁ˡ ⟩∘⟨refl ⟩ + α⇐ ∘ w +₁ (α⇒ ∘ +-swap +₁ id ∘ (y +₁ x) +₁ z ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ refl⟩∘⟨ assoc-commute-to) ⟩∘⟨refl ⟨ + α⇐ ∘ w +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐ ∘ y +₁ (x +₁ z)) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ assoc²εβ ⟩∘⟨refl ⟩ + α⇐ ∘ w +₁ ((α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ y +₁ (x +₁ z)) ∘ α⇒ ≈⟨ refl⟩∘⟨ pushˡ split₂ˡ ⟩ + α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ w +₁ (y +₁ (x +₁ z)) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc-commute-from ⟨ + α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ∘ (w +₁ y) +₁ (x +₁ z) ≈⟨ assoc²εβ ⟩ + swap-inner ∘ (w +₁ y) +₁ (x +₁ z) ∎ + + μ∘μ+μ∘swap-inner : {X : Obj} → μ {X} ∘ μ +₁ μ ∘ swap-inner ≈ μ ∘ μ +₁ μ {X} + μ∘μ+μ∘swap-inner = begin + μ ∘ μ +₁ μ ∘ α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ push-center serialize₁₂ ⟩ + μ ∘ μ +₁ id ∘ id +₁ μ ∘ α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⊗.identity ⟩⊗⟨refl ⟩∘⟨refl ⟨ + μ ∘ μ +₁ id ∘ (id +₁ id) +₁ μ ∘ α⇐ ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ extendʳ assoc-commute-to ⟨ + μ ∘ μ +₁ id ∘ α⇐ ∘ id +₁ (id +₁ μ) ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ pullˡ μ-assoc ⟩ + (μ ∘ id +₁ μ ∘ α⇒) ∘ α⇐ ∘ id +₁ (id +₁ μ) ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ extendʳ (pullʳ (cancelʳ associator.isoʳ)) ⟩ + μ ∘ id +₁ μ ∘ id +₁ (id +₁ μ) ∘ id +₁ (α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ pull-center merge₂ˡ ⟩ + μ ∘ id +₁ μ ∘ id +₁ (id +₁ μ ∘ α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ pull-center merge₂ʳ ⟩ + μ ∘ id +₁ (μ ∘ id +₁ μ ∘ α⇒ ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ pull-center refl ⟩∘⟨refl ⟩ + μ ∘ id +₁ (μ ∘ (id +₁ μ ∘ α⇒) ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ extendʳ μ-assoc ⟩∘⟨refl ⟨ + μ ∘ id +₁ (μ ∘ μ +₁ id ∘ +-swap +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ pull-center (sym split₁ˡ) ⟩∘⟨refl ⟩ + μ ∘ id +₁ (μ ∘ (μ ∘ +-swap) +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ μ∘σ ⟩⊗⟨refl ⟩∘⟨refl) ⟩∘⟨refl ⟩ + μ ∘ id +₁ (μ ∘ μ +₁ id ∘ α⇐) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (sym-assoc ○ flip-iso associator (μ-assoc ○ sym-assoc)) ⟩∘⟨refl ⟩ + μ ∘ id +₁ (μ ∘ id +₁ μ) ∘ α⇒ ≈⟨ push-center split₂ʳ ⟩ + μ ∘ id +₁ μ ∘ id +₁ (id +₁ μ) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc-commute-from ⟨ + μ ∘ id +₁ μ ∘ α⇒ ∘ (id +₁ id) +₁ μ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩∘⟨ ⊗.identity ⟩⊗⟨refl ⟩ + μ ∘ id +₁ μ ∘ α⇒ ∘ id +₁ μ ≈⟨ refl⟩∘⟨ sym-assoc ⟩ + μ ∘ (id +₁ μ ∘ α⇒) ∘ id +₁ μ ≈⟨ extendʳ μ-assoc ⟨ + μ ∘ μ +₁ id ∘ id +₁ μ ≈⟨ refl⟩∘⟨ serialize₁₂ ⟨ + μ ∘ μ +₁ μ ∎ + + ≅∘[]+[]∘σ₄ : (Q+Q′≅Q″.from ∘ [ i₁ , i₂ ]′ +₁ [ i₁′ , i₂′ ]′) ∘ swap-inner ≈ [ i₁″ , i₂″ ]′ + ≅∘[]+[]∘σ₄ = begin + (≅ ∘ [ i₁ , i₂ ]′ +₁ [ i₁′ , i₂′ ]′) ∘ _ ≈⟨ pushˡ ≅∘[]+[]≈μ∘μ+μ ⟩ + (μ ∘ (μ +₁ μ)) ∘ ((i₁″ ∘ ι₁) +₁ (i₂″ ∘ ι₁)) +₁ ((i₁″ ∘ ι₂) +₁ (i₂″ ∘ ι₂)) ∘ (α⇐ ∘ _) ≈⟨ refl⟩∘⟨ swap-inner-natural ⟩ + (μ ∘ (μ +₁ μ)) ∘ swap-inner ∘ _ ≈⟨ pullˡ assoc ⟩ + (μ ∘ (μ +₁ μ) ∘ swap-inner) ∘ _ ≈⟨ pushˡ μ∘μ+μ∘swap-inner ⟩ + μ ∘ (μ +₁ μ) ∘ ((i₁″ ∘ ι₁) +₁ (i₁″ ∘ ι₂)) +₁ ((i₂″ ∘ ι₁) +₁ (i₂″ ∘ ι₂)) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟩⊗⟨ ⊗-distrib-over-∘ ⟩ + μ ∘ (μ +₁ μ) ∘ (i₁″ +₁ i₁″ ∘ ι₁ +₁ ι₂) +₁ (i₂″ +₁ i₂″ ∘ ι₁ +₁ ι₂) ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟨ + μ ∘ (μ ∘ i₁″ +₁ i₁″ ∘ ι₁ +₁ ι₂) +₁ (μ ∘ i₂″ +₁ i₂″ ∘ ι₁ +₁ ι₂) ≈⟨ refl⟩∘⟨ extendʳ μ-natural ⟩⊗⟨ extendʳ μ-natural ⟨ + μ ∘ (i₁″ ∘ μ ∘ ι₁ +₁ ι₂) +₁ (i₂″ ∘ μ ∘ ι₁ +₁ ι₂) ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ μ∘+) ⟩⊗⟨ (refl⟩∘⟨ μ∘+) ⟨ + μ ∘ (i₁″ ∘ [ ι₁ , ι₂ ]′) +₁ (i₂″ ∘ [ ι₁ , ι₂ ]′) ≈⟨ refl⟩∘⟨ elimʳ +-η ⟩⊗⟨ elimʳ +-η ⟩ + μ ∘ i₁″ +₁ i₂″ ≈⟨ μ∘+ ⟨ + [ i₁″ , i₂″ ]′ ∎ + + module _ where + + open 𝒟 using (U; _⊗₁_; module ⊗; module unitorˡ; module unitorʳ; monoidal; assoc-commute-from; assoc-commute-to) + open Category U + open ⇒-Reasoning U + open Equiv + open ⊗-Reasoning monoidal + open smc𝒞 using () renaming (associator to α) + open 𝒟 using () renaming (associator to α′) + open Morphism._≅_ + + swap-unit : 𝒟.braiding.⇐.η (𝒟.unit , 𝒟.unit) ≈ 𝒟.id + swap-unit = cancel-toʳ 𝒟.unitorˡ + ( braiding-coherence-inv 𝒟.braided + ○ sym (coherence-inv₃ monoidal) + ○ sym 𝒟.identityˡ) + + φ-swap-inner : φ (N + M , N′ + M′) ∘ (φ (N , M) ∘ s ⊗₁ t) ⊗₁ (φ (N′ , M′) ∘ s′ ⊗₁ t′) + ≈ F.F₁ swap-inner ∘ φ (N + N′ , M + M′) ∘ (φ (N , N′) ∘ s ⊗₁ s′) ⊗₁ (φ (M , M′) ∘ t ⊗₁ t′) + φ-swap-inner = refl⟩∘⟨ ⊗-distrib-over-∘ + ○ extendʳ + ( insertˡ ([ F.F ]-resp-≅ α .isoˡ) ⟩∘⟨ serialize₁₂ + ○ center (assoc ○ F.associativity) + ○ refl⟩∘⟨ + (extendˡ + ( refl⟩∘⟨ sym ⊗.identity ⟩⊗⟨refl + ○ extendˡ assoc-commute-from + ○ ( merge₂ʳ + ○ sym F.identity ⟩⊗⟨ + ( switch-fromtoʳ α′ (assoc ○ (sym F.associativity)) + ○ ( refl⟩∘⟨ + ( refl⟩∘⟨ + ( switch-fromtoʳ 𝒟.braiding.FX≅GX (sym F.braiding-compat) ⟩⊗⟨refl + ○ assoc ⟩⊗⟨refl + ○ split₁ʳ + ○ refl⟩⊗⟨ sym F.identity ⟩∘⟨refl) + ○ extendʳ (φ-commute (_ , 𝒞.id)) + ○ refl⟩∘⟨ + ( refl⟩∘⟨ split₁ˡ + ○ extendʳ (switch-fromtoˡ ([ F.F ]-resp-≅ α) F.associativity)) + ○ pullˡ (sym F.homomorphism)) + ○ pullˡ (sym F.homomorphism)) ⟩∘⟨refl + ○ assoc) + ○ split₂ʳ) ⟩∘⟨refl) + ○ ( extendʳ (φ-commute (𝒞.id , _)) + ○ refl⟩∘⟨ + ( refl⟩∘⟨ split₂ʳ + ○ extendʳ + ( refl⟩∘⟨ (refl⟩⊗⟨ assoc ○ split₂ʳ) + ○ extendʳ (switch-fromtoʳ α′ (assoc ○ (sym F.associativity))) + ○ refl⟩∘⟨ + ( refl⟩∘⟨ (refl⟩⊗⟨ assoc ○ split₂ʳ) + ○ extendʳ assoc-commute-to + ○ ⊗.identity ⟩⊗⟨refl ⟩∘⟨refl) + ○ extendʳ (assoc ○ refl⟩∘⟨ (assoc ○ refl⟩∘⟨ sym serialize₁₂)))) + ○ pullˡ (sym F.homomorphism) + ○ refl⟩∘⟨ (assoc ○ refl⟩∘⟨ pullʳ merge₂ʳ) ) ⟩∘⟨refl ) + ○ center⁻¹ (sym F.homomorphism) assoc) + ○ refl⟩∘⟨ + ( pullʳ + ( extendˡ assoc-commute-from + ○ ( pullʳ + ( merge₂ˡ + ○ refl⟩⊗⟨ + ( extendˡ assoc-commute-to + ○ ( pullʳ (sym split₁ˡ ○ (𝒟.braiding.⇐.commute (s′ , t) ○ elimʳ swap-unit) ⟩⊗⟨refl) + ○ assoc-commute-from ) ⟩∘⟨refl + ○ cancelʳ 𝒟.associator.isoʳ)) + ○ assoc-commute-to) ⟩∘⟨refl + ○ cancelʳ 𝒟.associator.isoˡ) + ○ pullʳ (sym ⊗-distrib-over-∘)) + + open Shorthands monoidal + + same-deco + : (F.₁ Q+Q′≅Q″.from ∘ φ (Q , Q′) ∘ (F.₁ [ i₁ , i₂ ]′ ∘ φ (N , M) ∘ s ⊗₁ t ∘ ρ⇐) ⊗₁ (F.₁ [ i₁′ , i₂′ ]′ ∘ φ (N′ , M′) ∘ s′ ⊗₁ t′ ∘ ρ⇐) ∘ ρ⇐) + ≈ (F.₁ [ i₁″ , i₂″ ]′ ∘ φ (N + N′ , M + M′) ∘ (φ (N , N′) ∘ s ⊗₁ s′ ∘ ρ⇐) ⊗₁ (φ (M , M′) ∘ t ⊗₁ t′ ∘ ρ⇐) ∘ ρ⇐) + same-deco = + refl⟩∘⟨ + ( refl⟩∘⟨ pushˡ ⊗-distrib-over-∘ + ○ extendʳ (φ-commute ([ i₁ , i₂ ]′ , [ i₁′ , i₂′ ]′)) + ○ refl⟩∘⟨ refl⟩∘⟨ sym-assoc ⟩⊗⟨ sym-assoc ⟩∘⟨refl + ○ refl⟩∘⟨ refl⟩∘⟨ pushˡ ⊗-distrib-over-∘ + ○ refl⟩∘⟨ sym-assoc) + ○ pullˡ (sym F.homomorphism) + ○ extendʳ (pushʳ φ-swap-inner) + ○ (sym F.homomorphism ○ F.F-resp-≈ ≅∘[]+[]∘σ₄) ⟩∘⟨refl + ○ refl⟩∘⟨ + ( assoc + ○ refl⟩∘⟨ pullˡ (sym ⊗-distrib-over-∘) + ○ refl⟩∘⟨ assoc ⟩⊗⟨ assoc ⟩∘⟨refl) + +⊗-resp-≈ + : {A A′ B B′ : Obj} + {f f′ : Cospans [ A , B ]} + {g g′ : Cospans [ A′ , B′ ]} + → Cospans [ f ≈ f′ ] + → Cospans [ g ≈ g′ ] + → Cospans [ together f g ≈ together f′ g′ ] +⊗-resp-≈ {_} {_} {_} {_} {f} {f′} {g} {g′} ≈f ≈g = record + { cospans-≈ = Stack.⊗-resp-≈ (≈f .cospans-≈) (≈g .cospans-≈) + ; same-deco = same-deco + } + where + + open _≈_′ using (cospans-≈) + + module SameNames where + open _≈_′ ≈f using () renaming (same-deco to ≅∘s≈t) public + open _≈_′ ≈g using () renaming (same-deco to ≅∘s≈t′) public + open Cospan._≈_ (≈f .cospans-≈) using (module ≅N) public + open Cospan._≈_ (≈g .cospans-≈) using () renaming (module ≅N to ≅N′) public + + open SameNames + + module DecorationNames where + open DecoratedCospan f using (N) renaming (decoration to s) public + open DecoratedCospan f′ using () renaming (decoration to t; N to M) public + open DecoratedCospan g using () renaming (decoration to s′; N to N′) public + open DecoratedCospan g′ using () renaming (decoration to t′; N to M′) public + + open DecorationNames + + module _ where + open 𝒞 using (_⇒_) + ≅ : N ⇒ M + ≅ = ≅N.from + ≅′ : N′ ⇒ M′ + ≅′ = ≅N′.from + + open 𝒞 using (_+₁_) + + module _ where + + open 𝒟 using (U; monoidal; _⊗₁_) + open Category U + open Equiv + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open F.⊗-homo using () renaming (η to φ; commute to φ-commute) + open F using (F₁) + open Shorthands monoidal + + same-deco : F₁ (≅ +₁ ≅′) ∘ φ (N , N′) ∘ s ⊗₁ s′ ∘ ρ⇐ ≈ φ (M , M′) ∘ t ⊗₁ t′ ∘ ρ⇐ + same-deco = begin + F₁ (≅ +₁ ≅′) ∘ φ (N , N′) ∘ s ⊗₁ s′ ∘ ρ⇐ ≈⟨ extendʳ (φ-commute (_ , _)) ⟨ + φ (M , M′) ∘ F₁ ≅ ⊗₁ F₁ ≅′ ∘ s ⊗₁ s′ ∘ ρ⇐ ≈⟨ pull-center (sym ⊗-distrib-over-∘) ⟩ + φ (M , M′) ∘ (F₁ ≅ ∘ s) ⊗₁ (F₁ ≅′ ∘ s′) ∘ ρ⇐ ≈⟨ refl⟩∘⟨ ≅∘s≈t ⟩⊗⟨ ≅∘s≈t′ ⟩∘⟨refl ⟩ + φ (M , M′) ∘ t ⊗₁ t′ ∘ ρ⇐ ∎ + +⊗ : 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 + } + +module ⊗ = Functor ⊗ diff --git a/Functor/Instance/Endo/List.agda b/Functor/Instance/Endo/List.agda new file mode 100644 index 0000000..67e3d0b --- /dev/null +++ b/Functor/Instance/Endo/List.agda @@ -0,0 +1,15 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module Functor.Instance.Endo.List {ℓ : Level} where + +import Functor.Instance.List {ℓ} {ℓ} as List + +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Endofunctor) + +-- List is only an endofunctor when the carrier sets and +-- equivalence relations live at the same level +List : Endofunctor (Setoids ℓ ℓ) +List = List.List diff --git a/Functor/Instance/List.agda b/Functor/Instance/List.agda new file mode 100644 index 0000000..a280218 --- /dev/null +++ b/Functor/Instance/List.agda @@ -0,0 +1,67 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module Functor.Instance.List {c ℓ : Level} where + +import Data.List.Properties as ListProps +import Data.List.Relation.Binary.Pointwise as PW + +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Setoid using (∣_∣; _⇒ₛ_) +open import Function.Base using (_∘_; id) +open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_) +open import Relation.Binary using (Setoid) + +open Functor +open Setoid using (reflexive) +open Func + +open import Data.Opaque.List as L hiding (List) + +private + variable + A B C : Setoid c ℓ + +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (_∙_) + +opaque + + unfolding L.List + + map-id + : (xs : ∣ Listₛ A ∣) + → (open Setoid (Listₛ A)) + → mapₛ (Id _) ⟨$⟩ xs ≈ xs + map-id {A} = reflexive (Listₛ A) ∘ ListProps.map-id + + List-homo + : (f : A ⟶ₛ B) + (g : B ⟶ₛ C) + → (xs : ∣ Listₛ A ∣) + → (open Setoid (Listₛ C)) + → mapₛ (g ∙ f) ⟨$⟩ xs ≈ mapₛ g ⟨$⟩ (mapₛ f ⟨$⟩ xs) + List-homo {C = C} f g = reflexive (Listₛ C) ∘ ListProps.map-∘ + + List-resp-≈ + : (f g : A ⟶ₛ B) + → (let open Setoid (A ⇒ₛ B) in f ≈ g) + → (let open Setoid (Listₛ A ⇒ₛ Listₛ B) in mapₛ f ≈ mapₛ g) + List-resp-≈ f g f≈g = PW.map⁺ (to f) (to g) (PW.refl f≈g) + +-- the List functor takes a carrier A to lists of A +-- and the equivalence on A to pointwise equivalence on lists of A + +-- List on morphisms is the familiar map operation +-- which applies the same function to every element of a list + +List : Functor (Setoids c ℓ) (Setoids c (c ⊔ ℓ)) +List .F₀ = Listₛ +List .F₁ = mapₛ +List .identity {_} {xs} = map-id xs +List .homomorphism {f = f} {g} {xs} = List-homo f g xs +List .F-resp-≈ {f = f} {g} f≈g = List-resp-≈ f g f≈g + +module List = Functor List diff --git a/Functor/Instance/Monoidalize.agda b/Functor/Instance/Monoidalize.agda new file mode 100644 index 0000000..6423109 --- /dev/null +++ b/Functor/Instance/Monoidalize.agda @@ -0,0 +1,43 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +open import Categories.Category using (Category) +open import Categories.Category.Monoidal using (MonoidalCategory) +open import Categories.Category.Cocartesian using (Cocartesian) + +module Functor.Instance.Monoidalize + {o o′ ℓ ℓ′ e e ′ : Level} + {C : Category o ℓ e} + (cocartesian : Cocartesian C) + (D : MonoidalCategory o ℓ e) + where + +open import Categories.Category.Cocartesian using (module CocartesianMonoidal) + +open import Categories.Functor using (Functor) +open import Categories.Functor.Monoidal using (MonoidalFunctor) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Construction.Functors using (Functors) +open import Categories.Category.Construction.MonoidalFunctors using (module Lax) +open import Functor.Monoidal.Construction.MonoidValued cocartesian {D} using () renaming (F,⊗,ε to MonoidalFunctorOf) +open import NaturalTransformation.Monoidal.Construction.MonoidValued cocartesian {D} using () renaming (β,⊗,ε to MonoidalNaturalTransformationOf) + +C-MC : MonoidalCategory o ℓ e +C-MC = record { monoidal = +-monoidal } + where + open CocartesianMonoidal C cocartesian + +module C = MonoidalCategory C-MC +module D = MonoidalCategory D + +open Lax using (MonoidalFunctors) +open Functor + +Monoidalize : Functor (Functors C.U (Monoids D.monoidal)) (MonoidalFunctors C-MC D) +Monoidalize .F₀ = MonoidalFunctorOf +Monoidalize .F₁ α = MonoidalNaturalTransformationOf α +Monoidalize .identity = D.Equiv.refl +Monoidalize .homomorphism = D.Equiv.refl +Monoidalize .F-resp-≈ x = x + +module Monoidalize = Functor Monoidalize diff --git a/Functor/Instance/Multiset.agda b/Functor/Instance/Multiset.agda new file mode 100644 index 0000000..b961c7b --- /dev/null +++ b/Functor/Instance/Multiset.agda @@ -0,0 +1,72 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module Functor.Instance.Multiset {c ℓ : Level} where + +import Data.Opaque.List as L +import Data.List.Properties as ListProps +import Data.List.Relation.Binary.Pointwise as PW + +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.List.Relation.Binary.Permutation.Setoid using (↭-setoid; ↭-reflexive-≋) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (map⁺) +open import Data.Opaque.Multiset using (Multisetₛ; mapₛ) +open import Data.Setoid using (∣_∣; _⇒ₛ_) +open import Function.Base using (_∘_; id) +open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (_∙_) +open import Relation.Binary using (Setoid) + +open Functor +open Setoid using (reflexive) +open Func + +private + variable + A B C : Setoid c ℓ + +-- the Multiset functor takes a carrier A to lists of A +-- and the equivalence on A to permutation equivalence on lists of A + +-- Multiset on morphisms applies the same function to every element of a multiset + +opaque + unfolding mapₛ + + map-id + : (xs : ∣ Multisetₛ A ∣) + → (open Setoid (Multisetₛ A)) + → mapₛ (Id A) ⟨$⟩ xs ≈ xs + map-id {A} = reflexive (Multisetₛ A) ∘ ListProps.map-id + +opaque + unfolding mapₛ + + Multiset-homo + : (f : A ⟶ₛ B) + (g : B ⟶ₛ C) + → (xs : ∣ Multisetₛ A ∣) + → (open Setoid (Multisetₛ C)) + → mapₛ (g ∙ f) ⟨$⟩ xs ≈ mapₛ g ⟨$⟩ (mapₛ f ⟨$⟩ xs) + Multiset-homo {C = C} f g = reflexive (Multisetₛ C) ∘ ListProps.map-∘ + +opaque + unfolding mapₛ + + Multiset-resp-≈ + : (f g : A ⟶ₛ B) + → (let open Setoid (A ⇒ₛ B) in f ≈ g) + → (let open Setoid (Multisetₛ A ⇒ₛ Multisetₛ B) in mapₛ f ≈ mapₛ g) + Multiset-resp-≈ {A} {B} f g f≈g = ↭-reflexive-≋ B (PW.map⁺ (to f) (to g) (PW.refl f≈g)) + +Multiset : Functor (Setoids c ℓ) (Setoids c (c ⊔ ℓ)) +Multiset .F₀ = Multisetₛ +Multiset .F₁ = mapₛ +Multiset .identity {A} {xs} = map-id {A} xs +Multiset .homomorphism {f = f} {g} {xs} = Multiset-homo f g xs +Multiset .F-resp-≈ {A} {B} {f} {g} f≈g = Multiset-resp-≈ f g f≈g + +module Multiset = Functor Multiset diff --git a/Functor/Instance/Nat/Circ.agda b/Functor/Instance/Nat/Circ.agda new file mode 100644 index 0000000..88a6ec6 --- /dev/null +++ b/Functor/Instance/Nat/Circ.agda @@ -0,0 +1,56 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) + +module Functor.Instance.Nat.Circ {ℓ : Level} where + +import Data.List.Relation.Binary.Permutation.Setoid as ↭ + +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Functor using (Functor; _∘F_) +open import Categories.Morphism.Notation using (_[_≅_]) +open import Category.Instance.Setoids.SymmetricMonoidal {ℓ} {ℓ} using (Setoids-×) +open import Data.Circuit using (mk≈) +open import Data.Circuit {ℓ} using (Circuitₛ; mkCircuitₛ; edgesₛ) +open import Data.Circuit.Gate using (Gates) +open import Data.Nat using (ℕ) +open import Data.Opaque.Multiset using (Multisetₛ) +open import Data.Product using (proj₁; proj₂; Σ-syntax) +open import Functor.Free.Instance.CMonoid using (Free) +open import Functor.Instance.Nat.Edge {ℓ} Gates using (Edge) +open import Functor.Properties using (define-by-pw-iso) + +open import Category.Construction.CMonoids Setoids-×.symmetric using (CMonoids) +open import Category.Construction.CMonoids.Properties Setoids-×.symmetric using (transport-by-iso) +open import Object.Monoid.Commutative Setoids-×.symmetric using (CommutativeMonoid) + +Edges : Functor Nat CMonoids +Edges = Free ∘F Edge + +module Edges = Functor Edges +module Edge = Functor Edge + +opaque + unfolding Multisetₛ + Edges≅Circₛ : (n : ℕ) → Setoids-×.U [ Multisetₛ (Edge.₀ n) ≅ Circuitₛ n ] + Edges≅Circₛ n = record + { from = mkCircuitₛ + ; to = edgesₛ + ; iso = record + { isoˡ = ↭.↭-refl (Edge.₀ n) + ; isoʳ = mk≈ (↭.↭-refl (Edge.₀ n)) + } + } + +private + Edges≅ : (n : ℕ) → Σ[ M ∈ CommutativeMonoid ] CMonoids [ Edges.₀ n ≅ M ] + Edges≅ n = transport-by-iso (Edges.₀ n) (Edges≅Circₛ n) + +Circuitₘ : ℕ → CommutativeMonoid +Circuitₘ n = proj₁ (Edges≅ n) + +Edges≅Circₘ : (n : ℕ) → CMonoids [ Edges.₀ n ≅ Circuitₘ n ] +Edges≅Circₘ n = proj₂ (Edges≅ n) + +Circ : Functor Nat CMonoids +Circ = proj₁ (define-by-pw-iso Edges Circuitₘ Edges≅Circₘ) diff --git a/Functor/Instance/Nat/Edge.agda b/Functor/Instance/Nat/Edge.agda new file mode 100644 index 0000000..c69a1db --- /dev/null +++ b/Functor/Instance/Nat/Edge.agda @@ -0,0 +1,60 @@ +{-# OPTIONS --without-K --safe #-} + +open import Data.Hypergraph.Label using (HypergraphLabel) +open import Level using (Level; 0ℓ) + +module Functor.Instance.Nat.Edge {ℓ : Level} (HL : HypergraphLabel) where + +import Data.Vec as Vec +import Data.Vec.Properties as VecProps + +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Fin using (Fin) +open import Data.Fin.Properties using (cast-is-id) +open import Data.Hypergraph.Edge {ℓ} HL as Edge using (Edgeₛ; map; mapₛ; _≈_) +open import Data.Nat using (ℕ) +open import Data.Vec.Relation.Binary.Equality.Cast using (≈-reflexive) +open import Function using (id; _∘_; Func; _⟶ₛ_) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) + +module HL = HypergraphLabel HL + +open Edge.Edge +open Edge._≈_ +open Func +open Functor + +map-id : {v : ℕ} {e : Edge.Edge v} → map id e ≈ e +map-id .≡arity = ≡.refl +map-id .≡label = HL.≈-reflexive ≡.refl +map-id {_} {e} .≡ports = ≡.cong (ports e) ∘ ≡.sym ∘ cast-is-id ≡.refl + +map-∘ + : {n m o : ℕ} + (f : Fin n → Fin m) + (g : Fin m → Fin o) + {e : Edge.Edge n} + → map (g ∘ f) e ≈ map g (map f e) +map-∘ f g .≡arity = ≡.refl +map-∘ f g .≡label = HL.≈-reflexive ≡.refl +map-∘ f g {e} .≡ports = ≡.cong (g ∘ f ∘ ports e) ∘ ≡.sym ∘ cast-is-id ≡.refl + +map-resp-≗ + : {n m : ℕ} + {f g : Fin n → Fin m} + → f ≗ g + → {e : Edge.Edge n} + → map f e ≈ map g e +map-resp-≗ f≗g .≡arity = ≡.refl +map-resp-≗ f≗g .≡label = HL.≈-reflexive ≡.refl +map-resp-≗ {g = g} f≗g {e} .≡ports i = ≡.trans (f≗g (ports e i)) (≡.cong (g ∘ ports e) (≡.sym (cast-is-id ≡.refl i))) + +Edge : Functor Nat (Setoids ℓ ℓ) +Edge .F₀ = Edgeₛ +Edge .F₁ = mapₛ +Edge .identity = map-id +Edge .homomorphism {f = f} {g} = map-∘ f g +Edge .F-resp-≈ = map-resp-≗ diff --git a/Functor/Instance/Nat/Preimage.agda b/Functor/Instance/Nat/Preimage.agda new file mode 100644 index 0000000..7da00f4 --- /dev/null +++ b/Functor/Instance/Nat/Preimage.agda @@ -0,0 +1,65 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Preimage where + +open import Categories.Category.Instance.Nat using (Natop) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Fin.Base using (Fin) +open import Data.Bool.Base using (Bool) +open import Data.Nat.Base using (ℕ) +open import Data.Subset.Functional using (Subset) +open import Data.Vec.Functional.Relation.Binary.Equality.Setoid using (≋-setoid) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) + +open Functor +open Func + +_≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ +_≈_ {X} {Y} = Setoid._≈_ (setoid X Y) +infixr 4 _≈_ + +private + variable A B C : ℕ + +-- action on objects (Subset n) +Subsetₛ : ℕ → Setoid 0ℓ 0ℓ +Subsetₛ = ≋-setoid (≡.setoid Bool) + +-- action of Preimage on morphisms (contravariant) +Preimage₁ : (Fin A → Fin B) → Subsetₛ B ⟶ₛ Subsetₛ A +to (Preimage₁ f) i = i ∘ f +cong (Preimage₁ f) x≗y = x≗y ∘ f + +-- Preimage respects identity +Preimage-identity : Preimage₁ id ≈ Id (Subsetₛ A) +Preimage-identity {A} = Setoid.refl (Subsetₛ A) + +-- Preimage flips composition +Preimage-homomorphism + : {A B C : ℕ} + (f : Fin A → Fin B) + (g : Fin B → Fin C) + → Preimage₁ (g ∘ f) ≈ Preimage₁ f ∙ Preimage₁ g +Preimage-homomorphism {A} _ _ = Setoid.refl (Subsetₛ A) + +-- Preimage respects equality +Preimage-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Preimage₁ f ≈ Preimage₁ g +Preimage-resp-≈ f≗g {v} = ≡.cong v ∘ f≗g + +-- the Preimage functor +Preimage : Functor Natop (Setoids 0ℓ 0ℓ) +F₀ Preimage = Subsetₛ +F₁ Preimage = Preimage₁ +identity Preimage = Preimage-identity +homomorphism Preimage {f = f} {g} {v} = Preimage-homomorphism g f {v} +F-resp-≈ Preimage = Preimage-resp-≈ diff --git a/Functor/Instance/Nat/Pull.agda b/Functor/Instance/Nat/Pull.agda new file mode 100644 index 0000000..b1764d9 --- /dev/null +++ b/Functor/Instance/Nat/Pull.agda @@ -0,0 +1,81 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Pull where + +open import Categories.Category.Instance.Nat using (Natop) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Fin.Base using (Fin) +open import Data.Nat.Base using (ℕ) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) +open import Data.Circuit.Value using (Monoid) +open import Data.System.Values Monoid using (Values) +open import Data.Unit using (⊤; tt) + +open Functor +open Func + +-- Pull takes a natural number n to the setoid Values n + +private + + variable A B C : ℕ + + _≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ + _≈_ {X} {Y} = Setoid._≈_ (setoid X Y) + + infixr 4 _≈_ + + opaque + + unfolding Values + + -- action of Pull on morphisms (contravariant) + Pull₁ : (Fin A → Fin B) → Values B ⟶ₛ Values A + to (Pull₁ f) i = i ∘ f + cong (Pull₁ f) x≗y = x≗y ∘ f + + -- Pull respects identity + Pull-identity : Pull₁ id ≈ Id (Values A) + Pull-identity {A} = Setoid.refl (Values A) + + opaque + + unfolding Pull₁ + + -- Pull flips composition + Pull-homomorphism + : (f : Fin A → Fin B) + (g : Fin B → Fin C) + → Pull₁ (g ∘ f) ≈ Pull₁ f ∙ Pull₁ g + Pull-homomorphism {A} _ _ = Setoid.refl (Values A) + + -- Pull respects equality + Pull-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Pull₁ f ≈ Pull₁ g + Pull-resp-≈ f≗g {v} = ≡.cong v ∘ f≗g + +opaque + + unfolding Pull₁ + + Pull-defs : ⊤ + Pull-defs = tt + +-- the Pull functor +Pull : Functor Natop (Setoids 0ℓ 0ℓ) +F₀ Pull = Values +F₁ Pull = Pull₁ +identity Pull = Pull-identity +homomorphism Pull {f = f} {g} = Pull-homomorphism g f +F-resp-≈ Pull = Pull-resp-≈ + +module Pull = Functor Pull diff --git a/Functor/Instance/Nat/Push.agda b/Functor/Instance/Nat/Push.agda new file mode 100644 index 0000000..8126006 --- /dev/null +++ b/Functor/Instance/Nat/Push.agda @@ -0,0 +1,79 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Push where + +open import Categories.Functor using (Functor) +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Data.Circuit.Merge using (merge; merge-cong₁; merge-cong₂; merge-⁅⁆; merge-preimage) +open import Data.Fin.Base using (Fin) +open import Data.Fin.Preimage using (preimage; preimage-cong₁) +open import Data.Nat.Base using (ℕ) +open import Data.Subset.Functional using (⁅_⁆) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) +open import Data.Circuit.Value using (Monoid) +open import Data.System.Values Monoid using (Values) +open import Data.Unit using (⊤; tt) + +open Func +open Functor + +-- Push sends a natural number n to Values n + +private + + variable A B C : ℕ + + _≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ + _≈_ {X} {Y} = Setoid._≈_ (setoid X Y) + infixr 4 _≈_ + + opaque + + unfolding Values + + -- action of Push on morphisms (covariant) + Push₁ : (Fin A → Fin B) → Values A ⟶ₛ Values B + to (Push₁ f) v = merge v ∘ preimage f ∘ ⁅_⁆ + cong (Push₁ f) x≗y = merge-cong₁ x≗y ∘ preimage f ∘ ⁅_⁆ + + -- Push respects identity + Push-identity : Push₁ id ≈ Id (Values A) + Push-identity {_} {v} = merge-⁅⁆ v + + -- Push respects composition + Push-homomorphism + : {f : Fin A → Fin B} + {g : Fin B → Fin C} + → Push₁ (g ∘ f) ≈ Push₁ g ∙ Push₁ f + Push-homomorphism {f = f} {g} {v} = merge-preimage f v ∘ preimage g ∘ ⁅_⁆ + + -- Push respects equality + Push-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Push₁ f ≈ Push₁ g + Push-resp-≈ f≗g {v} = merge-cong₂ v ∘ preimage-cong₁ f≗g ∘ ⁅_⁆ + +opaque + + unfolding Push₁ + + Push-defs : ⊤ + Push-defs = tt + +-- the Push functor +Push : Functor Nat (Setoids 0ℓ 0ℓ) +F₀ Push = Values +F₁ Push = Push₁ +identity Push = Push-identity +homomorphism Push = Push-homomorphism +F-resp-≈ Push = Push-resp-≈ + +module Push = Functor Push diff --git a/Functor/Instance/Nat/System.agda b/Functor/Instance/Nat/System.agda new file mode 100644 index 0000000..05e1e7b --- /dev/null +++ b/Functor/Instance/Nat/System.agda @@ -0,0 +1,110 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.System where + + +open import Level using (suc; 0ℓ) + +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor.Core using (Functor) +open import Data.Circuit.Value using (Monoid) +open import Data.Fin.Base using (Fin) +open import Data.Nat.Base using (ℕ) +open import Data.Product.Base using (_,_; _×_) +open import Data.System {suc 0ℓ} using (System; _≤_; Systemₛ) +open import Data.System.Values Monoid using (module ≋) +open import Data.Unit using (⊤; tt) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (_∙_) +open import Functor.Instance.Nat.Pull using (Pull) +open import Functor.Instance.Nat.Push using (Push) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) + +open Func +open Functor +open _≤_ + +private + + variable A B C : ℕ + + opaque + + map : (Fin A → Fin B) → System A → System B + map f X = let open System X in record + { S = S + ; fₛ = fₛ ∙ Pull.₁ f + ; fₒ = Push.₁ f ∙ fₒ + } + + ≤-cong : (f : Fin A → Fin B) {X Y : System A} → Y ≤ X → map f Y ≤ map f X + ⇒S (≤-cong f x≤y) = ⇒S x≤y + ≗-fₛ (≤-cong f x≤y) = ≗-fₛ x≤y ∘ to (Pull.₁ f) + ≗-fₒ (≤-cong f x≤y) = cong (Push.₁ f) ∘ ≗-fₒ x≤y + + System₁ : (Fin A → Fin B) → Systemₛ A ⟶ₛ Systemₛ B + to (System₁ f) = map f + cong (System₁ f) (x≤y , y≤x) = ≤-cong f x≤y , ≤-cong f y≤x + + opaque + + unfolding System₁ + + id-x≤x : {X : System A} → System₁ id ⟨$⟩ X ≤ X + ⇒S (id-x≤x) = Id _ + ≗-fₛ (id-x≤x {_} {x}) i s = cong (System.fₛ x) Pull.identity + ≗-fₒ (id-x≤x {A} {x}) s = Push.identity + + x≤id-x : {x : System A} → x ≤ System₁ id ⟨$⟩ x + ⇒S x≤id-x = Id _ + ≗-fₛ (x≤id-x {A} {x}) i s = cong (System.fₛ x) (≋.sym Pull.identity) + ≗-fₒ (x≤id-x {A} {x}) s = ≋.sym Push.identity + + System-homomorphism + : {f : Fin A → Fin B} + {g : Fin B → Fin C} + {X : System A} + → System₁ (g ∘ f) ⟨$⟩ X ≤ System₁ g ⟨$⟩ (System₁ f ⟨$⟩ X) + × System₁ g ⟨$⟩ (System₁ f ⟨$⟩ X) ≤ System₁ (g ∘ f) ⟨$⟩ X + System-homomorphism {f = f} {g} {X} = left , right + where + open System X + left : map (g ∘ f) X ≤ map g (map f X) + left .⇒S = Id S + left .≗-fₛ i s = cong fₛ Pull.homomorphism + left .≗-fₒ s = Push.homomorphism + right : map g (map f X) ≤ map (g ∘ f) X + right .⇒S = Id S + right .≗-fₛ i s = cong fₛ (≋.sym Pull.homomorphism) + right .≗-fₒ s = ≋.sym Push.homomorphism + + System-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → {X : System A} + → System₁ f ⟨$⟩ X ≤ System₁ g ⟨$⟩ X + × System₁ g ⟨$⟩ X ≤ System₁ f ⟨$⟩ X + System-resp-≈ {A} {B} {f = f} {g} f≗g {X} = both f≗g , both (≡.sym ∘ f≗g) + where + open System X + both : {f g : Fin A → Fin B} → f ≗ g → map f X ≤ map g X + both f≗g .⇒S = Id S + both f≗g .≗-fₛ i s = cong fₛ (Pull.F-resp-≈ f≗g {i}) + both {f} {g} f≗g .≗-fₒ s = Push.F-resp-≈ f≗g + +opaque + unfolding System₁ + Sys-defs : ⊤ + Sys-defs = tt + +Sys : Functor Nat (Setoids (suc 0ℓ) (suc 0ℓ)) +Sys .F₀ = Systemₛ +Sys .F₁ = System₁ +Sys .identity = id-x≤x , x≤id-x +Sys .homomorphism {x = X} = System-homomorphism {X = X} +Sys .F-resp-≈ = System-resp-≈ + +module Sys = Functor Sys diff --git a/Functor/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda b/Functor/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda new file mode 100644 index 0000000..80b7b2f --- /dev/null +++ b/Functor/Instance/Underlying/SymmetricMonoidal/FinitelyCocomplete.agda @@ -0,0 +1,229 @@ +{-# OPTIONS --without-K --safe #-} +{-# OPTIONS --no-require-unique-meta-solutions #-} + +open import Level using (Level) +module Functor.Instance.Underlying.SymmetricMonoidal.FinitelyCocomplete {o ℓ e : Level} where + +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Object.Coproduct as Coproduct +import Categories.Object.Initial as Initial + +open import Categories.Functor using (Functor; _∘F_) +open import Categories.Functor.Properties using ([_]-resp-square; [_]-resp-∘) +open import Categories.Functor.Monoidal using (IsMonoidalFunctor) +open import Categories.Functor.Monoidal.Braided using (module Lax) +open import Categories.Functor.Monoidal.Properties using (idF-SymmetricMonoidal; ∘-SymmetricMonoidal) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.Category using (Category; _[_,_]; _[_≈_]; _[_∘_]) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory; BraidedMonoidalCategory; MonoidalCategory) +open import Categories.Category.Product using (_⁂_) +open import Categories.Morphism using (_≅_) +open import Categories.Morphism.Notation using (_[_≅_]) +open import Categories.NaturalTransformation.Core using (NaturalTransformation; ntHelper) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper) renaming (refl to ≃-refl) +open import Categories.NaturalTransformation.NaturalIsomorphism.Monoidal.Symmetric using (module Lax) +open import Data.Product.Base using (_,_) + +open import Category.Instance.FinitelyCocompletes {o} {ℓ} {e} using (FinitelyCocompletes) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Category.Instance.SymMonCat {o} {ℓ} {e} using (SymMonCat) +open import Functor.Exact using (RightExactFunctor; idREF; ∘-RightExactFunctor) + +open FinitelyCocompleteCategory using () renaming (symmetricMonoidalCategory to smc) +open SymmetricMonoidalCategory using (unit) renaming (braidedMonoidalCategory to bmc) +open BraidedMonoidalCategory using () renaming (monoidalCategory to mc) + +private + variable + A B C : FinitelyCocompleteCategory o ℓ e + +F₀ : FinitelyCocompleteCategory o ℓ e → SymmetricMonoidalCategory o ℓ e +F₀ C = smc C +{-# INJECTIVE_FOR_INFERENCE F₀ #-} + +F₁ : RightExactFunctor A B → Lax.SymmetricMonoidalFunctor (F₀ A) (F₀ B) +F₁ {A} {B} F = record + { F = F.F + ; isBraidedMonoidal = record + { isMonoidal = record + { ε = ε-iso.from + ; ⊗-homo = ⊗-homo + ; associativity = assoc + ; unitaryˡ = unitaryˡ + ; unitaryʳ = unitaryʳ + } + ; braiding-compat = braiding-compat + } + } + where + module F = RightExactFunctor F + module A = SymmetricMonoidalCategory (F₀ A) + module B = SymmetricMonoidalCategory (F₀ B) + module A′ = FinitelyCocompleteCategory A + module B′ = FinitelyCocompleteCategory B + ε-iso : B.U [ B.unit ≅ F.₀ A.unit ] + ε-iso = Initial.up-to-iso B.U B′.initial (record { ⊥ = F.₀ A′.⊥ ; ⊥-is-initial = F.F-resp-⊥ A′.⊥-is-initial }) + module ε-iso = _≅_ ε-iso + +-iso : ∀ {X Y} → B.U [ F.₀ X B′.+ F.₀ Y ≅ F.₀ (X A′.+ Y) ] + +-iso = Coproduct.up-to-iso B.U B′.coproduct (Coproduct.IsCoproduct⇒Coproduct B.U (F.F-resp-+ (Coproduct.Coproduct⇒IsCoproduct A.U A′.coproduct))) + module +-iso {X Y} = _≅_ (+-iso {X} {Y}) + module B-proofs where + open ⇒-Reasoning B.U + open B.HomReasoning + open B.Equiv + open B using (_∘_; _≈_) + open B′ using (_+₁_; []-congˡ; []-congʳ; []-cong₂) + open A′ using (_+_; i₁; i₂) + ⊗-homo : NaturalTransformation (B.⊗ ∘F (F.F ⁂ F.F)) (F.F ∘F A.⊗) + ⊗-homo = ntHelper record + { η = λ { (X , Y) → +-iso.from {X} {Y} } + ; commute = λ { {X , Y} {X′ , Y′} (f , g) → + B′.coproduct.∘-distribˡ-[] + ○ B′.coproduct.[]-cong₂ + (pullˡ B′.coproduct.inject₁ ○ [ F.F ]-resp-square (A.Equiv.sym A′.coproduct.inject₁)) + (pullˡ B′.coproduct.inject₂ ○ [ F.F ]-resp-square (A.Equiv.sym A′.coproduct.inject₂)) + ○ sym B′.coproduct.∘-distribˡ-[] } + } + assoc + : {X Y Z : A.Obj} + → F.₁ A′.+-assocˡ + ∘ +-iso.from {X + Y} {Z} + ∘ (+-iso.from {X} {Y} +₁ B.id {F.₀ Z}) + ≈ +-iso.from {X} {Y + Z} + ∘ (B.id {F.₀ X} +₁ +-iso.from {Y} {Z}) + ∘ B′.+-assocˡ + assoc {X} {Y} {Z} = begin + F.₁ A′.+-assocˡ ∘ +-iso.from ∘ (+-iso.from +₁ B.id) ≈⟨ refl⟩∘⟨ B′.[]∘+₁ ⟩ + F.₁ A′.+-assocˡ ∘ B′.[ F.₁ i₁ ∘ +-iso.from , F.₁ i₂ ∘ B.id ] ≈⟨ refl⟩∘⟨ []-congʳ B′.coproduct.∘-distribˡ-[] ⟩ + F.₁ A′.+-assocˡ ∘ B′.[ B′.[ F.₁ i₁ ∘ F.₁ i₁ , F.₁ i₁ ∘ F.₁ i₂ ] , F.₁ i₂ ∘ B.id ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟩ + B′.[ F.₁ A′.+-assocˡ ∘ B′.[ F.₁ i₁ ∘ F.₁ i₁ , F.₁ i₁ ∘ F.₁ i₂ ] , _ ] ≈⟨ []-congʳ B′.coproduct.∘-distribˡ-[] ⟩ + B′.[ B′.[ F.₁ A′.+-assocˡ ∘ F.₁ i₁ ∘ F.₁ i₁ , F.₁ A′.+-assocˡ ∘ _ ] , _ ] ≈⟨ []-congʳ ([]-congʳ (pullˡ ([ F.F ]-resp-∘ A′.coproduct.inject₁))) ⟩ + B′.[ B′.[ F.₁ A′.[ i₁ , i₂ A′.∘ i₁ ] ∘ F.₁ i₁ , F.₁ A′.+-assocˡ ∘ _ ] , _ ] ≈⟨ []-congʳ ([]-congʳ ([ F.F ]-resp-∘ A′.coproduct.inject₁)) ⟩ + B′.[ B′.[ F.₁ i₁ , F.₁ A′.+-assocˡ ∘ F.₁ i₁ ∘ F.₁ i₂ ] , _ ] ≈⟨ []-congʳ ([]-congˡ (pullˡ ([ F.F ]-resp-∘ A′.coproduct.inject₁))) ⟩ + B′.[ B′.[ F.₁ i₁ , F.₁ A′.[ i₁ , i₂ A′.∘ i₁ ] ∘ F.₁ i₂ ] , _ ] ≈⟨ []-congʳ ([]-congˡ ([ F.F ]-resp-∘ A′.coproduct.inject₂)) ⟩ + B′.[ B′.[ F.₁ i₁ , F.₁ (i₂ A′.∘ i₁) ] , F.₁ A′.+-assocˡ ∘ F.₁ i₂ ∘ B.id ] ≈⟨ []-congˡ (pullˡ ([ F.F ]-resp-∘ A′.coproduct.inject₂)) ⟩ + B′.[ B′.[ F.₁ i₁ , F.₁ (i₂ A′.∘ i₁) ] , F.₁ (i₂ A′.∘ i₂) ∘ B.id ] ≈⟨ []-cong₂ ([]-congˡ F.homomorphism) (B.identityʳ ○ F.homomorphism) ⟩ + B′.[ B′.[ F.₁ i₁ , F.₁ i₂ B′.∘ F.₁ i₁ ] , F.₁ i₂ ∘ F.₁ i₂ ] ≈⟨ []-congʳ ([]-congˡ B′.coproduct.inject₁) ⟨ + B′.[ B′.[ F.₁ i₁ , B′.[ F.₁ i₂ B′.∘ F.₁ i₁ , _ ] ∘ B′.i₁ ] , _ ] ≈⟨ []-congʳ ([]-cong₂ (sym B′.coproduct.inject₁) (pushˡ (sym B′.coproduct.inject₂))) ⟩ + B′.[ B′.[ B′.[ F.₁ i₁ , _ ] ∘ B′.i₁ , B′.[ F.₁ i₁ , _ ] ∘ B′.i₂ ∘ B′.i₁ ] , _ ] ≈⟨ []-congʳ B′.coproduct.∘-distribˡ-[] ⟨ + B′.[ B′.[ F.₁ i₁ , _ ] ∘ B′.[ B′.i₁ , B′.i₂ ∘ B′.i₁ ] , F.₁ i₂ ∘ F.₁ i₂ ] ≈⟨ []-congˡ B′.coproduct.inject₂ ⟨ + B′.[ B′.[ F.₁ i₁ , _ ] ∘ B′.[ _ , _ ] , B′.[ _ , F.₁ i₂ ∘ F.₁ i₂ ] ∘ B′.i₂ ] ≈⟨ []-congˡ (pushˡ (sym B′.coproduct.inject₂)) ⟩ + B′.[ B′.[ F.₁ i₁ , _ ] ∘ B′.[ _ , _ ] , B′.[ F.₁ i₁ , _ ] ∘ B′.i₂ ∘ B′.i₂ ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟨ + B′.[ F.₁ i₁ , B′.[ F.₁ i₂ ∘ F.₁ i₁ , F.₁ i₂ ∘ F.₁ i₂ ] ] ∘ B′.+-assocˡ ≈⟨ []-cong₂ B.identityʳ (B′.coproduct.∘-distribˡ-[]) ⟩∘⟨refl ⟨ + B′.[ F.₁ i₁ B′.∘ B′.id , F.₁ i₂ ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] ] ∘ B′.+-assocˡ ≈⟨ pushˡ (sym B′.[]∘+₁) ⟩ + +-iso.from ∘ (B.id +₁ +-iso.from) ∘ B′.+-assocˡ ∎ + unitaryˡ + : {X : A.Obj} + → F.₁ A′.[ A′.initial.! , A.id {X} ] + ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] + ∘ B′.[ B′.i₁ ∘ B′.initial.! , B′.i₂ ∘ B.id ] + ≈ B′.[ B′.initial.! , B.id ] + unitaryˡ {X} = begin + F.₁ A′.[ A′.initial.! , A.id ] ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.[ _ , B′.i₂ ∘ B.id ] ≈⟨ refl⟩∘⟨ B′.coproduct.∘-distribˡ-[] ⟩ + _ ∘ B′.[ _ ∘ B′.i₁ ∘ B′.initial.! , B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.i₂ ∘ B.id ] ≈⟨ refl⟩∘⟨ []-cong₂ (sym (B′.¡-unique _)) (pullˡ B′.coproduct.inject₂) ⟩ + F.₁ A′.[ A′.initial.! , A.id ] ∘ B′.[ B′.initial.! , F.₁ i₂ ∘ B.id ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟩ + B′.[ _ ∘ B′.initial.! , F.₁ A′.[ A′.initial.! , A.id ] ∘ F.₁ i₂ ∘ B.id ] ≈⟨ []-cong₂ (sym (B′.¡-unique _)) (pullˡ ([ F.F ]-resp-∘ A′.coproduct.inject₂)) ⟩ + B′.[ B′.initial.! , F.₁ A.id ∘ B.id ] ≈⟨ []-congˡ (elimˡ F.identity) ⟩ + B′.[ B′.initial.! , B.id ] ∎ + unitaryʳ + : {X : A.Obj} + → F.₁ A′.[ A′.id {X} , A′.initial.! ] + ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] + ∘ B′.[ B′.i₁ ∘ B.id , B′.i₂ ∘ B′.initial.! ] + ≈ B′.[ B.id , B′.initial.! ] + unitaryʳ {X} = begin + F.₁ A′.[ A.id , A′.initial.! ] ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.[ B′.i₁ ∘ B.id , _ ] ≈⟨ refl⟩∘⟨ B′.coproduct.∘-distribˡ-[] ⟩ + _ ∘ B′.[ B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.i₁ ∘ B.id , _ ∘ B′.i₂ ∘ B′.initial.! ] ≈⟨ refl⟩∘⟨ []-cong₂ (pullˡ B′.coproduct.inject₁) (sym (B′.¡-unique _)) ⟩ + F.₁ A′.[ A.id , A′.initial.! ] ∘ B′.[ F.₁ i₁ ∘ B.id , B′.initial.! ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟩ + B′.[ F.₁ A′.[ A.id , A′.initial.! ] ∘ F.₁ i₁ ∘ B.id , _ ∘ B′.initial.! ] ≈⟨ []-cong₂ (pullˡ ([ F.F ]-resp-∘ A′.coproduct.inject₁)) (sym (B′.¡-unique _)) ⟩ + B′.[ F.₁ A.id ∘ B.id , B′.initial.! ] ≈⟨ []-congʳ (elimˡ F.identity) ⟩ + B′.[ B.id , B′.initial.! ] ∎ + braiding-compat + : {X Y : A.Obj} + → F.₁ A′.[ i₂ {X} {Y} , i₁ ] ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] + ≈ B′.[ F.F₁ i₁ , F.F₁ i₂ ] ∘ B′.[ B′.i₂ , B′.i₁ ] + braiding-compat = begin + F.₁ A′.[ i₂ , i₁ ] ∘ B′.[ F.₁ i₁ , F.₁ i₂ ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟩ + B′.[ F.₁ A′.[ i₂ , i₁ ] ∘ F.₁ i₁ , F.₁ A′.[ i₂ , i₁ ] ∘ F.₁ i₂ ] ≈⟨ []-cong₂ ([ F.F ]-resp-∘ A′.coproduct.inject₁) ([ F.F ]-resp-∘ A′.coproduct.inject₂) ⟩ + B′.[ F.₁ i₂ , F.₁ i₁ ] ≈⟨ []-cong₂ B′.coproduct.inject₂ B′.coproduct.inject₁ ⟨ + B′.[ B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.i₂ , B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.i₁ ] ≈⟨ B′.coproduct.∘-distribˡ-[] ⟨ + B′.[ F.₁ i₁ , F.₁ i₂ ] ∘ B′.[ B′.i₂ , B′.i₁ ] ∎ + open B-proofs + +identity : Lax.SymmetricMonoidalNaturalIsomorphism (F₁ (Functor.Exact.idREF {o} {ℓ} {e} {A})) (idF-SymmetricMonoidal (F₀ A)) +identity {A} = record + { U = ≃-refl + ; F⇒G-isMonoidal = record + { ε-compat = ¡-unique₂ (id ∘ ¡) id + ; ⊗-homo-compat = refl⟩∘⟨ sym ([]-cong₂ identityʳ identityʳ) + } + } + where + open FinitelyCocompleteCategory A + open HomReasoning + open Equiv + +homomorphism + : {F : RightExactFunctor A B} + {G : RightExactFunctor B C} + → Lax.SymmetricMonoidalNaturalIsomorphism + (F₁ {A} {C} (∘-RightExactFunctor G F)) + (∘-SymmetricMonoidal (F₁ {B} {C} G) (F₁ {A} {B} F)) +homomorphism {A} {B} {C} {F} {G} = record + { U = ≃-refl + ; F⇒G-isMonoidal = record + { ε-compat = ¡-unique₂ (id ∘ ¡) (G.₁ B.¡ ∘ ¡) + ; ⊗-homo-compat = + identityˡ + ○ sym + ([]-cong₂ + ([ G.F ]-resp-∘ B.coproducts.inject₁) + ([ G.F ]-resp-∘ B.coproducts.inject₂)) + ○ sym ∘-distribˡ-[] + ○ pushʳ (introʳ C.⊗.identity) + } + } + where + module A = FinitelyCocompleteCategory A + module B = FinitelyCocompleteCategory B + open FinitelyCocompleteCategory C + module C = SymmetricMonoidalCategory (F₀ C) + open HomReasoning + open Equiv + open ⇒-Reasoning U + module F = RightExactFunctor F + module G = RightExactFunctor G + +module _ {F G : RightExactFunctor A B} where + + module F = RightExactFunctor F + module G = RightExactFunctor G + + F-resp-≈ + : NaturalIsomorphism F.F G.F + → Lax.SymmetricMonoidalNaturalIsomorphism (F₁ {A} {B} F) (F₁ {A} {B} G) + F-resp-≈ F≃G = record + { U = F≃G + ; F⇒G-isMonoidal = record + { ε-compat = sym (¡-unique (⇒.η A.⊥ ∘ ¡)) + ; ⊗-homo-compat = + ∘-distribˡ-[] + ○ []-cong₂ (⇒.commute A.i₁) (⇒.commute A.i₂) + ○ sym []∘+₁ + } + } + where + module A = FinitelyCocompleteCategory A + open NaturalIsomorphism F≃G + open FinitelyCocompleteCategory B + open HomReasoning + open Equiv + +Underlying : Functor FinitelyCocompletes SymMonCat +Underlying = record + { F₀ = F₀ + ; F₁ = F₁ + ; identity = λ { {X} → identity {X} } + ; homomorphism = λ { {X} {Y} {Z} {F} {G} → homomorphism {X} {Y} {Z} {F} {G} } + ; F-resp-≈ = λ { {X} {Y} {F} {G} → F-resp-≈ {X} {Y} {F} {G} } + } diff --git a/Functor/Monoidal/Braided/Strong/Properties.agda b/Functor/Monoidal/Braided/Strong/Properties.agda new file mode 100644 index 0000000..66dc4c0 --- /dev/null +++ b/Functor/Monoidal/Braided/Strong/Properties.agda @@ -0,0 +1,59 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +open import Categories.Category.Monoidal using (BraidedMonoidalCategory) +open import Categories.Functor.Monoidal.Braided using (module Strong) +open Strong using (BraidedMonoidalFunctor) + +module Functor.Monoidal.Braided.Strong.Properties + {o o′ ℓ ℓ′ e e′ : Level} + {C : BraidedMonoidalCategory o ℓ e} + {D : BraidedMonoidalCategory o′ ℓ′ e′} + (F,φ,ε : BraidedMonoidalFunctor C D) where + +import Categories.Category.Construction.Core as Core +import Categories.Category.Monoidal.Utilities as ⊗-Utilities +import Functor.Monoidal.Strong.Properties as MonoidalProp + +open import Categories.Functor.Properties using ([_]-resp-≅) + +private + + module C = BraidedMonoidalCategory C + module D = BraidedMonoidalCategory D + +open D +open Core.Shorthands U using (_∘ᵢ_; idᵢ; _≈ᵢ_; ⌞_⌟; to-≈; _≅_; module HomReasoningᵢ) +open ⊗-Utilities monoidal using (_⊗ᵢ_) +open BraidedMonoidalFunctor F,φ,ε +open MonoidalProp monoidalFunctor public + +private + + variable + A B : Obj + X Y : C.Obj + + σ : A ⊗₀ B ≅ B ⊗₀ A + σ = braiding.FX≅GX + + σ⇐ : B ⊗₀ A ⇒ A ⊗₀ B + σ⇐ = braiding.⇐.η _ + + Fσ : F₀ (X C.⊗₀ Y) ≅ F₀ (Y C.⊗₀ X) + Fσ = [ F ]-resp-≅ C.braiding.FX≅GX + + Fσ⇐ : F₀ (Y C.⊗₀ X) ⇒ F₀ (X C.⊗₀ Y) + Fσ⇐ = F₁ (C.braiding.⇐.η _) + + φ : F₀ X ⊗₀ F₀ Y ≅ F₀ (X C.⊗₀ Y) + φ = ⊗-homo.FX≅GX + +open HomReasoning +open Shorthands using (φ⇐) + +braiding-compatᵢ : Fσ {X} {Y} ∘ᵢ φ ≈ᵢ φ ∘ᵢ σ +braiding-compatᵢ = ⌞ braiding-compat ⌟ + +braiding-compat-inv : φ⇐ ∘ Fσ⇐ {X} {Y} ≈ σ⇐ ∘ φ⇐ +braiding-compat-inv = to-≈ braiding-compatᵢ diff --git a/Functor/Monoidal/Construction/MonoidValued.agda b/Functor/Monoidal/Construction/MonoidValued.agda new file mode 100644 index 0000000..937714d --- /dev/null +++ b/Functor/Monoidal/Construction/MonoidValued.agda @@ -0,0 +1,214 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category using (Category) +open import Categories.Category.Cocartesian using (Cocartesian) +open import Categories.Category.Cocartesian.Bundle using (CocartesianCategory) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory) +open import Categories.Functor using (Functor) renaming (_∘F_ to _∙_) +open import Level using (Level; _⊔_) + +-- A functor from a cocartesian category 𝒞 to Monoids[S] +-- can be turned into a monoidal functor from 𝒞 to S + +module Functor.Monoidal.Construction.MonoidValued + {o o′ ℓ ℓ′ e e′ : Level} + {𝒞 : Category o ℓ e} + (𝒞-+ : Cocartesian 𝒞) + {S : MonoidalCategory o′ ℓ′ e′} + (let module S = MonoidalCategory S) + (M : Functor 𝒞 (Monoids S.monoidal)) + where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Category.Monoidal.Utilities as ⊗-Util +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Object.Monoid as MonoidObject + +open import Categories.Category using (module Definitions) +open import Categories.Category.Cocartesian using (module CocartesianMonoidal) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor.Monoidal using (MonoidalFunctor; IsMonoidalFunctor) +open import Categories.Functor.Properties using ([_]-resp-square; [_]-resp-∘) +open import Categories.Morphism using (_≅_) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Data.Product using (_,_) +open import Functor.Forgetful.Instance.Monoid S.monoidal using (Forget) + +private + + G : Functor 𝒞 S.U + G = Forget ∙ M + + module 𝒞 = CocartesianCategory (record { cocartesian = 𝒞-+ }) + module 𝒞-M = CocartesianMonoidal 𝒞 𝒞-+ + + 𝒞-MC : MonoidalCategory o ℓ e + 𝒞-MC = record { monoidal = 𝒞-M.+-monoidal } + + module +-assoc {n} {m} {o} = _≅_ (𝒞.+-assoc {n} {m} {o}) + module +-λ {n} = _≅_ (𝒞-M.⊥+A≅A {n}) + module +-ρ {n} = _≅_ (𝒞-M.A+⊥≅A {n}) + + module G = Functor G + module M = Functor M + + open MonoidObject S.monoidal using (Monoid; Monoid⇒) + open Monoid renaming (assoc to μ-assoc; identityˡ to μ-identityˡ; identityʳ to μ-identityʳ) + open Monoid⇒ + + open 𝒞 using (-+-; _+_; _+₁_; i₁; i₂; inject₁; inject₂) + + module _ where + + open Category 𝒞 + open ⇒-Reasoning 𝒞 + open ⊗-Reasoning 𝒞-M.+-monoidal + + module _ {n m o : Obj} where + + private + + +-α : (n + m) + o 𝒞.⇒ n + (m + o) + +-α = +-assoc.to {n} {m} {o} + + +-α∘i₂ : +-α ∘ i₂ ≈ i₂ ∘ i₂ + +-α∘i₂ = inject₂ + + +-α∘i₁∘i₁ : (+-α ∘ i₁) ∘ i₁ ≈ i₁ + +-α∘i₁∘i₁ = inject₁ ⟩∘⟨refl ○ inject₁ + + +-α∘i₁∘i₂ : (+-α ∘ i₁) ∘ i₂ ≈ i₂ ∘ i₁ + +-α∘i₁∘i₂ = inject₁ ⟩∘⟨refl ○ inject₂ + + module _ {n : Obj} where + + +-ρ∘i₁ : +-ρ.from {n} ∘ i₁ ≈ id + +-ρ∘i₁ = inject₁ + + +-λ∘i₂ : +-λ.from {n} ∘ i₂ ≈ id + +-λ∘i₂ = inject₂ + + open S + open ⇒-Reasoning U + open ⊗-Reasoning monoidal + open ⊗-Util.Shorthands monoidal + + ⊲ : {A : 𝒞.Obj} → G.₀ A ⊗₀ G.₀ A ⇒ G.₀ A + ⊲ {A} = μ (M.₀ A) + + ⇒⊲ : {A B : 𝒞.Obj} (f : A 𝒞.⇒ B) → G.₁ f ∘ ⊲ ≈ ⊲ ∘ G.₁ f ⊗₁ G.₁ f + ⇒⊲ f = preserves-μ (M.₁ f) + + ε : {A : 𝒞.Obj} → unit ⇒ G.₀ A + ε {A} = η (M.₀ A) + + ⇒ε : {A B : 𝒞.Obj} (f : A 𝒞.⇒ B) → G.₁ f ∘ ε ≈ ε + ⇒ε f = preserves-η (M.₁ f) + + ⊲-⊗ : {n m : 𝒞.Obj} → G.₀ n ⊗₀ G.₀ m ⇒ G.₀ (n + m) + ⊲-⊗ = ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂ + + module _ {n n′ m m′ : 𝒞.Obj} (f : n 𝒞.⇒ n′) (g : m 𝒞.⇒ m′) where + + open Definitions S.U using (CommutativeSquare) + + left₁ : CommutativeSquare (G.₁ i₁) (G.₁ f) (G.₁ (f +₁ g)) (G.₁ i₁) + left₁ = [ G ]-resp-square inject₁ + + left₂ : CommutativeSquare (G.₁ i₂) (G.₁ g) (G.₁ (f +₁ g)) (G.₁ i₂) + left₂ = [ G ]-resp-square inject₂ + + right : CommutativeSquare ⊲ (G.₁ (f +₁ g) ⊗₁ G.₁ (f +₁ g)) (G.₁ (f +₁ g)) ⊲ + right = ⇒⊲ (f +₁ g) + + ⊲-⊗-commute : + CommutativeSquare + (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + (G.₁ f ⊗₁ G.₁ g) + (G.₁ (f +₁ g)) + (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + ⊲-⊗-commute = glue′ right (parallel left₁ left₂) + + ⊲-⊗-homo : NaturalTransformation (⊗ ∙ (G ⁂ G)) (G ∙ -+-) + ⊲-⊗-homo = ntHelper record + { η = λ (n , m) → ⊲-⊗ {n} {m} + ; commute = λ (f , g) → Equiv.sym (⊲-⊗-commute f g) + } + + ⊲-⊗-α + : {n m o : 𝒞.Obj} + → G.₁ (+-assoc.to {n} {m} {o}) + ∘ (μ (M.₀ ((n + m) + o)) ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + ∘ (μ (M.₀ (n + m)) ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ id + ≈ (μ (M.₀ (n + m + o)) ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + ∘ id ⊗₁ (μ (M.₀ (m + o)) ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + ∘ α⇒ + ⊲-⊗-α {n} {m} {o} = begin + G.₁ +-α ∘ (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ id ≈⟨ refl⟩∘⟨ pullʳ merge₁ʳ ⟩ + G.₁ +-α ∘ ⊲ ∘ (G.₁ i₁ ∘ ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ G.₁ i₂ ≈⟨ extendʳ (⇒⊲ +-α) ⟩ + ⊲ ∘ G.₁ +-α ⊗₁ G.₁ +-α ∘ (G.₁ i₁ ∘ ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ G.₁ i₂ ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟨ + ⊲ ∘ (G.₁ +-α ∘ G.₁ i₁ ∘ ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ (G.₁ +-α ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ pullˡ (Equiv.sym G.homomorphism) ⟩⊗⟨ [ G ]-resp-square +-α∘i₂ ⟩ + ⊲ ∘ (G.₁ (+-α 𝒞.∘ i₁) ∘ ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ extendʳ (⇒⊲ (+-α 𝒞.∘ i₁)) ⟩⊗⟨refl ⟩ + ⊲ ∘ (⊲ ∘ G.₁ (+-α 𝒞.∘ i₁) ⊗₁ G.₁ (+-α 𝒞.∘ i₁) ∘ _) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ ⊗-distrib-over-∘) ⟩⊗⟨refl ⟨ + ⊲ ∘ (⊲ ∘ _ ⊗₁ (G.₁ (+-α 𝒞.∘ i₁) ∘ G.₁ i₂)) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ (refl⟩∘⟨ [ G ]-resp-∘ +-α∘i₁∘i₁ ⟩⊗⟨ [ G ]-resp-square +-α∘i₁∘i₂) ⟩⊗⟨refl ⟩ + ⊲ ∘ (⊲ ∘ G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ G.₁ i₁)) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ split₁ˡ ⟩ + ⊲ ∘ ⊲ ⊗₁ id ∘ (G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ G.₁ i₁)) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ extendʳ (μ-assoc (M.₀ (n + (m + o)))) ⟩ + ⊲ ∘ (id ⊗₁ ⊲ ∘ α⇒) ∘ (G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ G.₁ i₁)) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ assoc ⟩ + ⊲ ∘ id ⊗₁ ⊲ ∘ α⇒ ∘ (G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ G.₁ i₁)) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ assoc-commute-from ⟩ + ⊲ ∘ id ⊗₁ ⊲ ∘ G.₁ i₁ ⊗₁ ((G.₁ i₂ ∘ G.₁ i₁) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂)) ∘ α⇒ ≈⟨ refl⟩∘⟨ pullˡ merge₂ˡ ⟩ + ⊲ ∘ G.₁ i₁ ⊗₁ (⊲ ∘ (G.₁ i₂ ∘ G.₁ i₁) ⊗₁ (G.₁ i₂ ∘ G.₁ i₂)) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (refl⟩∘⟨ ⊗-distrib-over-∘) ⟩∘⟨refl ⟩ + ⊲ ∘ G.₁ i₁ ⊗₁ (⊲ ∘ G.₁ i₂ ⊗₁ G.₁ i₂ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ α⇒ ≈⟨ refl⟩∘⟨ refl⟩⊗⟨ (extendʳ (⇒⊲ i₂)) ⟩∘⟨refl ⟨ + ⊲ ∘ G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ ⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ α⇒ ≈⟨ pushʳ (pushˡ split₂ʳ) ⟩ + (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ id ⊗₁ (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ α⇒ ∎ + where + +-α : (n + m) + o 𝒞.⇒ n + (m + o) + +-α = +-assoc.to {n} {m} {o} + + module _ {A B : 𝒞.Obj} (f : A 𝒞.⇒ B) where + + ⊲-εʳ : ⊲ ∘ G.₁ f ⊗₁ ε ≈ G.₁ f ∘ ρ⇒ + ⊲-εʳ = begin + ⊲ ∘ G.₁ f ⊗₁ ε ≈⟨ refl⟩∘⟨ serialize₂₁ ⟩ + ⊲ ∘ id ⊗₁ ε ∘ G.₁ f ⊗₁ id ≈⟨ pullˡ (Equiv.sym (μ-identityʳ (M.₀ B))) ⟩ + ρ⇒ ∘ G.₁ f ⊗₁ id ≈⟨ unitorʳ-commute-from ⟩ + G.₁ f ∘ ρ⇒ ∎ + + ⊲-εˡ : ⊲ ∘ ε ⊗₁ G.₁ f ≈ G.₁ f ∘ λ⇒ + ⊲-εˡ = begin + ⊲ ∘ ε ⊗₁ G.₁ f ≈⟨ refl⟩∘⟨ serialize₁₂ ⟩ + ⊲ ∘ ε ⊗₁ id ∘ id ⊗₁ G.₁ f ≈⟨ pullˡ (Equiv.sym (μ-identityˡ (M.₀ B))) ⟩ + λ⇒ ∘ id ⊗₁ G.₁ f ≈⟨ unitorˡ-commute-from ⟩ + G.₁ f ∘ λ⇒ ∎ + + module _ {n : 𝒞.Obj} where + + ⊲-⊗-λ : G.₁ (+-λ.from {n}) ∘ ⊲-⊗ ∘ ε ⊗₁ id ≈ λ⇒ + ⊲-⊗-λ = begin + G.₁ +-λ.from ∘ (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ ε ⊗₁ id ≈⟨ refl⟩∘⟨ pullʳ merge₁ʳ ⟩ + G.₁ +-λ.from ∘ ⊲ ∘ (G.₁ i₁ ∘ ε) ⊗₁ G.₁ i₂ ≈⟨ refl⟩∘⟨ refl⟩∘⟨ ⇒ε i₁ ⟩⊗⟨refl ⟩ + G.₁ +-λ.from ∘ ⊲ ∘ ε ⊗₁ G.₁ i₂ ≈⟨ refl⟩∘⟨ ⊲-εˡ i₂ ⟩ + G.₁ +-λ.from ∘ G.₁ i₂ ∘ λ⇒ ≈⟨ cancelˡ ([ G ]-resp-∘ +-λ∘i₂ ○ G.identity) ⟩ + λ⇒ ∎ + + ⊲-⊗-ρ : G.₁ (+-ρ.from {n}) ∘ ⊲-⊗ ∘ id ⊗₁ ε ≈ ρ⇒ + ⊲-⊗-ρ = begin + G.₁ +-ρ.from ∘ (⊲ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) ∘ id ⊗₁ ε ≈⟨ refl⟩∘⟨ pullʳ merge₂ʳ ⟩ + G.₁ +-ρ.from ∘ ⊲ ∘ G.₁ i₁ ⊗₁ (G.₁ i₂ ∘ ε) ≈⟨ refl⟩∘⟨ refl⟩∘⟨ refl⟩⊗⟨ ⇒ε i₂ ⟩ + G.₁ +-ρ.from ∘ ⊲ ∘ G.₁ i₁ ⊗₁ ε ≈⟨ refl⟩∘⟨ ⊲-εʳ i₁ ⟩ + G.₁ +-ρ.from ∘ G.₁ i₁ ∘ ρ⇒ ≈⟨ cancelˡ ([ G ]-resp-∘ +-ρ∘i₁ ○ G.identity) ⟩ + ρ⇒ ∎ + +F,⊗,ε : MonoidalFunctor 𝒞-MC S +F,⊗,ε = record + { F = G + ; isMonoidal = record + { ε = ε + ; ⊗-homo = ⊲-⊗-homo + ; associativity = ⊲-⊗-α + ; unitaryˡ = ⊲-⊗-λ + ; unitaryʳ = ⊲-⊗-ρ + } + } + +module F,⊗,ε = MonoidalFunctor F,⊗,ε diff --git a/Functor/Monoidal/Construction/MultisetOf.agda b/Functor/Monoidal/Construction/MultisetOf.agda new file mode 100644 index 0000000..83bdf52 --- /dev/null +++ b/Functor/Monoidal/Construction/MultisetOf.agda @@ -0,0 +1,89 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category using (Category) +open import Categories.Category.Cocartesian.Bundle using (CocartesianCategory) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Functor using (Functor) renaming (_∘F_ to _∙_) +open import Category.Construction.CMonoids using (CMonoids) + +open import Level using (Level) + +module Functor.Monoidal.Construction.MultisetOf + {o o′ ℓ ℓ′ e e′ : Level} + {𝒞 : CocartesianCategory o ℓ e} + {S : SymmetricMonoidalCategory o′ ℓ′ e′} + (let module 𝒞 = CocartesianCategory 𝒞) + (let module S = SymmetricMonoidalCategory S) + (G : Functor 𝒞.U S.U) + (M : Functor S.U (CMonoids S.symmetric)) + where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Category.Monoidal.Utilities as ⊗-Util +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Object.Monoid.Commutative as CMonoidObject + +open import Categories.Category.Cocartesian using (module CocartesianSymmetricMonoidal) +open import Categories.Category.Monoidal.Symmetric.Properties using (module Shorthands) +open import Categories.Functor.Monoidal using (MonoidalFunctor) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.Functor.Properties using ([_]-resp-∘) +open import Data.Product using (_,_) + +module G = Functor G +module M = Functor M +module 𝒞-SM = CocartesianSymmetricMonoidal 𝒞.U 𝒞.cocartesian + +open 𝒞 using (⊥; -+-; _+_; _+₁_; i₁; i₂; inject₁; inject₂; +-swap) +open Lax using (SymmetricMonoidalFunctor) + +open S +open Functor +open CMonoidObject symmetric using (CommutativeMonoid; CommutativeMonoid⇒) +open CommutativeMonoid renaming (assoc to μ-assoc; identityˡ to μ-identityˡ; identityʳ to μ-identityʳ; commutative to μ-commutative) +open CommutativeMonoid⇒ + +Forget : Functor (CMonoids symmetric) (Monoids monoidal) +Forget .F₀ = monoid +Forget .F₁ = monoid⇒ +Forget .identity = Equiv.refl +Forget .homomorphism = Equiv.refl +Forget .F-resp-≈ x = x + +𝒞-SMC : SymmetricMonoidalCategory o ℓ e +𝒞-SMC = record { symmetric = 𝒞-SM.+-symmetric } + +open import Functor.Monoidal.Construction.ListOf {𝒞 = 𝒞} G (Forget ∙ M) + using (List∘G; ListOf,++,[]; module LG; ++; module List; ++⇒) + +open Shorthands symmetric + +++-swap : {A : Obj} → ++ {A} ≈ ++ ∘ σ⇒ +++-swap {A} = μ-commutative (M.₀ A) + +open ⇒-Reasoning U +open ⊗-Reasoning monoidal + +++-⊗-σ + : {X Y : 𝒞.Obj} + → LG.₁ (+-swap {X} {Y}) ∘ ++ ∘ LG.₁ i₁ ⊗₁ LG.₁ i₂ + ≈ (++ ∘ LG.₁ i₁ ⊗₁ LG.₁ i₂) ∘ σ⇒ +++-⊗-σ {X} {Y} = begin + LG.₁ +-swap ∘ ++ ∘ LG.₁ i₁ ⊗₁ LG.₁ i₂ ≈⟨ extendʳ (++⇒ (G.₁ +-swap)) ⟩ + ++ ∘ LG.₁ +-swap ⊗₁ LG.₁ +-swap ∘ LG.₁ i₁ ⊗₁ LG.₁ i₂ ≈⟨ refl⟩∘⟨ ⊗-distrib-over-∘ ⟨ + ++ ∘ (LG.₁ +-swap ∘ LG.₁ i₁) ⊗₁ (LG.₁ +-swap ∘ LG.₁ i₂) ≈⟨ refl⟩∘⟨ [ List∘G ]-resp-∘ inject₁ ⟩⊗⟨ [ List∘G ]-resp-∘ inject₂ ⟩ + ++ ∘ LG.₁ i₂ ⊗₁ LG.₁ i₁ ≈⟨ pushˡ ++-swap ⟩ + ++ ∘ σ⇒ ∘ LG.₁ i₂ ⊗₁ LG.₁ i₁ ≈⟨ pushʳ (braiding.⇒.commute (LG.₁ i₂ , LG.₁ i₁ )) ⟩ + (++ ∘ LG.₁ i₁ ⊗₁ LG.₁ i₂) ∘ σ⇒ ∎ + +open SymmetricMonoidalFunctor + +module ListOf,++,[] = MonoidalFunctor ListOf,++,[] + +MultisetOf,++,[] : SymmetricMonoidalFunctor 𝒞-SMC S +MultisetOf,++,[] .F = List∘G +MultisetOf,++,[] .isBraidedMonoidal = record + { isMonoidal = ListOf,++,[].isMonoidal + ; braiding-compat = ++-⊗-σ + } diff --git a/Functor/Monoidal/Instance/Nat/Circ.agda b/Functor/Monoidal/Instance/Nat/Circ.agda new file mode 100644 index 0000000..1b45a75 --- /dev/null +++ b/Functor/Monoidal/Instance/Nat/Circ.agda @@ -0,0 +1,87 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_; 0ℓ; suc) + +module Functor.Monoidal.Instance.Nat.Circ where + +import Categories.Object.Monoid as MonoidObject +import Data.Permutation.Sort as ↭-Sort +import Function.Reasoning as →-Reasoning + +open import Category.Instance.Setoids.SymmetricMonoidal {suc 0ℓ} {suc 0ℓ} using (Setoids-×) +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +open import Category.Monoidal.Instance.Nat using (Nat,+,0) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Instance.Nat using (Nat; Nat-Cocartesian) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Data.Setoid.Unit using (⊤ₛ) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Cartesian using (Cartesian) +open Cartesian (Setoids-Cartesian {suc 0ℓ} {suc 0ℓ}) using (products) +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Functor using (_∘F_) +open BinaryProducts products using (-×-) +open import Categories.Category.Product using (_⁂_) +open import Categories.Category.Cocartesian using (Cocartesian) +open import Categories.Category.Instance.Nat using (Nat-Cocartesian) +open import Categories.Functor.Monoidal.Symmetric using (module Lax) +open import Categories.Functor using (Functor) +open import Categories.Category.Cocartesian.Bundle using (CocartesianCategory) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Data.Circuit using (Circuit; Circuitₛ; mkCircuit; mkCircuitₛ; _≈_; mk≈; map) +open import Data.Circuit.Gate using (Gates) +open import Data.Nat using (ℕ; _+_) +open import Data.Product using (_,_) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Function using (_⟶ₛ_; Func; _⟨$⟩_; _∘_; id) +open import Functor.Instance.Nat.Circ {suc 0ℓ} using (Circ; module Multiset∘Edge) +open import Functor.Instance.Nat.Edge {suc 0ℓ} using (Edge) +open import Function.Construct.Setoid using (_∙_) + +module Setoids-× = SymmetricMonoidalCategory Setoids-× + +open import Functor.Instance.FreeCMonoid {suc 0ℓ} {suc 0ℓ} using (FreeCMonoid) + +Nat-Cocartesian-Category : CocartesianCategory 0ℓ 0ℓ 0ℓ +Nat-Cocartesian-Category = record { cocartesian = Nat-Cocartesian } + +open import Functor.Monoidal.Construction.MultisetOf + {𝒞 = Nat-Cocartesian-Category} (Edge Gates) FreeCMonoid using (MultisetOf,++,[]) + +open Lax using (SymmetricMonoidalFunctor) + +module MultisetOf,++,[] = SymmetricMonoidalFunctor MultisetOf,++,[] + +open SymmetricMonoidalFunctor + +ε⇒ : ⊤ₛ ⟶ₛ Circuitₛ 0 +ε⇒ = mkCircuitₛ ∙ MultisetOf,++,[].ε + +open Cocartesian Nat-Cocartesian using (-+-) + +open Func + +η : {n m : ℕ} → Circuitₛ n ×ₛ Circuitₛ m ⟶ₛ Circuitₛ (n + m) +η {n} {m} .to (mkCircuit X , mkCircuit Y) = mkCircuit (MultisetOf,++,[].⊗-homo.η (n , m) ⟨$⟩ (X , Y)) +η {n} {m} .cong (mk≈ x , mk≈ y) = mk≈ (cong (MultisetOf,++,[].⊗-homo.η (n , m)) (x , y)) + +⊗-homomorphism : NaturalTransformation (-×- ∘F (Circ ⁂ Circ)) (Circ ∘F -+-) +⊗-homomorphism = ntHelper record + { η = λ (n , m) → η {n} {m} + ; commute = λ { (f , g) {mkCircuit X , mkCircuit Y} → mk≈ (MultisetOf,++,[].⊗-homo.commute (f , g) {X , Y}) } + } + +Circ,⊗,ε : SymmetricMonoidalFunctor Nat,+,0 Setoids-× +Circ,⊗,ε .F = Circ +Circ,⊗,ε .isBraidedMonoidal = record + { isMonoidal = record + { ε = ε⇒ + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {n} {m} {o} {(mkCircuit x , mkCircuit y) , mkCircuit z} → + mk≈ (MultisetOf,++,[].associativity {n} {m} {o} {(x , y) , z}) } + ; unitaryˡ = λ { {n} {_ , mkCircuit x} → mk≈ (MultisetOf,++,[].unitaryˡ {n} {_ , x}) } + ; unitaryʳ = λ { {n} {mkCircuit x , _} → mk≈ (MultisetOf,++,[].unitaryʳ {n} {x , _}) } + } + ; braiding-compat = λ { {n} {m} {mkCircuit x , mkCircuit y} → + mk≈ (MultisetOf,++,[].braiding-compat {n} {m} {x , y}) } + } diff --git a/Functor/Monoidal/Instance/Nat/Preimage.agda b/Functor/Monoidal/Instance/Nat/Preimage.agda new file mode 100644 index 0000000..844df79 --- /dev/null +++ b/Functor/Monoidal/Instance/Nat/Preimage.agda @@ -0,0 +1,164 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Monoidal.Instance.Nat.Preimage where + +open import Category.Monoidal.Instance.Nat using (Natop,+,0; Natop-Cartesian) +open import Categories.Category.Instance.Nat using (Nat-Cocartesian) +open import Data.Setoid.Unit using (⊤ₛ) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cocartesian using (Cocartesian; BinaryCoproducts) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using (_∘F_) +open import Data.Subset.Functional using (Subset) +open import Data.Nat.Base using (ℕ; _+_) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Product.Base using (_,_; _×_; Σ) +open import Data.Vec.Functional using ([]; _++_) +open import Data.Vec.Functional.Properties using (++-cong) +open import Data.Vec.Functional using (Vector; []) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Functor.Instance.Nat.Preimage using (Preimage; Subsetₛ) +open import Level using (0ℓ) + +open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) +open BinaryProducts products using (-×-) +open Cocartesian Nat-Cocartesian using (module Dual; _+₁_; +-assocʳ; +-swap; +₁∘+-swap) +open Dual.op-binaryProducts using () renaming (-×- to -+-; assocˡ∘⟨⟩ to []∘assocʳ; swap∘⟨⟩ to []∘swap) + +open import Data.Fin.Base using (Fin; splitAt; join; _↑ˡ_; _↑ʳ_) +open import Data.Fin.Properties using (splitAt-join; splitAt-↑ˡ) +open import Data.Sum.Base using ([_,_]′; map; map₁; map₂; inj₁; inj₂) +open import Data.Sum.Properties using ([,]-map; [,]-cong; [-,]-cong; [,-]-cong; [,]-∘) +open import Data.Fin.Preimage using (preimage) +open import Function.Base using (_∘_; id) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; module ≡-Reasoning) +open import Data.Bool.Base using (Bool) + +open Func +Preimage-ε : ⊤ₛ {0ℓ} {0ℓ} ⟶ₛ Subsetₛ 0 +to Preimage-ε x = [] +cong Preimage-ε x () + +++ₛ : {n m : ℕ} → Subsetₛ n ×ₛ Subsetₛ m ⟶ₛ Subsetₛ (n + m) +to ++ₛ (xs , ys) = xs ++ ys +cong ++ₛ (≗xs , ≗ys) = ++-cong _ _ ≗xs ≗ys + +preimage-++ + : {n n′ m m′ : ℕ} + (f : Fin n → Fin n′) + (g : Fin m → Fin m′) + {xs : Subset n′} + {ys : Subset m′} + → preimage f xs ++ preimage g ys ≗ preimage (f +₁ g) (xs ++ ys) +preimage-++ {n} {n′} {m} {m′} f g {xs} {ys} e = begin + (xs ∘ f ++ ys ∘ g) e ≡⟨ [,]-map (splitAt n e) ⟨ + [ xs , ys ]′ (map f g (splitAt n e)) ≡⟨ ≡.cong [ xs , ys ]′ (splitAt-join n′ m′ (map f g (splitAt n e))) ⟨ + [ xs , ys ]′ (splitAt n′ (join n′ m′ (map f g (splitAt n e)))) ≡⟨ ≡.cong ([ xs , ys ]′ ∘ splitAt n′) ([,]-map (splitAt n e)) ⟩ + [ xs , ys ]′ (splitAt n′ ((f +₁ g) e)) ∎ + where + open ≡-Reasoning + +⊗-homomorphism : NaturalTransformation (-×- ∘F (Preimage ⁂ Preimage)) (Preimage ∘F -+-) +⊗-homomorphism = ntHelper record + { η = λ (n , m) → ++ₛ {n} {m} + ; commute = λ { {n′ , m′} {n , m} (f , g) {xs , ys} e → preimage-++ f g e } + } + +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×) +open import Categories.Functor.Monoidal.Symmetric Natop,+,0 Setoids-× using (module Lax) +open Lax using (SymmetricMonoidalFunctor) + +++-assoc + : {m n o : ℕ} + (X : Subset m) + (Y : Subset n) + (Z : Subset o) + → ((X ++ Y) ++ Z) ∘ +-assocʳ {m} ≗ X ++ (Y ++ Z) +++-assoc {m} {n} {o} X Y Z i = begin + ((X ++ Y) ++ Z) (+-assocʳ {m} i) ≡⟨⟩ + [ [ X , Y ]′ ∘ splitAt m , Z ]′ (splitAt (m + n) (+-assocʳ {m} i)) ≡⟨ [,]-cong ([,]-cong (inv ∘ X) (inv ∘ Y) ∘ splitAt m) (inv ∘ Z) (splitAt (m + n) (+-assocʳ {m} i)) ⟨ + [ [ b ∘ X′ , b ∘ Y′ ]′ ∘ splitAt m , b ∘ Z′ ]′ (splitAt _ (+-assocʳ {m} i)) ≡⟨ [-,]-cong ([,]-∘ b ∘ splitAt m) (splitAt (m + n) (+-assocʳ {m} i)) ⟨ + [ b ∘ [ X′ , Y′ ]′ ∘ splitAt m , b ∘ Z′ ]′ (splitAt _ (+-assocʳ {m} i)) ≡⟨ [,]-∘ b (splitAt (m + n) (+-assocʳ {m} i)) ⟨ + b ([ [ X′ , Y′ ]′ ∘ splitAt m , Z′ ]′ (splitAt _ (+-assocʳ {m} i))) ≡⟨ ≡.cong b ([]∘assocʳ {2} {m} i) ⟩ + b ([ X′ , [ Y′ , Z′ ]′ ∘ splitAt n ]′ (splitAt m i)) ≡⟨ [,]-∘ b (splitAt m i) ⟩ + [ b ∘ X′ , b ∘ [ Y′ , Z′ ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,-]-cong ([,]-∘ b ∘ splitAt n) (splitAt m i) ⟩ + [ b ∘ X′ , [ b ∘ Y′ , b ∘ Z′ ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨ [,]-cong (inv ∘ X) ([,]-cong (inv ∘ Y) (inv ∘ Z) ∘ splitAt n) (splitAt m i) ⟩ + [ X , [ Y , Z ]′ ∘ splitAt n ]′ (splitAt m i) ≡⟨⟩ + (X ++ (Y ++ Z)) i ∎ + where + open Bool + open Fin + f : Bool → Fin 2 + f false = zero + f true = suc zero + b : Fin 2 → Bool + b zero = false + b (suc zero) = true + inv : b ∘ f ≗ id + inv false = ≡.refl + inv true = ≡.refl + X′ : Fin m → Fin 2 + X′ = f ∘ X + Y′ : Fin n → Fin 2 + Y′ = f ∘ Y + Z′ : Fin o → Fin 2 + Z′ = f ∘ Z + open ≡-Reasoning + +Preimage-unitaryˡ + : {n : ℕ} + (X : Subset n) + → (X ++ []) ∘ (_↑ˡ 0) ≗ X +Preimage-unitaryˡ {n} X i = begin + [ X , [] ]′ (splitAt _ (i ↑ˡ 0)) ≡⟨ ≡.cong ([ X , [] ]′) (splitAt-↑ˡ n i 0) ⟩ + [ X , [] ]′ (inj₁ i) ≡⟨⟩ + X i ∎ + where + open ≡-Reasoning + +++-swap + : {n m : ℕ} + (X : Subset n) + (Y : Subset m) + → (X ++ Y) ∘ +-swap {n} ≗ Y ++ X +++-swap {n} {m} X Y i = begin + [ X , Y ]′ (splitAt n (+-swap {n} i)) ≡⟨ [,]-cong (inv ∘ X) (inv ∘ Y) (splitAt n (+-swap {n} i)) ⟨ + [ b ∘ X′ , b ∘ Y′ ]′ (splitAt n (+-swap {n} i)) ≡⟨ [,]-∘ b (splitAt n (+-swap {n} i)) ⟨ + b ([ X′ , Y′ ]′ (splitAt n (+-swap {n} i))) ≡⟨ ≡.cong b ([]∘swap {2} {n} i) ⟩ + b ([ Y′ , X′ ]′ (splitAt m i)) ≡⟨ [,]-∘ b (splitAt m i) ⟩ + [ b ∘ Y′ , b ∘ X′ ]′ (splitAt m i) ≡⟨ [,]-cong (inv ∘ Y) (inv ∘ X) (splitAt m i) ⟩ + [ Y , X ]′ (splitAt m i) ∎ + where + open Bool + open Fin + f : Bool → Fin 2 + f false = zero + f true = suc zero + b : Fin 2 → Bool + b zero = false + b (suc zero) = true + inv : b ∘ f ≗ id + inv false = ≡.refl + inv true = ≡.refl + X′ : Fin n → Fin 2 + X′ = f ∘ X + Y′ : Fin m → Fin 2 + Y′ = f ∘ Y + open ≡-Reasoning + +open SymmetricMonoidalFunctor +Preimage,++,[] : SymmetricMonoidalFunctor +Preimage,++,[] .F = Preimage +Preimage,++,[] .isBraidedMonoidal = record + { isMonoidal = record + { ε = Preimage-ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {m} {n} {o} {(X , Y) , Z} i → ++-assoc X Y Z i } + ; unitaryˡ = λ _ → ≡.refl + ; unitaryʳ = λ { {n} {X , _} i → Preimage-unitaryˡ X i } + } + ; braiding-compat = λ { {n} {m} {X , Y} i → ++-swap X Y i } + } diff --git a/Functor/Monoidal/Instance/Nat/Pull.agda b/Functor/Monoidal/Instance/Nat/Pull.agda new file mode 100644 index 0000000..b267f97 --- /dev/null +++ b/Functor/Monoidal/Instance/Nat/Pull.agda @@ -0,0 +1,166 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Monoidal.Instance.Nat.Pull where + +import Categories.Morphism as Morphism + +open import Level using (0ℓ; Level) + +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×) +open import Category.Monoidal.Instance.Nat using (Natop,+,0; Natop-Cartesian) + +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cocartesian using (Cocartesian; BinaryCoproducts) +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Nat using (Nat-Cocartesian) +open import Data.Setoid.Unit using (⊤ₛ) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using (_∘F_) +open import Categories.Functor.Monoidal.Symmetric Natop,+,0 Setoids-× using (module Strong) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism; niHelper) +open import Data.Circuit.Value using (Monoid) +open import Data.Vector using (++-assoc) +open import Data.Fin.Base using (Fin; splitAt; join) +open import Data.Fin.Permutation using (Permutation; _⟨$⟩ʳ_; _⟨$⟩ˡ_) +open import Data.Fin.Preimage using (preimage) +open import Data.Fin.Properties using (splitAt-join; splitAt-↑ˡ; splitAt-↑ʳ; join-splitAt) +open import Data.Nat.Base using (ℕ; _+_) +open import Data.Product.Base using (_,_; _×_; Σ; proj₁; proj₂) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid using (∣_∣) +open import Data.Subset.Functional using (Subset) +open import Data.Sum.Base using ([_,_]′; map; map₁; map₂; inj₁; inj₂) +open import Data.Sum.Properties using ([,]-map; [,]-cong; [-,]-cong; [,-]-cong; [,]-∘) +open import Data.System.Values Monoid using (Values; <ε>; []-unique; _++_; ++ₛ; splitₛ; _≋_; []) +open import Data.Unit.Polymorphic using (tt) +open import Function using (Func; _⟶ₛ_; _⟨$⟩_; _∘_) +open import Function.Construct.Constant using () renaming (function to Const) +open import Functor.Instance.Nat.Pull using (Pull; Pull-defs) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_; module ≡-Reasoning) + +open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) + +open BinaryProducts products using (-×-) +open Cocartesian Nat-Cocartesian using (module Dual; _+₁_; +-assocʳ; +-comm; +-swap; +₁∘+-swap; i₁; i₂) +open Dual.op-binaryProducts using () renaming (-×- to -+-; assocˡ∘⟨⟩ to []∘assocʳ; swap∘⟨⟩ to []∘swap) +open Func +open Morphism (Setoids-×.U) using (_≅_; module Iso) +open Strong using (SymmetricMonoidalFunctor) +open ≡-Reasoning + +private + + open _≅_ + open Iso + + Pull-ε : ⊤ₛ ≅ Values 0 + from Pull-ε = Const ⊤ₛ (Values 0) [] + to Pull-ε = Const (Values 0) ⊤ₛ tt + isoˡ (iso Pull-ε) = tt + isoʳ (iso Pull-ε) {x} = []-unique [] x + + opaque + unfolding _++_ + unfolding Pull-defs + Pull-++ + : {n n′ m m′ : ℕ} + (f : Fin n → Fin n′) + (g : Fin m → Fin m′) + {xs : ∣ Values n′ ∣} + {ys : ∣ Values m′ ∣} + → (Pull.₁ f ⟨$⟩ xs) ++ (Pull.₁ g ⟨$⟩ ys) ≋ Pull.₁ (f +₁ g) ⟨$⟩ (xs ++ ys) + Pull-++ {n} {n′} {m} {m′} f g {xs} {ys} e = begin + (xs ∘ f ++ ys ∘ g) e ≡⟨ [,]-map (splitAt n e) ⟨ + [ xs , ys ]′ (map f g (splitAt n e)) ≡⟨ ≡.cong [ xs , ys ]′ (splitAt-join n′ m′ (map f g (splitAt n e))) ⟨ + (xs ++ ys) (join n′ m′ (map f g (splitAt n e))) ≡⟨ ≡.cong (xs ++ ys) ([,]-map (splitAt n e)) ⟩ + (xs ++ ys) ((f +₁ g) e) ∎ + + module _ {n m : ℕ} where + + opaque + unfolding splitₛ + + open import Function.Construct.Setoid using (setoid) + open module ⇒ₛ {A} {B} = Setoid (setoid {0ℓ} {0ℓ} {0ℓ} {0ℓ} A B) using (_≈_) + open import Function.Construct.Setoid using (_∙_) + open import Function.Construct.Identity using () renaming (function to Id) + + split∘++ : splitₛ ∙ ++ₛ ≈ Id (Values n ×ₛ Values m) + split∘++ {xs , ys} .proj₁ i = ≡.cong [ xs , ys ]′ (splitAt-↑ˡ n i m) + split∘++ {xs , ys} .proj₂ i = ≡.cong [ xs , ys ]′ (splitAt-↑ʳ n m i) + + ++∘split : ++ₛ {n} ∙ splitₛ ≈ Id (Values (n + m)) + ++∘split {x} i = ≡.trans (≡.sym ([,]-∘ x (splitAt n i))) (≡.cong x (join-splitAt n m i)) + + ⊗-homomorphism : NaturalIsomorphism (-×- ∘F (Pull ⁂ Pull)) (Pull ∘F -+-) + ⊗-homomorphism = niHelper record + { η = λ (n , m) → ++ₛ {n} {m} + ; η⁻¹ = λ (n , m) → splitₛ {n} {m} + ; commute = λ { {n , m} {n′ , m′} (f , g) {xs , ys} → Pull-++ f g } + ; iso = λ (n , m) → record + { isoˡ = split∘++ + ; isoʳ = ++∘split + } + } + + module _ {n m : ℕ} where + + opaque + unfolding Pull-++ + + Pull-i₁ + : (X : ∣ Values n ∣) + (Y : ∣ Values m ∣) + → Pull.₁ i₁ ⟨$⟩ (X ++ Y) ≋ X + Pull-i₁ X Y i = ≡.cong [ X , Y ]′ (splitAt-↑ˡ n i m) + + Pull-i₂ + : (X : ∣ Values n ∣) + (Y : ∣ Values m ∣) + → Pull.₁ i₂ ⟨$⟩ (X ++ Y) ≋ Y + Pull-i₂ X Y i = ≡.cong [ X , Y ]′ (splitAt-↑ʳ n m i) + + opaque + unfolding Pull-++ + + Push-assoc + : {m n o : ℕ} + (X : ∣ Values m ∣) + (Y : ∣ Values n ∣) + (Z : ∣ Values o ∣) + → Pull.₁ (+-assocʳ {m} {n} {o}) ⟨$⟩ ((X ++ Y) ++ Z) ≋ X ++ (Y ++ Z) + Push-assoc {m} {n} {o} X Y Z i = ++-assoc X Y Z i + + Pull-swap + : {n m : ℕ} + (X : ∣ Values n ∣) + (Y : ∣ Values m ∣) + → Pull.₁ (+-swap {n}) ⟨$⟩ (X ++ Y) ≋ Y ++ X + Pull-swap {n} {m} X Y i = begin + ((X ++ Y) ∘ +-swap {n}) i ≡⟨ [,]-∘ (X ++ Y) (splitAt m i) ⟩ + [ (X ++ Y) ∘ i₂ , (X ++ Y) ∘ i₁ ]′ (splitAt m i) ≡⟨ [-,]-cong (Pull-i₂ X Y) (splitAt m i) ⟩ + [ Y , (X ++ Y) ∘ i₁ ]′ (splitAt m i) ≡⟨ [,-]-cong (Pull-i₁ X Y) (splitAt m i) ⟩ + [ Y , X ]′ (splitAt m i) ≡⟨⟩ + (Y ++ X) i ∎ + +open SymmetricMonoidalFunctor + +Pull,++,[] : SymmetricMonoidalFunctor +Pull,++,[] .F = Pull +Pull,++,[] .isBraidedMonoidal = record + { isStrongMonoidal = record + { ε = Pull-ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {_} {_} {_} {(X , Y) , Z} → Push-assoc X Y Z } + ; unitaryˡ = λ { {n} {_ , X} → Pull-i₂ {0} {n} [] X } + ; unitaryʳ = λ { {n} {X , _} → Pull-i₁ {n} {0} X [] } + } + ; braiding-compat = λ { {n} {m} {X , Y} → Pull-swap X Y } + } + +module Pull,++,[] = SymmetricMonoidalFunctor Pull,++,[] diff --git a/Functor/Monoidal/Instance/Nat/Push.agda b/Functor/Monoidal/Instance/Nat/Push.agda new file mode 100644 index 0000000..2e8c0cf --- /dev/null +++ b/Functor/Monoidal/Instance/Nat/Push.agda @@ -0,0 +1,209 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Monoidal.Instance.Nat.Push where + +open import Categories.Category.Instance.Nat using (Nat) +open import Data.Bool.Base using (Bool; false) +open import Data.Subset.Functional using (Subset; ⁅_⁆; ⊥) +open import Function.Base using (_∘_; case_of_; _$_; id) +open import Function.Bundles using (Func; _⟶ₛ_; _⟨$⟩_) +open import Level using (0ℓ; Level) +open import Relation.Binary using (Rel; Setoid) +open import Functor.Instance.Nat.Push using (Push; Push-defs) +open import Data.Setoid.Unit using (⊤ₛ) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Data.Vec.Functional as Vec using (Vector) +open import Data.Vector using (++-assoc; ++-↑ˡ; ++-↑ʳ) +-- open import Data.Vec.Functional.Properties using (++-cong) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Function.Construct.Constant using () renaming (function to Const) +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using (products) +open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory) +open import Categories.Category.Instance.Nat using (Nat-Cocartesian) +open import Categories.Category.Cocartesian using (Cocartesian) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using () renaming (_∘F_ to _∘′_) +open Cocartesian Nat-Cocartesian using (module Dual; i₁; i₂; -+-; _+₁_; +-assoc; +-assocʳ; +-assocˡ; +-comm; +-swap; +₁∘+-swap) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Nat using (ℕ; _+_) +open import Data.Fin using (Fin) +open import Data.Product.Base using (_,_; _×_; Σ) +open import Data.Fin.Preimage using (preimage; preimage-⊥; preimage-cong₂) +open import Functor.Monoidal.Instance.Nat.Preimage using (preimage-++) +open import Data.Sum.Base using ([_,_]; [_,_]′; inj₁; inj₂) +open import Data.Sum.Properties using ([,]-cong; [,-]-cong; [-,]-cong; [,]-∘; [,]-map) +open import Data.Circuit.Merge using (merge-with; merge; merge-⊥; merge-[]; ⁅⁆-++; merge-++; merge-cong₁; merge-cong₂; merge-suc; _when_; join-merge; merge-preimage-ρ; merge-⁅⁆) +open import Data.Circuit.Value using (Value; join; join-comm; join-assoc; Monoid) +open import Data.Fin.Base using (splitAt; _↑ˡ_; _↑ʳ_) renaming (join to joinAt) +open import Data.Fin.Properties using (splitAt-↑ˡ; splitAt-↑ʳ; splitAt⁻¹-↑ˡ; splitAt⁻¹-↑ʳ; ↑ˡ-injective; ↑ʳ-injective; _≟_) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≢_; _≗_; module ≡-Reasoning) +open BinaryProducts products using (-×-) +open Value using (U) +open Bool using (false) + +open import Function.Bundles using (Equivalence) +open import Category.Monoidal.Instance.Nat using (Nat,+,0) +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using (Setoids-×) +open import Categories.Functor.Monoidal.Symmetric Nat,+,0 Setoids-× using (module Lax) +open Lax using (SymmetricMonoidalFunctor) +open import Categories.Morphism Nat using (_≅_) +open import Function.Bundles using (Inverse) +open import Data.Fin.Permutation using (Permutation; _⟨$⟩ʳ_; _⟨$⟩ˡ_) +open Dual.op-binaryProducts using () renaming (assocˡ∘⟨⟩ to []∘assocʳ; swap∘⟨⟩ to []∘swap) +open import Relation.Nullary.Decidable using (does; does-⇔; dec-false) +open import Data.Setoid using (∣_∣) + +open ℕ + +open import Data.System.Values Monoid using (Values; <ε>; ++ₛ; _++_; head; tail; _≋_) + +open Func +open ≡-Reasoning + +private + + Push-ε : ⊤ₛ {0ℓ} {0ℓ} ⟶ₛ Values 0 + Push-ε = Const ⊤ₛ (Values 0) <ε> + + opaque + + unfolding _++_ + + unfolding Push-defs + Push-++ + : {n n′ m m′ : ℕ } + → (f : Fin n → Fin n′) + → (g : Fin m → Fin m′) + → (xs : ∣ Values n ∣) + → (ys : ∣ Values m ∣) + → (Push.₁ f ⟨$⟩ xs) ++ (Push.₁ g ⟨$⟩ ys) + ≋ Push.₁ (f +₁ g) ⟨$⟩ (xs ++ ys) + Push-++ {n} {n′} {m} {m′} f g xs ys i = begin + ((merge xs ∘ preimage f ∘ ⁅_⁆) ++ (merge ys ∘ preimage g ∘ ⁅_⁆)) i + ≡⟨ [,]-cong left right (splitAt n′ i) ⟩ + [ (λ x → merge (xs ++ ys) _) , (λ x → merge (xs ++ ys) _) ]′ (splitAt n′ i) + ≡⟨ [,]-∘ (merge (xs ++ ys) ∘ (preimage (f +₁ g))) (splitAt n′ i) ⟨ + merge (xs ++ ys) (preimage (f +₁ g) ((⁅⁆++⊥ Vec.++ ⊥++⁅⁆) i)) ≡⟨ merge-cong₂ (xs ++ ys) (preimage-cong₂ (f +₁ g) (⁅⁆-++ {n′} i)) ⟩ + merge (xs ++ ys) (preimage (f +₁ g) ⁅ i ⁆) ∎ + where + ⁅⁆++⊥ : Vector (Subset (n′ + m′)) n′ + ⁅⁆++⊥ x = ⁅ x ⁆ Vec.++ ⊥ + ⊥++⁅⁆ : Vector (Subset (n′ + m′)) m′ + ⊥++⁅⁆ x = ⊥ Vec.++ ⁅ x ⁆ + left : (x : Fin n′) → merge xs (preimage f ⁅ x ⁆) ≡ merge (xs ++ ys) (preimage (f +₁ g) (⁅ x ⁆ Vec.++ ⊥)) + left x = begin + merge xs (preimage f ⁅ x ⁆) ≡⟨ join-comm U (merge xs (preimage f ⁅ x ⁆)) ⟩ + join (merge xs (preimage f ⁅ x ⁆)) U ≡⟨ ≡.cong (join (merge _ _)) (merge-⊥ ys) ⟨ + join (merge xs (preimage f ⁅ x ⁆)) (merge ys ⊥) ≡⟨ ≡.cong (join (merge _ _)) (merge-cong₂ ys (preimage-⊥ g)) ⟨ + join (merge xs (preimage f ⁅ x ⁆)) (merge ys (preimage g ⊥)) ≡⟨ merge-++ xs ys (preimage f ⁅ x ⁆) (preimage g ⊥) ⟨ + merge (xs ++ ys) ((preimage f ⁅ x ⁆) Vec.++ (preimage g ⊥)) ≡⟨ merge-cong₂ (xs ++ ys) (preimage-++ f g) ⟩ + merge (xs ++ ys) (preimage (f +₁ g) (⁅ x ⁆ Vec.++ ⊥)) ∎ + right : (x : Fin m′) → merge ys (preimage g ⁅ x ⁆) ≡ merge (xs ++ ys) (preimage (f +₁ g) (⊥ Vec.++ ⁅ x ⁆)) + right x = begin + merge ys (preimage g ⁅ x ⁆) ≡⟨⟩ + join U (merge ys (preimage g ⁅ x ⁆)) ≡⟨ ≡.cong (λ h → join h (merge _ _)) (merge-⊥ xs) ⟨ + join (merge xs ⊥) (merge ys (preimage g ⁅ x ⁆)) ≡⟨ ≡.cong (λ h → join h (merge _ _)) (merge-cong₂ xs (preimage-⊥ f)) ⟨ + join (merge xs (preimage f ⊥)) (merge ys (preimage g ⁅ x ⁆)) ≡⟨ merge-++ xs ys (preimage f ⊥) (preimage g ⁅ x ⁆) ⟨ + merge (xs ++ ys) ((preimage f ⊥) Vec.++ (preimage g ⁅ x ⁆)) ≡⟨ merge-cong₂ (xs ++ ys) (preimage-++ f g) ⟩ + merge (xs ++ ys) (preimage (f +₁ g) (⊥ Vec.++ ⁅ x ⁆)) ∎ + + ⊗-homomorphism : NaturalTransformation (-×- ∘′ (Push ⁂ Push)) (Push ∘′ -+-) + ⊗-homomorphism = ntHelper record + { η = λ (n , m) → ++ₛ {n} {m} + ; commute = λ { (f , g) {xs , ys} → Push-++ f g xs ys } + } + + opaque + + unfolding Push-defs + unfolding _++_ + + Push-assoc + : {m n o : ℕ} + (X : ∣ Values m ∣) + (Y : ∣ Values n ∣) + (Z : ∣ Values o ∣) + → (Push.₁ (+-assocˡ {m} {n} {o}) ⟨$⟩ ((X ++ Y) ++ Z)) ≋ X ++ Y ++ Z + Push-assoc {m} {n} {o} X Y Z i = begin + merge ((X ++ Y) ++ Z) (preimage (+-assocˡ {m}) ⁅ i ⁆) ≡⟨ merge-preimage-ρ ↔-mno ((X ++ Y) ++ Z) ⁅ i ⁆ ⟩ + merge (((X ++ Y) ++ Z) ∘ (+-assocʳ {m})) (⁅ i ⁆) ≡⟨⟩ + merge (((X ++ Y) ++ Z) ∘ (+-assocʳ {m})) (preimage id ⁅ i ⁆) ≡⟨ merge-cong₁ (++-assoc X Y Z) (preimage id ⁅ i ⁆) ⟩ + merge (X ++ (Y ++ Z)) (preimage id ⁅ i ⁆) ≡⟨ Push.identity i ⟩ + (X ++ (Y ++ Z)) i ∎ + where + open Inverse + module +-assoc = _≅_ (+-assoc {m} {n} {o}) + ↔-mno : Permutation ((m + n) + o) (m + (n + o)) + ↔-mno .to = +-assocˡ {m} + ↔-mno .from = +-assocʳ {m} + ↔-mno .to-cong ≡.refl = ≡.refl + ↔-mno .from-cong ≡.refl = ≡.refl + ↔-mno .inverse = (λ { ≡.refl → +-assoc.isoˡ _ }) , λ { ≡.refl → +-assoc.isoʳ _ } + + Push-unitaryˡ + : {n : ℕ} + (X : ∣ Values n ∣) + → Push.₁ id ⟨$⟩ (<ε> ++ X) ≋ X + Push-unitaryˡ = merge-⁅⁆ + + preimage-++′ + : {n m o : ℕ} + (f : Vector (Fin o) n) + (g : Vector (Fin o) m) + (S : Subset o) + → preimage (f Vec.++ g) S ≗ preimage f S Vec.++ preimage g S + preimage-++′ {n} f g S = [,]-∘ S ∘ splitAt n + + Push-unitaryʳ + : {n : ℕ} + (X : ∣ Values n ∣) + → Push.₁ (id Vec.++ (λ())) ⟨$⟩ (X ++ (<ε> {0})) ≋ X + Push-unitaryʳ {n} X i = begin + merge (X ++ <ε>) (preimage (id Vec.++ (λ ())) ⁅ i ⁆) ≡⟨ merge-cong₂ (X Vec.++ <ε>) (preimage-++′ id (λ ()) ⁅ i ⁆) ⟩ + merge (X ++ <ε>) (⁅ i ⁆ Vec.++ preimage (λ ()) ⁅ i ⁆) ≡⟨ merge-++ X <ε> ⁅ i ⁆ (preimage (λ ()) ⁅ i ⁆) ⟩ + join (merge X ⁅ i ⁆) (merge <ε> (preimage (λ ()) ⁅ i ⁆)) ≡⟨ ≡.cong (join (merge X ⁅ i ⁆)) (merge-[] <ε> (preimage (λ ()) ⁅ i ⁆)) ⟩ + join (merge X ⁅ i ⁆) U ≡⟨ join-comm (merge X ⁅ i ⁆) U ⟩ + merge X ⁅ i ⁆ ≡⟨ merge-⁅⁆ X i ⟩ + X i ∎ + + Push-swap + : {n m : ℕ} + (X : ∣ Values n ∣) + (Y : ∣ Values m ∣) + → Push.₁ (+-swap {m}) ⟨$⟩ (X ++ Y) ≋ (Y ++ X) + Push-swap {n} {m} X Y i = begin + merge (X ++ Y) (preimage (+-swap {m}) ⁅ i ⁆) ≡⟨ merge-preimage-ρ n+m↔m+n (X ++ Y) ⁅ i ⁆ ⟩ + merge ((X ++ Y) ∘ +-swap {n}) ⁅ i ⁆ ≡⟨ merge-⁅⁆ ((X ++ Y) ∘ (+-swap {n})) i ⟩ + ((X ++ Y) ∘ +-swap {n}) i ≡⟨ [,]-∘ (X ++ Y) (splitAt m i) ⟩ + [ (X ++ Y) ∘ i₂ , (X ++ Y) ∘ i₁ ]′ (splitAt m i) ≡⟨ [-,]-cong (++-↑ʳ X Y) (splitAt m i) ⟩ + [ Y , (X ++ Y) ∘ i₁ ]′ (splitAt m i) ≡⟨ [,-]-cong (++-↑ˡ X Y) (splitAt m i) ⟩ + [ Y , X ]′ (splitAt m i) ≡⟨⟩ + (Y ++ X) i ∎ + where + open ≡-Reasoning + open Inverse + module +-swap = _≅_ (+-comm {m} {n}) + n+m↔m+n : Permutation (n + m) (m + n) + n+m↔m+n .to = +-swap.to + n+m↔m+n .from = +-swap.from + n+m↔m+n .to-cong ≡.refl = ≡.refl + n+m↔m+n .from-cong ≡.refl = ≡.refl + n+m↔m+n .inverse = (λ { ≡.refl → +-swap.isoˡ _ }) , (λ { ≡.refl → +-swap.isoʳ _ }) + +open SymmetricMonoidalFunctor +Push,++,[] : SymmetricMonoidalFunctor +Push,++,[] .F = Push +Push,++,[] .isBraidedMonoidal = record + { isMonoidal = record + { ε = Push-ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {n} {m} {o} {(X , Y) , Z} → Push-assoc X Y Z } + ; unitaryˡ = λ { {n} {_ , X} → Push-unitaryˡ X } + ; unitaryʳ = λ { {n} {X , _} → Push-unitaryʳ X } + } + ; braiding-compat = λ { {n} {m} {X , Y} → Push-swap X Y } + } + +module Push,++,[] = SymmetricMonoidalFunctor Push,++,[] diff --git a/Functor/Monoidal/Instance/Nat/System.agda b/Functor/Monoidal/Instance/Nat/System.agda new file mode 100644 index 0000000..6659fb3 --- /dev/null +++ b/Functor/Monoidal/Instance/Nat/System.agda @@ -0,0 +1,394 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Monoidal.Instance.Nat.System where + +import Categories.Category.Monoidal.Utilities as ⊗-Util +import Data.Circuit.Value as Value +import Data.Vec.Functional as Vec +import Relation.Binary.PropositionalEquality as ≡ + +open import Level using (0ℓ; suc; Level) + +open import Category.Monoidal.Instance.Nat using (Nat,+,0; Natop,+,0) +open import Categories.Category.Monoidal.Bundle using (SymmetricMonoidalCategory; BraidedMonoidalCategory) +open import Category.Instance.Setoids.SymmetricMonoidal {0ℓ} {0ℓ} using () renaming (Setoids-× to 0ℓ-Setoids-×) +open import Category.Instance.Setoids.SymmetricMonoidal {suc 0ℓ} {suc 0ℓ} using (Setoids-×) + +module Natop,+,0 = SymmetricMonoidalCategory Natop,+,0 renaming (braidedMonoidalCategory to B) +module 0ℓ-Setoids-× = SymmetricMonoidalCategory 0ℓ-Setoids-× renaming (braidedMonoidalCategory to B) + +open import Functor.Monoidal.Instance.Nat.Pull using (Pull,++,[]) +open import Categories.Functor.Monoidal.Braided Natop,+,0.B 0ℓ-Setoids-×.B using (module Strong) + +Pull,++,[]B : Strong.BraidedMonoidalFunctor +Pull,++,[]B = record { isBraidedMonoidal = Pull,++,[].isBraidedMonoidal } +module Pull,++,[]B = Strong.BraidedMonoidalFunctor (record { isBraidedMonoidal = Pull,++,[].isBraidedMonoidal }) + +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Cocartesian using (Cocartesian) +open import Categories.Category.Instance.Nat using (Nat; Nat-Cocartesian; Natop) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Data.Setoid.Unit using (⊤ₛ) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Product using (Product) +open import Categories.Category.Product using (_⁂_) +open import Categories.Functor using (Functor) +open import Categories.Functor using (_∘F_) +open import Categories.Functor.Monoidal.Symmetric Nat,+,0 Setoids-× using (module Lax) +open import Categories.NaturalTransformation.Core using (NaturalTransformation; ntHelper) +open import Data.Circuit.Value using (Monoid) +open import Data.Fin using (Fin) +open import Data.Nat using (ℕ; _+_) +open import Data.Product using (_,_; dmap; _×_) renaming (map to ×-map) +open import Data.Product.Function.NonDependent.Setoid using (_×-function_; proj₁ₛ; proj₂ₛ; <_,_>ₛ; swapₛ) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Setoid using (_⇒ₛ_; ∣_∣) +open import Data.System {suc 0ℓ} using (Systemₛ; System; discrete; _≤_) +open import Data.System.Values Monoid using (++ₛ; splitₛ; Values; ++-cong; _++_; []) +open import Data.System.Values Value.Monoid using (_≋_; module ≋) +open import Data.Unit.Polymorphic using (⊤; tt) +open import Function using (Func; _⟶ₛ_; _⟨$⟩_; _∘_; id; case_of_) +open import Function.Construct.Constant using () renaming (function to Const) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (_∙_; setoid) +open import Functor.Instance.Nat.Pull using (Pull) +open import Functor.Instance.Nat.Push using (Push) +open import Functor.Instance.Nat.System using (Sys; Sys-defs) +open import Functor.Monoidal.Braided.Strong.Properties Pull,++,[]B using (braiding-compat-inv) +open import Functor.Monoidal.Instance.Nat.Push using (Push,++,[]) +open import Functor.Monoidal.Strong.Properties Pull,++,[].monoidalFunctor using (associativity-inv) +open import Functor.Monoidal.Strong.Properties Pull,++,[].monoidalFunctor using (unitaryʳ-inv; unitaryˡ-inv; module Shorthands) +open import Relation.Binary using (Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) + +open module ⇒ₛ {A} {B} = Setoid (setoid {0ℓ} {0ℓ} {0ℓ} {0ℓ} A B) using (_≈_) + +open Cartesian (Setoids-Cartesian {suc 0ℓ} {suc 0ℓ}) using (products) + +open BinaryProducts products using (-×-) +open Cocartesian Nat-Cocartesian using (module Dual; i₁; i₂; -+-; _+₁_; +-assocʳ; +-assocˡ; +-comm; +-swap; +₁∘+-swap; +₁∘i₁; +₁∘i₂) +open Dual.op-binaryProducts using () renaming (×-assoc to +-assoc) +open SymmetricMonoidalCategory using () renaming (braidedMonoidalCategory to B) + +open Func + +Sys-ε : ⊤ₛ {suc 0ℓ} {suc 0ℓ} ⟶ₛ Systemₛ 0 +Sys-ε = Const ⊤ₛ (Systemₛ 0) (discrete 0) + +private + + variable + n m o : ℕ + c₁ c₂ c₃ c₄ c₅ c₆ : Level + ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ ℓ₆ : Level + +_×-⇒_ + : {A : Setoid c₁ ℓ₁} + {B : Setoid c₂ ℓ₂} + {C : Setoid c₃ ℓ₃} + {D : Setoid c₄ ℓ₄} + {E : Setoid c₅ ℓ₅} + {F : Setoid c₆ ℓ₆} + → A ⟶ₛ B ⇒ₛ C + → D ⟶ₛ E ⇒ₛ F + → A ×ₛ D ⟶ₛ B ×ₛ E ⇒ₛ C ×ₛ F +_×-⇒_ f g .to (x , y) = to f x ×-function to g y +_×-⇒_ f g .cong (x , y) = cong f x , cong g y + +⊗ : System n × System m → System (n + m) +⊗ {n} {m} (S₁ , S₂) = record + { S = S₁.S ×ₛ S₂.S + ; fₛ = S₁.fₛ ×-⇒ S₂.fₛ ∙ splitₛ + ; fₒ = ++ₛ ∙ S₁.fₒ ×-function S₂.fₒ + } + where + module S₁ = System S₁ + module S₂ = System S₂ + +opaque + + _~_ : {A B : Setoid 0ℓ 0ℓ} → Func A B → Func A B → Set + _~_ = _≈_ + infix 4 _~_ + + sym-~ + : {A B : Setoid 0ℓ 0ℓ} + {x y : Func A B} + → x ~ y + → y ~ x + sym-~ {A} {B} {x} {y} = 0ℓ-Setoids-×.Equiv.sym {A} {B} {x} {y} + +⊗ₛ + : {n m : ℕ} + → Systemₛ n ×ₛ Systemₛ m ⟶ₛ Systemₛ (n + m) +⊗ₛ .to = ⊗ +⊗ₛ {n} {m} .cong {a , b} {c , d} ((a≤c , c≤a) , (b≤d , d≤b)) = left , right + where + module a = System a + module b = System b + module c = System c + module d = System d + module a≤c = _≤_ a≤c + module b≤d = _≤_ b≤d + module c≤a = _≤_ c≤a + module d≤b = _≤_ d≤b + + open _≤_ + left : ⊗ₛ ⟨$⟩ (a , b) ≤ ⊗ₛ ⟨$⟩ (c , d) + left .⇒S = a≤c.⇒S ×-function b≤d.⇒S + left .≗-fₛ i with (i₁ , i₂) ← splitₛ ⟨$⟩ i = dmap (a≤c.≗-fₛ i₁) (b≤d.≗-fₛ i₂) + left .≗-fₒ = cong ++ₛ ∘ dmap a≤c.≗-fₒ b≤d.≗-fₒ + + right : ⊗ₛ ⟨$⟩ (c , d) ≤ ⊗ₛ ⟨$⟩ (a , b) + right .⇒S = c≤a.⇒S ×-function d≤b.⇒S + right .≗-fₛ i with (i₁ , i₂) ← splitₛ ⟨$⟩ i = dmap (c≤a.≗-fₛ i₁) (d≤b.≗-fₛ i₂) + right .≗-fₒ = cong ++ₛ ∘ dmap c≤a.≗-fₒ d≤b.≗-fₒ + +opaque + + unfolding Sys-defs + + System-⊗-≤ + : {n n′ m m′ : ℕ} + (X : System n) + (Y : System m) + (f : Fin n → Fin n′) + (g : Fin m → Fin m′) + → ⊗ (Sys.₁ f ⟨$⟩ X , Sys.₁ g ⟨$⟩ Y) ≤ Sys.₁ (f +₁ g) ⟨$⟩ ⊗ (X , Y) + System-⊗-≤ {n} {n′} {m} {m′} X Y f g = record + { ⇒S = Id (X.S ×ₛ Y.S) + ; ≗-fₛ = λ i s → cong (X.fₛ ×-⇒ Y.fₛ) (Pull,++,[].⊗-homo.⇐.sym-commute (f , g) {i}) {s} + ; ≗-fₒ = λ (s₁ , s₂) → Push,++,[].⊗-homo.commute (f , g) {X.fₒ′ s₁ , Y.fₒ′ s₂} + } + where + module X = System X + module Y = System Y + + System-⊗-≥ + : {n n′ m m′ : ℕ} + (X : System n) + (Y : System m) + (f : Fin n → Fin n′) + (g : Fin m → Fin m′) + → Sys.₁ (f +₁ g) ⟨$⟩ (⊗ (X , Y)) ≤ ⊗ (Sys.₁ f ⟨$⟩ X , Sys.₁ g ⟨$⟩ Y) + System-⊗-≥ {n} {n′} {m} {m′} X Y f g = record + { ⇒S = Id (X.S ×ₛ Y.S) + ; ≗-fₛ = λ i s → cong (X.fₛ ×-⇒ Y.fₛ) (Pull,++,[].⊗-homo.⇐.commute (f , g) {i}) {s} + ; ≗-fₒ = λ (s₁ , s₂) → Push,++,[].⊗-homo.sym-commute (f , g) {X.fₒ′ s₁ , Y.fₒ′ s₂} + } + where + module X = System X + module Y = System Y + +⊗-homomorphism : NaturalTransformation (-×- ∘F (Sys ⁂ Sys)) (Sys ∘F -+-) +⊗-homomorphism = ntHelper record + { η = λ (n , m) → ⊗ₛ {n} {m} + ; commute = λ { (f , g) {X , Y} → System-⊗-≤ X Y f g , System-⊗-≥ X Y f g } + } + +opaque + + unfolding Sys-defs + + ⊗-assoc-≤ + : (X : System n) + (Y : System m) + (Z : System o) + → Sys.₁ (+-assocˡ {n}) ⟨$⟩ (⊗ (⊗ (X , Y) , Z)) ≤ ⊗ (X , ⊗ (Y , Z)) + ⊗-assoc-≤ {n} {m} {o} X Y Z = record + { ⇒S = assocˡ + ; ≗-fₛ = λ i ((s₁ , s₂) , s₃) → cong (X.fₛ ×-⇒ (Y.fₛ ×-⇒ Z.fₛ) ∙ assocˡ) (associativity-inv {x = i}) {s₁ , s₂ , s₃} + ; ≗-fₒ = λ ((s₁ , s₂) , s₃) → Push,++,[].associativity {x = (X.fₒ′ s₁ , Y.fₒ′ s₂) , Z.fₒ′ s₃} + } + where + open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using () renaming (products to 0ℓ-products) + open BinaryProducts 0ℓ-products using (assocˡ) + + module X = System X + module Y = System Y + module Z = System Z + + ⊗-assoc-≥ + : (X : System n) + (Y : System m) + (Z : System o) + → ⊗ (X , ⊗ (Y , Z)) ≤ Sys.₁ (+-assocˡ {n}) ⟨$⟩ (⊗ (⊗ (X , Y) , Z)) + ⊗-assoc-≥ {n} {m} {o} X Y Z = record + { ⇒S = ×-assocʳ + ; ≗-fₛ = λ i (s₁ , s₂ , s₃) → cong ((X.fₛ ×-⇒ Y.fₛ) ×-⇒ Z.fₛ) (sym-split-assoc {i}) {(s₁ , s₂) , s₃} + ; ≗-fₒ = λ (s₁ , s₂ , s₃) → sym-++-assoc {(X.fₒ′ s₁ , Y.fₒ′ s₂) , Z.fₒ′ s₃} + } + where + open Cartesian (Setoids-Cartesian {0ℓ} {0ℓ}) using () renaming (products to prod) + open BinaryProducts prod using () renaming (assocʳ to ×-assocʳ; assocˡ to ×-assocˡ) + + +-assocℓ : Fin ((n + m) + o) → Fin (n + (m + o)) + +-assocℓ = +-assocˡ {n} {m} {o} + + opaque + + unfolding _~_ + + associativity-inv-~ : splitₛ ×-function Id (Values o) ∙ splitₛ ∙ Pull.₁ +-assocℓ ~ ×-assocʳ ∙ Id (Values n) ×-function splitₛ ∙ splitₛ + associativity-inv-~ {i} = associativity-inv {n} {m} {o} {i} + + associativity-~ : Push.₁ (+-assocˡ {n} {m} {o}) ∙ ++ₛ ∙ ++ₛ ×-function Id (Values o) ~ ++ₛ ∙ Id (Values n) ×-function ++ₛ ∙ ×-assocˡ + associativity-~ {i} = Push,++,[].associativity {n} {m} {o} {i} + + sym-split-assoc-~ : ×-assocʳ ∙ Id (Values n) ×-function splitₛ ∙ splitₛ ~ splitₛ ×-function Id (Values o) ∙ splitₛ ∙ Pull.₁ +-assocℓ + sym-split-assoc-~ = sym-~ associativity-inv-~ + + sym-++-assoc-~ : ++ₛ ∙ Id (Values n) ×-function ++ₛ ∙ ×-assocˡ ~ Push.₁ (+-assocˡ {n} {m} {o}) ∙ ++ₛ ∙ ++ₛ ×-function Id (Values o) + sym-++-assoc-~ = sym-~ associativity-~ + + opaque + + unfolding _~_ + + sym-split-assoc : ×-assocʳ ∙ Id (Values n) ×-function splitₛ ∙ splitₛ ≈ splitₛ ×-function Id (Values o) ∙ splitₛ ∙ Pull.₁ +-assocℓ + sym-split-assoc {i} = sym-split-assoc-~ {i} + + sym-++-assoc : ++ₛ ∙ Id (Values n) ×-function ++ₛ ∙ ×-assocˡ ≈ Push.₁ (+-assocˡ {n} {m} {o}) ∙ ++ₛ ∙ ++ₛ ×-function Id (Values o) + sym-++-assoc {i} = sym-++-assoc-~ + + module X = System X + module Y = System Y + module Z = System Z + + Sys-unitaryˡ-≤ : (X : System n) → Sys.₁ id ⟨$⟩ (⊗ (discrete 0 , X)) ≤ X + Sys-unitaryˡ-≤ {n} X = record + { ⇒S = proj₂ₛ + ; ≗-fₛ = λ i (_ , s) → cong (X.fₛ ∙ proj₂ₛ {A = ⊤ₛ {0ℓ}}) (unitaryˡ-inv {n} {i}) + ; ≗-fₒ = λ (_ , s) → Push,++,[].unitaryˡ {n} {tt , X.fₒ′ s} + } + where + module X = System X + + Sys-unitaryˡ-≥ : (X : System n) → X ≤ Sys.₁ id ⟨$⟩ (⊗ (discrete 0 , X)) + Sys-unitaryˡ-≥ {n} X = record + { ⇒S = < Const X.S ⊤ₛ tt , Id X.S >ₛ + ; ≗-fₛ = λ i s → cong (disc.fₛ ×-⇒ X.fₛ ∙ ε⇒ ×-function Id (Values n)) (sym-split-unitaryˡ {i}) + ; ≗-fₒ = λ s → sym-++-unitaryˡ {_ , X.fₒ′ s} + } + where + module X = System X + open SymmetricMonoidalCategory 0ℓ-Setoids-× using (module Equiv) + open ⊗-Util.Shorthands 0ℓ-Setoids-×.monoidal using (λ⇐) + open Shorthands using (ε⇐; ε⇒) + module disc = System (discrete 0) + sym-split-unitaryˡ + : λ⇐ ≈ ε⇐ ×-function Id (Values n) ∙ splitₛ ∙ Pull.₁ ((λ ()) Vec.++ id) + sym-split-unitaryˡ = + 0ℓ-Setoids-×.Equiv.sym + {Values n} + {⊤ₛ ×ₛ Values n} + {ε⇐ ×-function Id (Values n) ∙ splitₛ ∙ Pull.₁ ((λ ()) Vec.++ id)} + {λ⇐} + (unitaryˡ-inv {n}) + sym-++-unitaryˡ : proj₂ₛ {A = ⊤ₛ {0ℓ} {0ℓ}} ≈ Push.₁ ((λ ()) Vec.++ id) ∙ ++ₛ ∙ Push,++,[].ε ×-function Id (Values n) + sym-++-unitaryˡ = + 0ℓ-Setoids-×.Equiv.sym + {⊤ₛ ×ₛ Values n} + {Values n} + {Push.₁ ((λ ()) Vec.++ id) ∙ ++ₛ ∙ Push,++,[].ε ×-function Id (Values n)} + {proj₂ₛ} + (Push,++,[].unitaryˡ {n}) + + + Sys-unitaryʳ-≤ : (X : System n) → Sys.₁ (id Vec.++ (λ ())) ⟨$⟩ (⊗ {n} {0} (X , discrete 0)) ≤ X + Sys-unitaryʳ-≤ {n} X = record + { ⇒S = proj₁ₛ + ; ≗-fₛ = λ i (s , _) → cong (X.fₛ ∙ proj₁ₛ {B = ⊤ₛ {0ℓ}}) (unitaryʳ-inv {n} {i}) + ; ≗-fₒ = λ (s , _) → Push,++,[].unitaryʳ {n} {X.fₒ′ s , tt} + } + where + module X = System X + + Sys-unitaryʳ-≥ : (X : System n) → X ≤ Sys.₁ (id Vec.++ (λ ())) ⟨$⟩ (⊗ {n} {0} (X , discrete 0)) + Sys-unitaryʳ-≥ {n} X = record + { ⇒S = < Id X.S , Const X.S ⊤ₛ tt >ₛ + ; ≗-fₛ = λ i s → cong (X.fₛ ×-⇒ disc.fₛ ∙ Id (Values n) ×-function ε⇒) sym-split-unitaryʳ {s , tt} + ; ≗-fₒ = λ s → sym-++-unitaryʳ {X.fₒ′ s , tt} + } + where + module X = System X + module disc = System (discrete 0) + open ⊗-Util.Shorthands 0ℓ-Setoids-×.monoidal using (ρ⇐) + open Shorthands using (ε⇐; ε⇒) + sym-split-unitaryʳ + : ρ⇐ ≈ Id (Values n) ×-function ε⇐ ∙ splitₛ ∙ Pull.₁ (id Vec.++ (λ ())) + sym-split-unitaryʳ = + 0ℓ-Setoids-×.Equiv.sym + {Values n} + {Values n ×ₛ ⊤ₛ} + {Id (Values n) ×-function ε⇐ ∙ splitₛ ∙ Pull.₁ (id Vec.++ (λ ()))} + {ρ⇐} + (unitaryʳ-inv {n}) + sym-++-unitaryʳ : proj₁ₛ {B = ⊤ₛ {0ℓ}} ≈ Push.₁ (id Vec.++ (λ ())) ∙ ++ₛ ∙ Id (Values n) ×-function Push,++,[].ε + sym-++-unitaryʳ = + 0ℓ-Setoids-×.Equiv.sym + {Values n ×ₛ ⊤ₛ} + {Values n} + {Push.₁ (id Vec.++ (λ ())) ∙ ++ₛ ∙ Id (Values n) ×-function Push,++,[].ε} + {proj₁ₛ} + (Push,++,[].unitaryʳ {n}) + + Sys-braiding-compat-≤ + : (X : System n) + (Y : System m) + → Sys.₁ (+-swap {m} {n}) ⟨$⟩ (⊗ (X , Y)) ≤ ⊗ (Y , X) + Sys-braiding-compat-≤ {n} {m} X Y = record + { ⇒S = swapₛ + ; ≗-fₛ = λ i (s₁ , s₂) → cong (Y.fₛ ×-⇒ X.fₛ ∙ swapₛ) (braiding-compat-inv {m} {n} {i}) {s₂ , s₁} + ; ≗-fₒ = λ (s₁ , s₂) → Push,++,[].braiding-compat {n} {m} {X.fₒ′ s₁ , Y.fₒ′ s₂} + } + where + module X = System X + module Y = System Y + + Sys-braiding-compat-≥ + : (X : System n) + (Y : System m) + → ⊗ (Y , X) ≤ Sys.₁ (+-swap {m} {n}) ⟨$⟩ ⊗ (X , Y) + Sys-braiding-compat-≥ {n} {m} X Y = record + { ⇒S = swapₛ + ; ≗-fₛ = λ i (s₂ , s₁) → cong (X.fₛ ×-⇒ Y.fₛ) (sym-braiding-compat-inv {i}) + ; ≗-fₒ = λ (s₂ , s₁) → sym-braiding-compat-++ {X.fₒ′ s₁ , Y.fₒ′ s₂} + } + where + module X = System X + module Y = System Y + sym-braiding-compat-inv : swapₛ ∙ splitₛ {m} ≈ splitₛ ∙ Pull.₁ (+-swap {m} {n}) + sym-braiding-compat-inv {i} = + 0ℓ-Setoids-×.Equiv.sym + {Values (m + n)} + {Values n ×ₛ Values m} + {splitₛ ∙ Pull.₁ (+-swap {m} {n})} + {swapₛ ∙ splitₛ {m}} + (λ {j} → braiding-compat-inv {m} {n} {j}) {i} + sym-braiding-compat-++ : ++ₛ {m} ∙ swapₛ ≈ Push.₁ (+-swap {m} {n}) ∙ ++ₛ + sym-braiding-compat-++ {i} = + 0ℓ-Setoids-×.Equiv.sym + {Values n ×ₛ Values m} + {Values (m + n)} + {Push.₁ (+-swap {m} {n}) ∙ ++ₛ} + {++ₛ {m} ∙ swapₛ} + (Push,++,[].braiding-compat {n} {m}) + +open Lax.SymmetricMonoidalFunctor + +Sys,⊗,ε : Lax.SymmetricMonoidalFunctor +Sys,⊗,ε .F = Sys +Sys,⊗,ε .isBraidedMonoidal = record + { isMonoidal = record + { ε = Sys-ε + ; ⊗-homo = ⊗-homomorphism + ; associativity = λ { {n} {m} {o} {(X , Y), Z} → ⊗-assoc-≤ X Y Z , ⊗-assoc-≥ X Y Z } + ; unitaryˡ = λ { {n} {_ , X} → Sys-unitaryˡ-≤ X , Sys-unitaryˡ-≥ X } + ; unitaryʳ = λ { {n} {X , _} → Sys-unitaryʳ-≤ X , Sys-unitaryʳ-≥ X } + } + ; braiding-compat = λ { {n} {m} {X , Y} → Sys-braiding-compat-≤ X Y , Sys-braiding-compat-≥ X Y } + } + +module Sys,⊗,ε = Lax.SymmetricMonoidalFunctor Sys,⊗,ε diff --git a/Functor/Monoidal/Strong/Properties.agda b/Functor/Monoidal/Strong/Properties.agda new file mode 100644 index 0000000..9eb7579 --- /dev/null +++ b/Functor/Monoidal/Strong/Properties.agda @@ -0,0 +1,104 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level) +open import Categories.Category.Monoidal using (MonoidalCategory) +open import Categories.Functor.Monoidal using (StrongMonoidalFunctor) + +module Functor.Monoidal.Strong.Properties + {o o′ ℓ ℓ′ e e′ : Level} + {C : MonoidalCategory o ℓ e} + {D : MonoidalCategory o′ ℓ′ e′} + (F,φ,ε : StrongMonoidalFunctor C D) where + +import Categories.Category.Monoidal.Utilities as ⊗-Utilities +import Categories.Category.Construction.Core as Core + +open import Categories.Functor.Monoidal using (StrongMonoidalFunctor) +open import Categories.Functor.Properties using ([_]-resp-≅) + +private + + module C where + open MonoidalCategory C public + open ⊗-Utilities.Shorthands monoidal public using (α⇐; λ⇐; ρ⇐) + + module D where + open MonoidalCategory D public + open ⊗-Utilities.Shorthands monoidal using (α⇐; λ⇐; ρ⇐) public + open Core.Shorthands U using (_∘ᵢ_; idᵢ; _≈ᵢ_; ⌞_⌟; to-≈; _≅_; module HomReasoningᵢ) public + open ⊗-Utilities monoidal using (_⊗ᵢ_) public + +open D + +open StrongMonoidalFunctor F,φ,ε + +private + + variable + X Y Z : C.Obj + + α : {A B C : Obj} → (A ⊗₀ B) ⊗₀ C ≅ A ⊗₀ (B ⊗₀ C) + α = associator + + Fα : F₀ ((X C.⊗₀ Y) C.⊗₀ Z) ≅ F₀ (X C.⊗₀ (Y C.⊗₀ Z)) + Fα = [ F ]-resp-≅ C.associator + + λ- : {A : Obj} → unit ⊗₀ A ≅ A + λ- = unitorˡ + + Fλ : F₀ (C.unit C.⊗₀ X) ≅ F₀ X + Fλ = [ F ]-resp-≅ C.unitorˡ + + ρ : {A : Obj} → A ⊗₀ unit ≅ A + ρ = unitorʳ + + Fρ : F₀ (X C.⊗₀ C.unit) ≅ F₀ X + Fρ = [ F ]-resp-≅ C.unitorʳ + + φ : F₀ X ⊗₀ F₀ Y ≅ F₀ (X C.⊗₀ Y) + φ = ⊗-homo.FX≅GX + +module Shorthands where + + φ⇒ : F₀ X ⊗₀ F₀ Y ⇒ F₀ (X C.⊗₀ Y) + φ⇒ = ⊗-homo.⇒.η _ + + φ⇐ : F₀ (X C.⊗₀ Y) ⇒ F₀ X ⊗₀ F₀ Y + φ⇐ = ⊗-homo.⇐.η _ + + ε⇒ : unit ⇒ F₀ C.unit + ε⇒ = ε.from + + ε⇐ : F₀ C.unit ⇒ unit + ε⇐ = ε.to + +open Shorthands +open HomReasoning + +associativityᵢ : Fα {X} {Y} {Z} ∘ᵢ φ ∘ᵢ φ ⊗ᵢ idᵢ ≈ᵢ φ ∘ᵢ idᵢ ⊗ᵢ φ ∘ᵢ α +associativityᵢ = ⌞ associativity ⌟ + +associativity-inv : φ⇐ ⊗₁ id ∘ φ⇐ ∘ F₁ (C.α⇐ {X} {Y} {Z}) ≈ α⇐ ∘ id ⊗₁ φ⇐ ∘ φ⇐ +associativity-inv = begin + φ⇐ ⊗₁ id ∘ φ⇐ ∘ F₁ C.α⇐ ≈⟨ sym-assoc ⟩ + (φ⇐ ⊗₁ id ∘ φ⇐) ∘ F₁ C.α⇐ ≈⟨ to-≈ associativityᵢ ⟩ + (α⇐ ∘ id ⊗₁ φ⇐) ∘ φ⇐ ≈⟨ assoc ⟩ + α⇐ ∘ id ⊗₁ φ⇐ ∘ φ⇐ ∎ + +unitaryˡᵢ : Fλ {X} ∘ᵢ φ ∘ᵢ ε ⊗ᵢ idᵢ ≈ᵢ λ- +unitaryˡᵢ = ⌞ unitaryˡ ⌟ + +unitaryˡ-inv : ε⇐ ⊗₁ id ∘ φ⇐ ∘ F₁ (C.λ⇐ {X}) ≈ λ⇐ +unitaryˡ-inv = begin + ε⇐ ⊗₁ id ∘ φ⇐ ∘ F₁ C.λ⇐ ≈⟨ sym-assoc ⟩ + (ε⇐ ⊗₁ id ∘ φ⇐) ∘ F₁ C.λ⇐ ≈⟨ to-≈ unitaryˡᵢ ⟩ + λ⇐ ∎ + +unitaryʳᵢ : Fρ {X} ∘ᵢ φ ∘ᵢ idᵢ ⊗ᵢ ε ≈ᵢ ρ +unitaryʳᵢ = ⌞ unitaryʳ ⌟ + +unitaryʳ-inv : id ⊗₁ ε⇐ ∘ φ⇐ ∘ F₁ (C.ρ⇐ {X}) ≈ ρ⇐ +unitaryʳ-inv = begin + id ⊗₁ ε⇐ ∘ φ⇐ ∘ F₁ C.ρ⇐ ≈⟨ sym-assoc ⟩ + (id ⊗₁ ε⇐ ∘ φ⇐) ∘ F₁ C.ρ⇐ ≈⟨ to-≈ unitaryʳᵢ ⟩ + ρ⇐ ∎ diff --git a/Functor/Properties.agda b/Functor/Properties.agda new file mode 100644 index 0000000..1bd3ba6 --- /dev/null +++ b/Functor/Properties.agda @@ -0,0 +1,77 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Properties where + +import Categories.Morphism.Reasoning as ⇒-Reasoning + +open import Categories.Category using (Category; _[_,_]) +open import Level using (Level) +open import Categories.Morphism.Notation using (_[_≅_]) +open import Categories.Morphism using (_≅_) +open import Categories.Functor using (Functor) +open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_; niHelper) +open import Data.Product using (Σ-syntax; _,_) + +module _ + {o o′ ℓ ℓ′ e e′ : Level} + {𝒞 : Category o ℓ e} + {𝒟 : Category o′ ℓ′ e′} + where + + module 𝒞 = Category 𝒞 + module 𝒟 = Category 𝒟 + + define-by-pw-iso + : (F : Functor 𝒞 𝒟) + (G₀ : 𝒞.Obj → 𝒟.Obj) + → (let module F = Functor F) + → ((X : 𝒞.Obj) → 𝒟 [ F.₀ X ≅ G₀ X ]) + → Σ[ G ∈ Functor 𝒞 𝒟 ] F ≃ G + define-by-pw-iso F G₀ α = G , F≃G + where + open Functor + module F = Functor F + open 𝒟 + open _≅_ + open HomReasoning + open ⇒-Reasoning 𝒟 + + G-homo + : {X Y Z : 𝒞.Obj} + → (f : 𝒞 [ X , Y ]) + → (g : 𝒞 [ Y , Z ]) + → from (α Z) ∘ F.₁ (g 𝒞.∘ f) ∘ to (α X) + ≈ (from (α Z) ∘ F.₁ g ∘ to (α Y)) ∘ from (α Y) ∘ F.₁ f ∘ to (α X) + G-homo {X} {Y} {Z} f g = begin + from (α Z) ∘ F.₁ (g 𝒞.∘ f) ∘ to (α X) ≈⟨ extendʳ (pushʳ F.homomorphism) ⟩ + (from (α Z) ∘ F.₁ g) ∘ F.₁ f ∘ to (α X) ≈⟨ extendˡ (pushˡ (insertʳ (isoˡ (α Y)))) ⟩ + (from (α Z) ∘ F.₁ g ∘ to (α Y)) ∘ from (α Y) ∘ F.₁ f ∘ to (α X) ∎ + + G-resp-≈ + : {X Y : 𝒞.Obj} + → {f g : 𝒞 [ X , Y ]} + → f 𝒞.≈ g + → from (α Y) ∘ F.₁ f ∘ to (α X) + ≈ from (α Y) ∘ F.₁ g ∘ to (α X) + G-resp-≈ f≈g = refl⟩∘⟨ F.F-resp-≈ f≈g ⟩∘⟨refl + + G-identity : {X : 𝒞.Obj} → from (α X) ∘ F.₁ 𝒞.id ∘ to (α X) ≈ id + G-identity {X} = refl⟩∘⟨ (F.identity ⟩∘⟨refl ○ identityˡ) ○ isoʳ (α X) + + G : Functor 𝒞 𝒟 + G .F₀ = G₀ + G .F₁ {X} {Y} f = from (α Y) ∘ F.₁ f ∘ to (α X) + G .identity {X} = G-identity + G .homomorphism {f = f} {g} = G-homo f g + G .F-resp-≈ = G-resp-≈ + + commute : {X Y : 𝒞.Obj} (f : 𝒞 [ X , Y ]) → from (α Y) ∘ F.F₁ f ≈ (from (α Y) ∘ F.₁ f ∘ to (α X)) ∘ from (α X) + commute {X} {Y} f = pushʳ (insertʳ (isoˡ (α X))) + + F≃G : F ≃ G + F≃G = niHelper record + { η = λ X → from (α X) + ; η⁻¹ = λ X → to (α X) + ; commute = commute + ; iso = λ X → iso (α X) + } diff --git a/Nat/Properties.agda b/Nat/Properties.agda index 6e04346..0b3da1e 100644 --- a/Nat/Properties.agda +++ b/Nat/Properties.agda @@ -19,16 +19,15 @@ open import Categories.Object.Coproduct Nat using (Coproduct; IsCoproduct; IsCop open import Coeq Nat using (coequalizer-on-coproduct) open import FinMerge using (glue-iter; glue-unglue-once) open import FinMerge.Properties using (lemma₂; merge-unmerge; unmerge-merge) -open import Util using (compare; less; equal; greater; _<_<_) +open import FinMerge.Util using (compare; less; equal; greater) open Category Nat - makeZeroCoequalizer : {B : ℕ} {f g : 0 ⇒ B} → Coequalizer f g makeZeroCoequalizer = record { arr = id ; isCoequalizer = record - { equality = λ() + { equality = λ () ; coequalize = λ {_} {h} _ → h ; universal = λ _ → refl ; unique = λ h≈i∘id → Equiv.sym h≈i∘id diff --git a/NaturalTransformation/Instance/EmptyList.agda b/NaturalTransformation/Instance/EmptyList.agda new file mode 100644 index 0000000..9f94955 --- /dev/null +++ b/NaturalTransformation/Instance/EmptyList.agda @@ -0,0 +1,35 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module NaturalTransformation.Instance.EmptyList {c ℓ : Level} where + +open import Categories.Functor using (Functor) +open import Categories.Functor.Construction.Constant using (const) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Data.Opaque.List using (Listₛ; []ₛ; mapₛ) +open import Data.Setoid using (_⇒ₛ_) +open import Data.Setoid.Unit using (⊤ₛ) +open import Function using (_⟶ₛ_) +open import Function.Construct.Constant using () renaming (function to Const) +open import Function.Construct.Setoid using (_∙_) +open import Functor.Instance.List {c} {ℓ} using (List) +open import Relation.Binary using (Setoid) + +opaque + + unfolding []ₛ + + map-[]ₛ : {A B : Setoid c ℓ} + → (f : A ⟶ₛ B) + → (open Setoid (⊤ₛ ⇒ₛ Listₛ B)) + → []ₛ ≈ mapₛ f ∙ []ₛ + map-[]ₛ {_} {B} f = refl + where + open Setoid (List.₀ B) + +⊤⇒[] : NaturalTransformation (const ⊤ₛ) List +⊤⇒[] = ntHelper record + { η = λ X → []ₛ + ; commute = map-[]ₛ + } diff --git a/NaturalTransformation/Instance/EmptyMultiset.agda b/NaturalTransformation/Instance/EmptyMultiset.agda new file mode 100644 index 0000000..bfec451 --- /dev/null +++ b/NaturalTransformation/Instance/EmptyMultiset.agda @@ -0,0 +1,34 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module NaturalTransformation.Instance.EmptyMultiset {c ℓ : Level} where + +import Function.Construct.Constant as Const + +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.Functor using (Functor) +open import Data.Setoid.Unit {c} {c ⊔ ℓ} using (⊤ₛ) +open import Categories.Functor.Construction.Constant using (const) +open import Data.Opaque.Multiset using (Multisetₛ; []ₛ; mapₛ) +open import Functor.Instance.Multiset {c} {ℓ} using (Multiset) +open import Function.Construct.Constant using () renaming (function to Const) +open import Relation.Binary using (Setoid) +open import Data.Setoid using (_⇒ₛ_) +open import Function using (Func; _⟶ₛ_) +open import Function.Construct.Setoid using (_∙_) + +opaque + unfolding mapₛ + map-[]ₛ + : {A B : Setoid c ℓ} + → (f : A ⟶ₛ B) + → (open Setoid (⊤ₛ ⇒ₛ Multisetₛ B)) + → []ₛ ≈ mapₛ f ∙ []ₛ + map-[]ₛ {B = B} f = Setoid.refl (Multisetₛ B) + +⊤⇒[] : NaturalTransformation (const ⊤ₛ) Multiset +⊤⇒[] = ntHelper record + { η = λ X → []ₛ {Aₛ = X} + ; commute = map-[]ₛ + } diff --git a/NaturalTransformation/Instance/ListAppend.agda b/NaturalTransformation/Instance/ListAppend.agda new file mode 100644 index 0000000..3f198e1 --- /dev/null +++ b/NaturalTransformation/Instance/ListAppend.agda @@ -0,0 +1,43 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module NaturalTransformation.Instance.ListAppend {c ℓ : Level} where + +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.Category.Product using (_※_) +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Functor using (Functor; _∘F_) +open import Data.Opaque.List as L using (mapₛ; ++ₛ) +open import Data.List.Properties using (map-++) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Product using (_,_) +open import Functor.Instance.List {c} {ℓ} using (List) +open import Function using (Func; _⟶ₛ_; _⟨$⟩_) +open import Relation.Binary using (Setoid) + +open Cartesian (Setoids-Cartesian {c} {c ⊔ ℓ}) using (products) +open BinaryProducts products using (-×-) +open Func + +opaque + + unfolding ++ₛ + + map-++ₛ + : {A B : Setoid c ℓ} + (f : Func A B) + (xs ys : Setoid.Carrier (L.Listₛ A)) + (open Setoid (L.Listₛ B)) + → ++ₛ ⟨$⟩ (mapₛ f ⟨$⟩ xs , mapₛ f ⟨$⟩ ys) ≈ mapₛ f ⟨$⟩ (++ₛ ⟨$⟩ (xs , ys)) + map-++ₛ {_} {B} f xs ys = sym (reflexive (map-++ (to f) xs ys)) + where + open Setoid (List.₀ B) + +++ : NaturalTransformation (-×- ∘F (List ※ List)) List +++ = ntHelper record + { η = λ X → ++ₛ {c} {ℓ} {X} + ; commute = λ { {A} {B} f {xs , ys} → map-++ₛ f xs ys } + } diff --git a/NaturalTransformation/Instance/MultisetAppend.agda b/NaturalTransformation/Instance/MultisetAppend.agda new file mode 100644 index 0000000..f786124 --- /dev/null +++ b/NaturalTransformation/Instance/MultisetAppend.agda @@ -0,0 +1,45 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module NaturalTransformation.Instance.MultisetAppend {c ℓ : Level} where + +import Data.Opaque.List as L + +open import Categories.Category.BinaryProducts using (module BinaryProducts) +open import Categories.Category.Cartesian using (Cartesian) +open import Categories.Category.Monoidal.Instance.Setoids using (Setoids-Cartesian) +open import Categories.Category.Product using (_※_) +open import Categories.Functor using (Functor; _∘F_) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Data.List.Properties using (map-++) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (++⁺) +open import Data.Opaque.Multiset using (Multisetₛ; mapₛ; ++ₛ) +open import Data.Product using (_,_) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Function using (Func; _⟶ₛ_; _⟨$⟩_) +open import Functor.Instance.Multiset {c} {ℓ} using (Multiset) +open import Relation.Binary using (Setoid) + +open Cartesian (Setoids-Cartesian {c} {c ⊔ ℓ}) using (products) +open BinaryProducts products using (-×-) +open Func + +opaque + unfolding ++ₛ mapₛ + + map-++ₛ + : {A B : Setoid c ℓ} + (f : Func A B) + (xs ys : Setoid.Carrier (Multiset.₀ A)) + → (open Setoid (Multiset.₀ B)) + → ++ₛ ⟨$⟩ (mapₛ f ⟨$⟩ xs , mapₛ f ⟨$⟩ ys) ≈ mapₛ f ⟨$⟩ (++ₛ ⟨$⟩ (xs , ys)) + map-++ₛ {A} {B} f xs ys = sym (reflexive (map-++ (to f) xs ys)) + where + open Setoid (Multiset.₀ B) + +++ : NaturalTransformation (-×- ∘F (Multiset ※ Multiset)) Multiset +++ = ntHelper record + { η = λ X → ++ₛ + ; commute = λ { {A} {B} f {xs , ys} → map-++ₛ f xs ys } + } diff --git a/NaturalTransformation/Monoidal/Construction/MonoidValued.agda b/NaturalTransformation/Monoidal/Construction/MonoidValued.agda new file mode 100644 index 0000000..2ef70d4 --- /dev/null +++ b/NaturalTransformation/Monoidal/Construction/MonoidValued.agda @@ -0,0 +1,110 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category using (Category) +open import Categories.Category.Cocartesian using (Cocartesian) +open import Categories.Category.Cocartesian.Bundle using (CocartesianCategory) +open import Categories.Category.Construction.Monoids using (Monoids) +open import Categories.Category.Monoidal.Bundle using (MonoidalCategory) +open import Categories.Functor using (Functor) renaming (_∘F_ to _∙_) +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper; _∘ˡ_) +open import Level using (Level) + +-- A natural transformation between functors F, G +-- from a cocartesian category 𝒞 to Monoids[S] +-- can be turned into a monoidal natural transformation +-- between monoidal functors F′ G′ from 𝒞 to S + +module NaturalTransformation.Monoidal.Construction.MonoidValued + {o o′ ℓ ℓ′ e e′ : Level} + {𝒞 : Category o ℓ e} + (𝒞-+ : Cocartesian 𝒞) + {S : MonoidalCategory o′ ℓ′ e′} + (let module S = MonoidalCategory S) + {M N : Functor 𝒞 (Monoids S.monoidal)} + (α : NaturalTransformation M N) + where + +import Categories.Category.Monoidal.Reasoning as ⊗-Reasoning +import Categories.Morphism.Reasoning as ⇒-Reasoning +import Categories.Object.Monoid as MonoidObject +import Functor.Monoidal.Construction.MonoidValued 𝒞-+ {S = S} as FamilyOfMonoids + +open import Categories.Category using (module Definitions) +open import Categories.Functor.Properties using ([_]-resp-square) +open import Categories.NaturalTransformation.Monoidal using (module Lax) +open import Data.Product using (_,_) +open import Functor.Forgetful.Instance.Monoid S.monoidal using (Forget) + +private + + module α = NaturalTransformation α + module M = Functor M + module N = Functor N + module 𝒞 = CocartesianCategory (record { cocartesian = 𝒞-+ }) + + open 𝒞 using (⊥; i₁; i₂; _+_) + open S + + open MonoidObject monoidal using (Monoid; Monoid⇒) + open Definitions U using (CommutativeSquare) + open FamilyOfMonoids M using (F,⊗,ε) + open FamilyOfMonoids N using () renaming (F,⊗,ε to G,⊗,ε) + open Monoid using (μ; η) + open Monoid⇒ + open ⇒-Reasoning U using (glue′) + open ⊗-Reasoning monoidal + + F : Functor 𝒞 U + F = Forget ∙ M + + G : Functor 𝒞 U + G = Forget ∙ N + + β : NaturalTransformation F G + β = Forget ∘ˡ α + + module F = Functor F + module G = Functor G + module β = NaturalTransformation β + + module _ {A : 𝒞.Obj} where + + εₘ : unit ⇒ F.₀ A + εₘ = η (M.₀ A) + + ⊲ₘ : F.₀ A ⊗₀ F.₀ A ⇒ F.₀ A + ⊲ₘ = μ (M.₀ A) + + εₙ : unit ⇒ G.₀ A + εₙ = η (N.₀ A) + + ⊲ₙ : G.₀ A ⊗₀ G.₀ A ⇒ G.₀ A + ⊲ₙ = μ (N.₀ A) + + ε-compat : β.η ⊥ ∘ εₘ ≈ εₙ + ε-compat = preserves-η (α.η ⊥) + + module _ {n m : 𝒞.Obj} where + + square : CommutativeSquare ⊲ₘ (β.η (n + m) ⊗₁ β.η (n + m)) (β.η (n + m)) ⊲ₙ + square = preserves-μ (α.η (n + m)) + + comm₁ : CommutativeSquare (F.₁ i₁) (β.η n) (β.η (n + m)) (G.₁ i₁) + comm₁ = β.commute i₁ + + comm₂ : CommutativeSquare (F.₁ i₂) (β.η m) (β.η (n + m)) (G.₁ i₂) + comm₂ = β.commute i₂ + + ⊗-homo-compat : CommutativeSquare (⊲ₘ ∘ F.₁ i₁ ⊗₁ F.₁ i₂) (β.η n ⊗₁ β.η m) (β.η (n + m)) (⊲ₙ ∘ G.₁ i₁ ⊗₁ G.₁ i₂) + ⊗-homo-compat = glue′ square (parallel comm₁ comm₂) + +open Lax.MonoidalNaturalTransformation hiding (ε-compat; ⊗-homo-compat) + +β,⊗,ε : Lax.MonoidalNaturalTransformation F,⊗,ε G,⊗,ε +β,⊗,ε .U = β +β,⊗,ε .isMonoidal = record + { ε-compat = ε-compat + ; ⊗-homo-compat = ⊗-homo-compat + } + +module β,⊗,ε = Lax.MonoidalNaturalTransformation β,⊗,ε diff --git a/Object/Monoid/Commutative.agda b/Object/Monoid/Commutative.agda new file mode 100644 index 0000000..8a8c756 --- /dev/null +++ b/Object/Monoid/Commutative.agda @@ -0,0 +1,48 @@ +{-# OPTIONS --without-K --safe #-} + +open import Categories.Category.Core using (Category) +open import Categories.Category.Monoidal.Core using (Monoidal) +open import Categories.Category.Monoidal.Symmetric using (Symmetric) +open import Level using (Level; _⊔_) + +module Object.Monoid.Commutative {o ℓ e : Level} {𝒞 : Category o ℓ e} {M : Monoidal 𝒞} (sym : Symmetric M) where + +open import Categories.Object.Monoid M using (IsMonoid; Monoid; Monoid⇒) + +-- a commutative monoid object in a symmetric monoidal category + +open Category 𝒞 +open Symmetric sym using (module braiding; _⊗₁_) + +record IsCommutativeMonoid (M : Obj) : Set (ℓ ⊔ e) where + + field + isMonoid : IsMonoid M + + open IsMonoid isMonoid public + + field + commutative : μ ≈ μ ∘ braiding.⇒.η _ + +record CommutativeMonoid : Set (o ⊔ ℓ ⊔ e) where + + field + Carrier : Obj + isCommutativeMonoid : IsCommutativeMonoid Carrier + + open IsCommutativeMonoid isCommutativeMonoid public + + monoid : Monoid + monoid = record { isMonoid = isMonoid } + +open CommutativeMonoid + +record CommutativeMonoid⇒ (M M′ : CommutativeMonoid) : Set (ℓ ⊔ e) where + + module M = CommutativeMonoid M + module M′ = CommutativeMonoid M′ + + field + monoid⇒ : Monoid⇒ M.monoid M′.monoid + + open Monoid⇒ monoid⇒ public diff --git a/racket/circuits/expander.rkt b/racket/circuits/expander.rkt new file mode 100644 index 0000000..66948c4 --- /dev/null +++ b/racket/circuits/expander.rkt @@ -0,0 +1,15 @@ +#lang racket/base + +;; Hypergraph module language +(require "hypergraph.rkt") +(provide #%app #%module-begin define new-node) + +;; Don't provide quote or new-edge + +;; Primitive cells (gates) +(define (and a b c) (new-edge 'and a b c)) +(define (or a b c) (new-edge 'or a b c)) +(define (not a b) (new-edge 'not a b)) +(define (zero a) (new-edge 'zero a)) +(define (one a) (new-edge 'one a)) +(provide and or not zero one) diff --git a/racket/circuits/hypergraph.rkt b/racket/circuits/hypergraph.rkt new file mode 100644 index 0000000..e0496e4 --- /dev/null +++ b/racket/circuits/hypergraph.rkt @@ -0,0 +1,36 @@ +#lang racket + +;; Top-level syntax transformer +;; displays hypergraph after evaluating module +(define-syntax-rule (hypergraph-mb expr ...) + (#%module-begin + expr + ... + (displayln `(hypergraph ,node-num)) + (for-each displayln graph))) + +;; Need application, quotation, and defines +(provide quote define #%app) + +;; As well as #%module-begin implicit form +(provide + (rename-out + [hypergraph-mb #%module-begin])) + +;; Internal state +(define graph empty) +(define node-num 0) + +;; Create a fresh node +(define (new-node) + (let ([fresh-num node-num]) + (set! node-num (+ node-num 1)) + fresh-num)) + +;; Create a new hyperedge +(define (new-edge label . nodes) + (set! graph (cons (cons label nodes) graph))) + +;; User code constructs hypergraph using +;; new-node and new-edge +(provide new-node new-edge) diff --git a/racket/circuits/lexer.rkt b/racket/circuits/lexer.rkt new file mode 100644 index 0000000..5d3dc82 --- /dev/null +++ b/racket/circuits/lexer.rkt @@ -0,0 +1,33 @@ +#lang racket/base +(require parser-tools/lex) +(require (prefix-in : parser-tools/lex-sre)) + +(define-tokens basic-tokens (ID)) +(define-empty-tokens punct-tokens (LPAREN RPAREN EOF COMMA SEMICOLON LBRACE RBRACE)) +(define-empty-tokens keyword-tokens (WIRE MODULE)) + +(define-lex-abbrev ident-special-char (char-set "_|~+-^&#!")) + +;; Lexer for circuits DSL +(define circuits-lexer + (lexer + [(eof) (token-EOF)] + ["(" (token-LPAREN)] + [")" (token-RPAREN)] + ["{" (token-LBRACE)] + ["}" (token-RBRACE)] + ["," (token-COMMA)] + [";" (token-SEMICOLON)] + ["wire" (token-WIRE)] + ["module" (token-MODULE)] + [(:: "//" (:* (:- any-char (char-set "\n"))) "\n") + (circuits-lexer input-port)] + [(:: + (:or alphabetic ident-special-char) + (:* (:or alphabetic numeric ident-special-char))) + (token-ID (string->symbol lexeme))] + [whitespace (circuits-lexer input-port)])) + +(provide basic-tokens punct-tokens keyword-tokens) + +(provide circuits-lexer) diff --git a/racket/circuits/main.rkt b/racket/circuits/main.rkt new file mode 100644 index 0000000..6c12362 --- /dev/null +++ b/racket/circuits/main.rkt @@ -0,0 +1,5 @@ +#lang racket/base + +(module reader racket + (require "reader.rkt") + (provide read read-syntax)) diff --git a/racket/circuits/parser.rkt b/racket/circuits/parser.rkt new file mode 100644 index 0000000..6287be3 --- /dev/null +++ b/racket/circuits/parser.rkt @@ -0,0 +1,35 @@ +#lang racket/base +(require parser-tools/yacc) + +;; Needed for tokens +(require "lexer.rkt") + +;; Parser for circuits DSL +(define circuits-parser + (parser + [start decls] + [end EOF] + [error void] + [tokens basic-tokens punct-tokens keyword-tokens] + [grammar + (decls + [(decl) $1] + [(decl decls) (append $1 $2)]) + (decl + [(wire-decl) $1] + [(module-decl) (list $1)] + [(module-inst-decl) (list $1)]) + (idents + [(ID) (list $1)] + [(ID COMMA idents) (cons $1 $3)]) + (wire-decl + [(WIRE idents SEMICOLON) + (map (lambda (x) `(define ,x (new-node))) $2)]) + (module-decl + [(MODULE ID LPAREN idents RPAREN LBRACE decls RBRACE) + `(define (,$2 ,@$4) ,@$7)]) + (module-inst-decl + [(ID LPAREN idents RPAREN SEMICOLON) + `(,$1 ,@$3)])])) + +(provide circuits-parser) diff --git a/racket/circuits/reader.rkt b/racket/circuits/reader.rkt new file mode 100644 index 0000000..66a8afa --- /dev/null +++ b/racket/circuits/reader.rkt @@ -0,0 +1,17 @@ +#lang racket/base + +(require "lexer.rkt" "parser.rkt") + +(define (read in) + (syntax->datum (read-syntax #f in))) + +;; Parse the input and produce a module syntax object +;; to be expanded using the circuits module language +(define (read-syntax path input-port) + (define (tokenizer) (circuits-lexer input-port)) + (define parse-tree (circuits-parser tokenizer)) + (datum->syntax #f + `(module circuits-module circuits/expander + ,@parse-tree))) + +(provide read read-syntax) |
