diff options
| author | Jacques Comeaux <jacquesrcomeaux@protonmail.com> | 2025-11-05 03:14:18 -0600 |
|---|---|---|
| committer | Jacques Comeaux <jacquesrcomeaux@protonmail.com> | 2025-11-05 03:14:18 -0600 |
| commit | 07a14947f6fee3219d575a15938bf33764cce791 (patch) | |
| tree | dc4392ed9327f8e316b752f6eced3f51da864aeb /NaturalTransformation | |
| parent | c4e5d5ee327e917dabb86a7de70bafaffcfa6d6f (diff) | |
Add free commutaive monoid functor
Diffstat (limited to 'NaturalTransformation')
| -rw-r--r-- | NaturalTransformation/Instance/EmptyMultiset.agda | 23 | ||||
| -rw-r--r-- | NaturalTransformation/Instance/MultisetAppend.agda | 46 |
2 files changed, 69 insertions, 0 deletions
diff --git a/NaturalTransformation/Instance/EmptyMultiset.agda b/NaturalTransformation/Instance/EmptyMultiset.agda new file mode 100644 index 0000000..9c3a779 --- /dev/null +++ b/NaturalTransformation/Instance/EmptyMultiset.agda @@ -0,0 +1,23 @@ +{-# 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 Categories.Category.Instance.SingletonSet using (SingletonSetoid) +open import Categories.Functor.Construction.Constant using (const) +open import Data.List using ([]) +open import Functor.Instance.Multiset {c} {ℓ} using (Multiset) +open import Relation.Binary using (Setoid) + +module Multiset = Functor Multiset + +⊤⇒[] : NaturalTransformation (const SingletonSetoid) Multiset +⊤⇒[] = ntHelper record + { η = λ X → Const.function SingletonSetoid (Multiset.₀ X) [] + ; commute = λ {_} {B} f → Setoid.refl (Multiset.₀ B) + } diff --git a/NaturalTransformation/Instance/MultisetAppend.agda b/NaturalTransformation/Instance/MultisetAppend.agda new file mode 100644 index 0000000..b0e8bc4 --- /dev/null +++ b/NaturalTransformation/Instance/MultisetAppend.agda @@ -0,0 +1,46 @@ +{-# OPTIONS --without-K --safe #-} + +open import Level using (Level; _⊔_) + +module NaturalTransformation.Instance.MultisetAppend {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.List using (List; _++_; map) +open import Data.List.Properties using (map-++) +open import Data.List.Relation.Binary.Permutation.Setoid.Properties using (++⁺) +open import Data.Product.Relation.Binary.Pointwise.NonDependent using (_×ₛ_) +open import Data.Product using (_,_) +open import Functor.Instance.Multiset {c} {ℓ} using (Multiset) +open import Function using (Func; _⟶ₛ_) +open import Relation.Binary using (Setoid) + +module Multiset = Functor Multiset + +open Cartesian (Setoids-Cartesian {c} {c ⊔ ℓ}) using (products) +open BinaryProducts products using (-×-) +open Func + +++ₛ : {X : Setoid c ℓ} → Multiset.₀ X ×ₛ Multiset.₀ X ⟶ₛ Multiset.₀ X +++ₛ .to (xs , ys) = xs ++ ys +++ₛ {A} .cong (≈xs , ≈ys) = ++⁺ A ≈xs ≈ys + +map-++ₛ + : {A B : Setoid c ℓ} + (f : Func A B) + (xs ys : List (Setoid.Carrier A)) + → (open Setoid (Multiset.₀ B)) + → map (to f) xs ++ map (to f) ys ≈ map (to f) (xs ++ ys) +map-++ₛ {_} {B} f xs ys = sym (reflexive (map-++ (to f) xs ys)) + where + open Setoid (Multiset.₀ B) + +++ : NaturalTransformation (-×- ∘F (Multiset ※ Multiset)) Multiset +++ = ntHelper record + { η = λ X → ++ₛ {X} + ; commute = λ { {A} {B} f {xs , ys} → map-++ₛ f xs ys } + } |
