Skip to content

Commit 018aa65

Browse files
committed
possible fix for agda#1363
1 parent bde655f commit 018aa65

File tree

6 files changed

+170
-0
lines changed

6 files changed

+170
-0
lines changed

CHANGELOG.md

+11
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,17 @@ Deprecated names
3333
New modules
3434
-----------
3535

36+
* Adding a distinguished point `` (\bu2) to `Monoid`, and using it to define
37+
abstract 'literals' for any `PointedMonoid`, with the intended mode-of-use
38+
being to instantiate the point with `1#` from any (semi)ring-like structure:
39+
```agda
40+
module Algebra.Bundles.Literals
41+
module Algebra.Bundles.Pointed
42+
module Algebra.Structures.Literals
43+
module Algebra.Structures.Pointed
44+
module Algebra.Literals
45+
```
46+
3647
Additions to existing modules
3748
-----------------------------
3849

src/Algebra/Bundles/Literals.agda

+19
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,19 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Algebra Literals, from a PointedMonoid bundle
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Algebra.Bundles.Pointed
10+
11+
module Algebra.Bundles.Literals
12+
{c ℓ} (pointedMonoid : PointedMonoid c ℓ)
13+
where
14+
15+
open PointedMonoid pointedMonoid
16+
17+
-- Re-export `Number` instance defined in Algebra.Structures.Literals
18+
19+
open import Algebra.Structures.Literals isPointedMonoid public using (number)

src/Algebra/Bundles/Pointed.agda

+50
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,50 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- `Pointed` intermediate between `Monoid` and `SemiringWithoutAnnihilatingZero`
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Algebra.Bundles
10+
open import Algebra.Bundles.Raw
11+
open import Algebra.Core
12+
open import Algebra.Structures.Pointed as Pointed using (IsPointedMonoid)
13+
import Algebra.Properties.Monoid.Mult as Mult
14+
open import Data.Nat.Base as ℕ using (ℕ)
15+
open import Data.Unit.Base
16+
open import Level using (Level; suc; _⊔_; Lift)
17+
open import Relation.Binary.Core using (Rel)
18+
19+
module Algebra.Bundles.Pointed where
20+
21+
private
22+
variable
23+
c ℓ : Level
24+
25+
------------------------------------------------------------------------
26+
-- Bundles with 1 binary operation & 2 elements
27+
------------------------------------------------------------------------
28+
29+
record PointedMonoid c ℓ : Set (suc (c ⊔ ℓ)) where
30+
field
31+
rawMonoid : RawMonoid c ℓ
32+
open RawMonoid rawMonoid using (Carrier)
33+
field
34+
: Carrier
35+
isPointedMonoid : IsPointedMonoid rawMonoid •
36+
37+
open IsPointedMonoid isPointedMonoid public
38+
39+
------------------------------------------------------------------------
40+
-- instance from any SemiringWithoutAnnihilatingZero
41+
42+
pointedMonoid : SemiringWithoutAnnihilatingZero c ℓ PointedMonoid c ℓ
43+
pointedMonoid semiringWithoutAnnihilatingZero
44+
= record { isPointedMonoid = isPointedMonoid }
45+
where
46+
open SemiringWithoutAnnihilatingZero semiringWithoutAnnihilatingZero
47+
using (1#; +-rawMonoid; +-isMonoid)
48+
isPointedMonoid : IsPointedMonoid +-rawMonoid 1#
49+
isPointedMonoid = record { isMonoid = +-isMonoid }
50+

src/Algebra/Literals.agda

+20
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,20 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Algebra Literals, from a SemiringWithoutAnnihilatingZero
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Algebra.Bundles using (SemiringWithoutAnnihilatingZero)
10+
11+
module Algebra.Literals {c ℓ}
12+
(semiringWithoutAnnihilatingZero : SemiringWithoutAnnihilatingZero c ℓ)
13+
where
14+
15+
open import Algebra.Bundles.Pointed
16+
17+
-- Re-export `Number` instance defined in Algebra.Bundles.Literals
18+
19+
open import Algebra.Bundles.Literals
20+
(pointedMonoid semiringWithoutAnnihilatingZero) public using (number)

src/Algebra/Structures/Literals.agda

+21
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,21 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Algebra Literals
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
open import Algebra.Bundles.Raw
10+
open import Algebra.Structures.Pointed
11+
12+
module Algebra.Structures.Literals
13+
{c ℓ} {rawMonoid : RawMonoid c ℓ} {•}
14+
(isPointedMonoid : IsPointedMonoid rawMonoid •)
15+
where
16+
17+
open import Agda.Builtin.FromNat
18+
open IsPointedMonoid isPointedMonoid
19+
20+
number : Number Carrier
21+
number = record { Literals }

src/Algebra/Structures/Pointed.agda

+49
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,49 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Defines `IsPointedMonoid`
5+
-- intermediate between `Monoid` and `SemiringWithoutAnnihilatingZero`
6+
--
7+
-- By contrast with the rest of `Algebra.Structures`, this is modelled
8+
-- on an underlying `RawMonoid`, rather than a 'flattened' such signature
9+
------------------------------------------------------------------------
10+
11+
{-# OPTIONS --cubical-compatible --safe #-}
12+
13+
open import Algebra.Bundles.Raw
14+
import Algebra.Structures as Structures
15+
open import Data.Nat.Base as ℕ using (ℕ)
16+
open import Data.Unit.Base
17+
open import Level using (Level; _⊔_; Lift)
18+
19+
import Algebra.Definitions.RawMonoid as Definitions
20+
21+
module Algebra.Structures.Pointed where
22+
23+
private
24+
variable
25+
c ℓ : Level
26+
27+
28+
------------------------------------------------------------------------
29+
-- Structures with 1 binary operation & 2 elements
30+
------------------------------------------------------------------------
31+
32+
record IsPointedMonoid
33+
(rawMonoid : RawMonoid c ℓ)
34+
(• : RawMonoid.Carrier rawMonoid) : Set (c ⊔ ℓ)
35+
where
36+
open RawMonoid rawMonoid public
37+
open Structures _≈_
38+
39+
field
40+
isMonoid : IsMonoid _∙_ ε
41+
42+
_ו : Carrier
43+
n ו = n × • where open Definitions rawMonoid
44+
45+
module Literals where
46+
Constraint : Set c
47+
Constraint _ = Lift c ⊤
48+
fromNat : n {{Constraint n}} Carrier
49+
fromNat n = n ו

0 commit comments

Comments
 (0)