Skip to content

Commit 1aa9865

Browse files
authored
Add Monoid for Add Point (#845) (#1751)
1 parent 69c93f5 commit 1aa9865

File tree

3 files changed

+117
-0
lines changed

3 files changed

+117
-0
lines changed

CHANGELOG.md

+10
Original file line numberDiff line numberDiff line change
@@ -809,6 +809,11 @@ Deprecated names
809809
New modules
810810
-----------
811811

812+
* Algebraic structures when freely adding an identity element:
813+
```
814+
Algebra.Construct.Add.Identity
815+
```
816+
812817
* Operations for module-like algebraic structures:
813818
```
814819
Algebra.Module.Core
@@ -991,6 +996,11 @@ New modules
991996
Other minor changes
992997
-------------------
993998

999+
* Added new proof to `Data.Maybe.Properties`
1000+
```agda
1001+
<∣>-idem : Idempotent _<∣>_
1002+
```
1003+
9941004
* The module `Algebra` now publicly re-exports the contents of
9951005
`Algebra.Structures.Biased`.
9961006

+103
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,103 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Definition of algebraic structures we get from freely adding an
5+
-- identity element
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --without-K --safe #-}
9+
10+
module Algebra.Construct.Add.Identity where
11+
12+
open import Algebra.Bundles
13+
open import Algebra.Core using (Op₂)
14+
open import Algebra.Definitions
15+
open import Algebra.Structures
16+
open import Relation.Binary.Construct.Add.Point.Equality renaming (_≈∙_ to lift≈)
17+
open import Data.Product using (_,_)
18+
open import Level using (Level; _⊔_)
19+
open import Relation.Binary.Core
20+
open import Relation.Binary.Definitions
21+
open import Relation.Binary.Structures
22+
open import Relation.Nullary.Construct.Add.Point
23+
24+
private
25+
variable
26+
a ℓ : Level
27+
A : Set a
28+
29+
liftOp : Op₂ A Op₂ (Pointed A)
30+
liftOp op [ p ] [ q ] = [ op p q ]
31+
liftOp _ [ p ] ∙ = [ p ]
32+
liftOp _ ∙ [ q ] = [ q ]
33+
liftOp _ ∙ ∙ =
34+
35+
module _ {_≈_ : Rel A ℓ} {op : Op₂ A} (refl-≈ : Reflexive _≈_) where
36+
private
37+
_≈∙_ = lift≈ _≈_
38+
op∙ = liftOp op
39+
40+
lift-≈ : {x y : A} x ≈ y [ x ] ≈∙ [ y ]
41+
lift-≈ = [_]
42+
43+
cong₂ : Congruent₂ _≈_ op Congruent₂ _≈∙_ (op∙)
44+
cong₂ R-cong [ eq-l ] [ eq-r ] = lift-≈ (R-cong eq-l eq-r)
45+
cong₂ R-cong [ eq ] ∙≈∙ = lift-≈ eq
46+
cong₂ R-cong ∙≈∙ [ eq ] = lift-≈ eq
47+
cong₂ R-cong ∙≈∙ ∙≈∙ = ≈∙-refl _≈_ refl-≈
48+
49+
assoc : Associative _≈_ op Associative _≈∙_ (op∙)
50+
assoc assoc [ p ] [ q ] [ r ] = lift-≈ (assoc p q r)
51+
assoc _ [ p ] [ q ] ∙ = ≈∙-refl _≈_ refl-≈
52+
assoc _ [ p ] ∙ [ r ] = ≈∙-refl _≈_ refl-≈
53+
assoc _ [ p ] ∙ ∙ = ≈∙-refl _≈_ refl-≈
54+
assoc _ ∙ [ r ] [ q ] = ≈∙-refl _≈_ refl-≈
55+
assoc _ ∙ [ q ] ∙ = ≈∙-refl _≈_ refl-≈
56+
assoc _ ∙ ∙ [ r ] = ≈∙-refl _≈_ refl-≈
57+
assoc _ ∙ ∙ ∙ = ≈∙-refl _≈_ refl-≈
58+
59+
identityˡ : LeftIdentity _≈∙_ ∙ (op∙)
60+
identityˡ [ p ] = ≈∙-refl _≈_ refl-≈
61+
identityˡ ∙ = ≈∙-refl _≈_ refl-≈
62+
63+
identityʳ : RightIdentity _≈∙_ ∙ (op∙)
64+
identityʳ [ p ] = ≈∙-refl _≈_ refl-≈
65+
identityʳ ∙ = ≈∙-refl _≈_ refl-≈
66+
67+
identity : Identity _≈∙_ ∙ (op∙)
68+
identity = identityˡ , identityʳ
69+
70+
module _ {_≈_ : Rel A ℓ} {op : Op₂ A} where
71+
private
72+
_≈∙_ = lift≈ _≈_
73+
op∙ = liftOp op
74+
75+
isMagma : IsMagma _≈_ op IsMagma _≈∙_ op∙
76+
isMagma M =
77+
record
78+
{ isEquivalence = ≈∙-isEquivalence _≈_ M.isEquivalence
79+
; ∙-cong = cong₂ M.refl M.∙-cong
80+
} where module M = IsMagma M
81+
82+
isSemigroup : IsSemigroup _≈_ op IsSemigroup _≈∙_ op∙
83+
isSemigroup S = record
84+
{ isMagma = isMagma S.isMagma
85+
; assoc = assoc S.refl S.assoc
86+
} where module S = IsSemigroup S
87+
88+
isMonoid : IsSemigroup _≈_ op IsMonoid _≈∙_ op∙ ∙
89+
isMonoid S = record
90+
{ isSemigroup = isSemigroup S
91+
; identity = identity S.refl
92+
} where module S = IsSemigroup S
93+
94+
semigroup : Semigroup a (a ⊔ ℓ) Semigroup a (a ⊔ ℓ)
95+
semigroup S = record
96+
{ Carrier = Pointed S.Carrier
97+
; isSemigroup = isSemigroup S.isSemigroup
98+
} where module S = Semigroup S
99+
100+
monoid : Semigroup a (a ⊔ ℓ) Monoid a (a ⊔ ℓ)
101+
monoid S = record
102+
{ isMonoid = isMonoid S.isSemigroup
103+
} where module S = Semigroup S

src/Data/Maybe/Properties.agda

+4
Original file line numberDiff line numberDiff line change
@@ -114,6 +114,10 @@ module _ {A : Set a} where
114114
<∣>-identity : Identity nothing _<∣>_
115115
<∣>-identity = <∣>-identityˡ , <∣>-identityʳ
116116

117+
<∣>-idem : Idempotent _<∣>_
118+
<∣>-idem (just x) = refl
119+
<∣>-idem nothing = refl
120+
117121
module _ (A : Set a) where
118122

119123
open Structures {A = Maybe A} _≡_

0 commit comments

Comments
 (0)