Skip to content

Commit d992900

Browse files
jamesmckinnaandreasabel
authored andcommitted
Add the Setoid-based Monoid on (List, [], _++_) (#2393)
* refactored from #2350 * enriched the `module` contents in response to review comments
1 parent aef9afc commit d992900

File tree

2 files changed

+64
-0
lines changed

2 files changed

+64
-0
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -113,6 +113,11 @@ New modules
113113
Algebra.Module.Bundles.Raw
114114
```
115115

116+
* Properties of `List` modulo `Setoid` equality (currently only the ([],++) monoid):
117+
```
118+
Data.List.Relation.Binary.Equality.Setoid.Properties
119+
```
120+
116121
* Refactoring of `Data.List.Base.{scanr|scanl}` and their properties:
117122
```
118123
Data.List.Scans.Base
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,59 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Properties of List modulo ≋
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Relation.Binary.Bundles using (Setoid)
10+
11+
module Data.List.Relation.Binary.Equality.Setoid.Properties
12+
{c ℓ} (S : Setoid c ℓ)
13+
where
14+
15+
open import Algebra.Bundles using (Magma; Semigroup; Monoid)
16+
import Algebra.Structures as Structures
17+
open import Data.List.Base using (List; []; _++_)
18+
import Data.List.Properties as List
19+
import Data.List.Relation.Binary.Equality.Setoid as ≋
20+
open import Data.Product.Base using (_,_)
21+
open import Function.Base using (_∘_)
22+
open import Level using (_⊔_)
23+
24+
open S using (_≋_; ≋-refl; ≋-reflexive; ≋-isEquivalence; ++⁺)
25+
open Structures _≋_ using (IsMagma; IsSemigroup; IsMonoid)
26+
27+
------------------------------------------------------------------------
28+
-- The []-++-Monoid
29+
30+
-- Structures
31+
32+
isMagma : IsMagma _++_
33+
isMagma = record
34+
{ isEquivalence = ≋-isEquivalence
35+
; ∙-cong = ++⁺
36+
}
37+
38+
isSemigroup : IsSemigroup _++_
39+
isSemigroup = record
40+
{ isMagma = isMagma
41+
; assoc = λ xs ys zs ≋-reflexive (List.++-assoc xs ys zs)
42+
}
43+
44+
isMonoid : IsMonoid _++_ []
45+
isMonoid = record
46+
{ isSemigroup = isSemigroup
47+
; identity = (λ _ ≋-refl) , ≋-reflexive ∘ List.++-identityʳ
48+
}
49+
50+
-- Bundles
51+
52+
magma : Magma c (c ⊔ ℓ)
53+
magma = record { isMagma = isMagma }
54+
55+
semigroup : Semigroup c (c ⊔ ℓ)
56+
semigroup = record { isSemigroup = isSemigroup }
57+
58+
monoid : Monoid c (c ⊔ ℓ)
59+
monoid = record { isMonoid = isMonoid }

0 commit comments

Comments
 (0)