diff options
Diffstat (limited to 'Functor/Monoidal/Construction')
| -rw-r--r-- | Functor/Monoidal/Construction/MonoidValued.agda | 214 | ||||
| -rw-r--r-- | Functor/Monoidal/Construction/MultisetOf.agda | 89 |
2 files changed, 303 insertions, 0 deletions
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 = ++-β-Ο + } |
