diff --git a/CHANGELOG.md b/CHANGELOG.md index 02014e8d9f..fd80f6ad53 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -1689,10 +1689,12 @@ Other minor changes deduplicateᵇ : (A → A → Bool) → List A → List A ``` -* Added new functions to `Data.List.Base`: +* Added new functions and definitions to `Data.List.Base`: ```agda catMaybes : List (Maybe A) → List A ap : List (A → B) → List A → List B + ++-rawMagma : Set a → RawMagma a _ + ++-[]-rawMonoid : Set a → RawMonoid a _ ``` * Added new proofs in `Data.List.Relation.Binary.Lex.Strict`: @@ -1725,6 +1727,9 @@ Other minor changes concatMap-pure : concatMap [_] ≗ id concatMap-map : concatMap g (map f xs) ≡ concatMap (g ∘′ f) xs map-concatMap : map f ∘′ concatMap g ≗ concatMap (map f ∘′ g) + + length-isMagmaHomomorphism : (A : Set a) → IsMagmaHomomorphism (++-rawMagma A) +-rawMagma length + length-isMonoidHomomorphism : (A : Set a) → IsMonoidHomomorphism (++-[]-rawMonoid A) +-0-rawMonoid length ``` * Added new patterns and definitions to `Data.Nat.Base`: diff --git a/src/Data/List/Base.agda b/src/Data/List/Base.agda index 3ab3c2f923..2cee335208 100644 --- a/src/Data/List/Base.agda +++ b/src/Data/List/Base.agda @@ -11,6 +11,7 @@ module Data.List.Base where +open import Algebra.Bundles.Raw using (RawMagma; RawMonoid) open import Data.Bool.Base as Bool using (Bool; false; true; not; _∧_; _∨_; if_then_else_) open import Data.Fin.Base using (Fin; zero; suc) @@ -27,6 +28,7 @@ open import Relation.Unary using (Pred; Decidable) open import Relation.Unary.Properties using (∁?) open import Relation.Binary.Core using (Rel) import Relation.Binary.Definitions as B +open import Relation.Binary.PropositionalEquality.Core using (_≡_) private variable @@ -463,6 +465,24 @@ infixl 6 _∷ʳ?_ _∷ʳ?_ : List A → Maybe A → List A xs ∷ʳ? x = maybe′ (xs ∷ʳ_) xs x +------------------------------------------------------------------------ +-- Raw algebraic bundles + +module _ (A : Set a) where + ++-rawMagma : RawMagma a _ + ++-rawMagma = record + { Carrier = List A + ; _≈_ = _≡_ + ; _∙_ = _++_ + } + + ++-[]-rawMonoid : RawMonoid a _ + ++-[]-rawMonoid = record + { Carrier = List A + ; _≈_ = _≡_ + ; _∙_ = _++_ + ; ε = [] + } ------------------------------------------------------------------------ -- DEPRECATED diff --git a/src/Data/List/Properties.agda b/src/Data/List/Properties.agda index d2866cfa81..5cfc4b3635 100644 --- a/src/Data/List/Properties.agda +++ b/src/Data/List/Properties.agda @@ -13,6 +13,7 @@ module Data.List.Properties where open import Algebra.Bundles open import Algebra.Definitions as AlgebraicDefinitions using (Involutive) +open import Algebra.Morphism.Structures using (IsMagmaHomomorphism; IsMonoidHomomorphism) import Algebra.Structures as AlgebraicStructures open import Data.Bool.Base using (Bool; false; true; not; if_then_else_) open import Data.Fin.Base using (Fin; zero; suc; cast; toℕ) @@ -240,6 +241,22 @@ module _ (A : Set a) where ; isMonoid = ++-isMonoid } +module _ (A : Set a) where + + length-isMagmaHomomorphism : IsMagmaHomomorphism (++-rawMagma A) +-rawMagma length + length-isMagmaHomomorphism = record + { isRelHomomorphism = record + { cong = cong length + } + ; homo = λ xs ys → length-++ xs {ys} + } + + length-isMonoidHomomorphism : IsMonoidHomomorphism (++-[]-rawMonoid A) +-0-rawMonoid length + length-isMonoidHomomorphism = record + { isMagmaHomomorphism = length-isMagmaHomomorphism + ; ε-homo = refl + } + ------------------------------------------------------------------------ -- alignWith