Skip to content

Commit

Permalink
regularise and specify/systematise, the conventions for symbol usage
Browse files Browse the repository at this point in the history
  • Loading branch information
jamesmckinna authored and MatthewDaggitt committed Oct 30, 2023
1 parent aba7a8a commit db6b8f1
Show file tree
Hide file tree
Showing 4 changed files with 213 additions and 175 deletions.
54 changes: 27 additions & 27 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -112,21 +112,21 @@ Non-backwards compatible changes
always true and cannot be assumed in user's code.

* Therefore the definitions have been changed as follows to make all their arguments explicit:
- `LeftCancellative __`
- From: `∀ x {y z} → (x y) ≈ (x z) → y ≈ z`
- To: `∀ x y z → (x y) ≈ (x z) → y ≈ z`
- `LeftCancellative __`
- From: `∀ x {y z} → (x y) ≈ (x z) → y ≈ z`
- To: `∀ x y z → (x y) ≈ (x z) → y ≈ z`

- `RightCancellative __`
- From: `∀ {x} y z → (y x) ≈ (z x) → y ≈ z`
- To: `∀ x y z → (y x) ≈ (z x) → y ≈ z`
- `RightCancellative __`
- From: `∀ {x} y z → (y x) ≈ (z x) → y ≈ z`
- To: `∀ x y z → (y x) ≈ (z x) → y ≈ z`

- `AlmostLeftCancellative e __`
- From: `∀ {x} y z → ¬ x ≈ e → (x y) ≈ (x z) → y ≈ z`
- To: `∀ x y z → ¬ x ≈ e → (x y) ≈ (x z) → y ≈ z`
- `AlmostLeftCancellative e __`
- From: `∀ {x} y z → ¬ x ≈ e → (x y) ≈ (x z) → y ≈ z`
- To: `∀ x y z → ¬ x ≈ e → (x y) ≈ (x z) → y ≈ z`

- `AlmostRightCancellative e __`
- From: `∀ {x} y z → ¬ x ≈ e → (y x) ≈ (z x) → y ≈ z`
- To: `∀ x y z → ¬ x ≈ e → (y x) ≈ (z x) → y ≈ z`
- `AlmostRightCancellative e __`
- From: `∀ {x} y z → ¬ x ≈ e → (y x) ≈ (z x) → y ≈ z`
- To: `∀ x y z → ¬ x ≈ e → (y x) ≈ (z x) → y ≈ z`

* Correspondingly some proofs of the above types will need additional arguments passed explicitly.
Instances can easily be fixed by adding additional underscores, e.g.
Expand Down Expand Up @@ -2152,16 +2152,16 @@ Additions to existing modules

* Added new proofs to `Algebra.Consequences.Propositional`:
```agda
comm+assoc⇒middleFour : Commutative __ → Associative __ → __ MiddleFourExchange __
identity+middleFour⇒assoc : Identity e __ → __ MiddleFourExchange __ → Associative __
identity+middleFour⇒comm : Identity e _+_ → __ MiddleFourExchange _+_ → Commutative __
comm+assoc⇒middleFour : Commutative __ → Associative __ → __ MiddleFourExchange __
identity+middleFour⇒assoc : Identity e __ → __ MiddleFourExchange __ → Associative __
identity+middleFour⇒comm : Identity e _+_ → __ MiddleFourExchange _+_ → Commutative __
```

* Added new proofs to `Algebra.Consequences.Setoid`:
```agda
comm+assoc⇒middleFour : Congruent₂ __ → Commutative __ → Associative __ → __ MiddleFourExchange __
identity+middleFour⇒assoc : Congruent₂ __ → Identity e __ → __ MiddleFourExchange __ → Associative __
identity+middleFour⇒comm : Congruent₂ __ → Identity e _+_ → __ MiddleFourExchange _+_ → Commutative __
comm+assoc⇒middleFour : Congruent₂ __ → Commutative __ → Associative __ → __ MiddleFourExchange __
identity+middleFour⇒assoc : Congruent₂ __ → Identity e __ → __ MiddleFourExchange __ → Associative __
identity+middleFour⇒comm : Congruent₂ __ → Identity e _+_ → __ MiddleFourExchange _+_ → Commutative __
involutive⇒surjective : Involutive f → Surjective f
selfInverse⇒involutive : SelfInverse f → Involutive f
Expand All @@ -2171,15 +2171,15 @@ Additions to existing modules
selfInverse⇒injective : SelfInverse f → Injective f
selfInverse⇒bijective : SelfInverse f → Bijective f
comm+idˡ⇒id : Commutative __ → LeftIdentity e __ → Identity e __
comm+idʳ⇒id : Commutative __ → RightIdentity e __ → Identity e __
comm+zeˡ⇒ze : Commutative __ → LeftZero e __ → Zero e __
comm+zeʳ⇒ze : Commutative __ → RightZero e __ → Zero e __
comm+invˡ⇒inv : Commutative __ → LeftInverse e _⁻¹ __ → Inverse e _⁻¹ __
comm+invʳ⇒inv : Commutative __ → RightInverse e _⁻¹ __ → Inverse e _⁻¹ __
comm+distrˡ⇒distr : Commutative __ → __ DistributesOverˡ _◦_ → __ DistributesOver _◦_
comm+distrʳ⇒distr : Commutative __ → __ DistributesOverʳ _◦_ → __ DistributesOver _◦_
distrib+absorbs⇒distribˡ : Associative __ → Commutative _◦_ → __ Absorbs _◦_ → _◦_ Absorbs __ → _◦_ DistributesOver __ → __ DistributesOverˡ _◦_
comm+idˡ⇒id : Commutative __ → LeftIdentity e __ → Identity e __
comm+idʳ⇒id : Commutative __ → RightIdentity e __ → Identity e __
comm+zeˡ⇒ze : Commutative __ → LeftZero e __ → Zero e __
comm+zeʳ⇒ze : Commutative __ → RightZero e __ → Zero e __
comm+invˡ⇒inv : Commutative __ → LeftInverse e _⁻¹ __ → Inverse e _⁻¹ __
comm+invʳ⇒inv : Commutative __ → RightInverse e _⁻¹ __ → Inverse e _⁻¹ __
comm+distrˡ⇒distr : Commutative __ → __ DistributesOverˡ _◦_ → __ DistributesOver _◦_
comm+distrʳ⇒distr : Commutative __ → __ DistributesOverʳ _◦_ → __ DistributesOver _◦_
distrib+absorbs⇒distribˡ : Associative __ → Commutative _◦_ → __ Absorbs _◦_ → _◦_ Absorbs __ → _◦_ DistributesOver __ → __ DistributesOverˡ _◦_
```

* Added new functions to `Algebra.Construct.DirectProduct`:
Expand Down
38 changes: 38 additions & 0 deletions notes/style-guide.md
Original file line number Diff line number Diff line change
Expand Up @@ -443,6 +443,44 @@ word within a compound word is capitalized except for the first word.
relations should be used over the `¬` symbol (e.g. `m+n≮n` should be
used instead of `¬m+n<n`).

#### Symbols for operators

* The stdlib aims to use a consistent set of notations, governed by a
consistent set of conventions, but sometimes, different
Unicode/emacs-input-method symbols neverthless can be rendered by
identical-*seeming* symbols, so this is an attempt to document these.

* The typical binary operator in the `Algebra` hierarchy, inheriting
from the root `Structure`/`Bundle` `isMagma`/`Magma`, is written as
infix ``, obtained as `\.`, NOT as `\bu2`. Nevertheless, there is
also a 'generic' operator, written as infix `·`, obtained as
`\cdot`. Do NOT attempt to use related, but typographically
indistinguishable, symbols.

* Similarly, 'primed' names and symbols, used to standardise names
apart, or to provide (more) simply-typed versions of
dependently-typed operations, should be written using `\'`, NOT the
unmarked `'` character.

* Likewise, standard infix symbols for eg, divisibility on numeric
datatypes/algebraic structure, should be written `\|`, NOT the
unmarked `|` character. An exception to this is the *strict*
ordering relation, written using `<`, NOT `\<` as might be expected.

* Since v2.0, the `Algebra` hierarchy systematically introduces
consistent symbolic notation for the negated versions of the usual
binary predicates for equality, ordering etc. These are obtained
from the corresponding input sequence by adding `n` to the symbol
name, so that ``, obtained as `\le`, becomes `` obtained as
`\len`, etc.

* Correspondingly, the flipped symbols (and their negations) for the
converse relations are systematically introduced, eg `` as `\ge`
and `` as `\gen`.

* Any exceptions to these conventions should be flagged on the GitHub
`agda-stdlib` issue tracker in the usual way.

#### Functions and relations over specific datatypes

* When defining a new relation `P` over a datatype `X` in a `Data.X.Relation` module,
Expand Down
54 changes: 27 additions & 27 deletions src/Algebra/Consequences/Propositional.agda
Original file line number Diff line number Diff line change
Expand Up @@ -44,16 +44,16 @@ open Base public
------------------------------------------------------------------------
-- Group-like structures

module _ {__ _⁻¹ ε} where
module _ {__ _⁻¹ ε} where

assoc+id+invʳ⇒invˡ-unique : Associative __ Identity ε __
RightInverse ε _⁻¹ __
x y (x y) ≡ ε x ≡ (y ⁻¹)
assoc+id+invʳ⇒invˡ-unique : Associative __ Identity ε __
RightInverse ε _⁻¹ __
x y (x y) ≡ ε x ≡ (y ⁻¹)
assoc+id+invʳ⇒invˡ-unique = Base.assoc+id+invʳ⇒invˡ-unique (cong₂ _)

assoc+id+invˡ⇒invʳ-unique : Associative __ Identity ε __
LeftInverse ε _⁻¹ __
x y (x y) ≡ ε y ≡ (x ⁻¹)
assoc+id+invˡ⇒invʳ-unique : Associative __ Identity ε __
LeftInverse ε _⁻¹ __
x y (x y) ≡ ε y ≡ (x ⁻¹)
assoc+id+invˡ⇒invʳ-unique = Base.assoc+id+invˡ⇒invʳ-unique (cong₂ _)

------------------------------------------------------------------------
Expand All @@ -76,43 +76,43 @@ module _ {_+_ _*_ -_ 0#} where
------------------------------------------------------------------------
-- Bisemigroup-like structures

module _ {__ _◦_ : Op₂ A} (-comm : Commutative __) where
module _ {__ _◦_ : Op₂ A} (-comm : Commutative __) where

comm+distrˡ⇒distrʳ : __ DistributesOverˡ _◦_ __ DistributesOverʳ _◦_
comm+distrˡ⇒distrʳ = Base.comm+distrˡ⇒distrʳ (cong₂ _) -comm
comm+distrˡ⇒distrʳ : __ DistributesOverˡ _◦_ __ DistributesOverʳ _◦_
comm+distrˡ⇒distrʳ = Base.comm+distrˡ⇒distrʳ (cong₂ _) -comm

comm+distrʳ⇒distrˡ : __ DistributesOverʳ _◦_ __ DistributesOverˡ _◦_
comm+distrʳ⇒distrˡ = Base.comm+distrʳ⇒distrˡ (cong₂ _) -comm
comm+distrʳ⇒distrˡ : __ DistributesOverʳ _◦_ __ DistributesOverˡ _◦_
comm+distrʳ⇒distrˡ = Base.comm+distrʳ⇒distrˡ (cong₂ _) -comm

comm⇒sym[distribˡ] : x Symmetric (λ y z (x ◦ (y z)) ≡ ((x ◦ y) (x ◦ z)))
comm⇒sym[distribˡ] = Base.comm⇒sym[distribˡ] (cong₂ _◦_) -comm
comm⇒sym[distribˡ] : x Symmetric (λ y z (x ◦ (y z)) ≡ ((x ◦ y) (x ◦ z)))
comm⇒sym[distribˡ] = Base.comm⇒sym[distribˡ] (cong₂ _◦_) -comm

------------------------------------------------------------------------
-- Selectivity

module _ {__ : Op₂ A} where
module _ {__ : Op₂ A} where

sel⇒idem : Selective __ Idempotent __
sel⇒idem : Selective __ Idempotent __
sel⇒idem = Base.sel⇒idem _≡_

------------------------------------------------------------------------
-- Middle-Four Exchange

module _ {__ : Op₂ A} where
module _ {__ : Op₂ A} where

comm+assoc⇒middleFour : Commutative __ Associative __
__ MiddleFourExchange __
comm+assoc⇒middleFour = Base.comm+assoc⇒middleFour (cong₂ __)
comm+assoc⇒middleFour : Commutative __ Associative __
__ MiddleFourExchange __
comm+assoc⇒middleFour = Base.comm+assoc⇒middleFour (cong₂ __)

identity+middleFour⇒assoc : {e : A} Identity e __
__ MiddleFourExchange __
Associative __
identity+middleFour⇒assoc = Base.identity+middleFour⇒assoc (cong₂ __)
identity+middleFour⇒assoc : {e : A} Identity e __
__ MiddleFourExchange __
Associative __
identity+middleFour⇒assoc = Base.identity+middleFour⇒assoc (cong₂ __)

identity+middleFour⇒comm : {_+_ : Op₂ A} {e : A} Identity e _+_
__ MiddleFourExchange _+_
Commutative __
identity+middleFour⇒comm = Base.identity+middleFour⇒comm (cong₂ __)
__ MiddleFourExchange _+_
Commutative __
identity+middleFour⇒comm = Base.identity+middleFour⇒comm (cong₂ __)

------------------------------------------------------------------------
-- Without Loss of Generality
Expand Down
Loading

0 comments on commit db6b8f1

Please sign in to comment.