Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[ refactor ] Restate, and use, the definitions of Monotonic etc. operations #2580

Draft
wants to merge 23 commits into
base: master
Choose a base branch
from
Draft
Show file tree
Hide file tree
Changes from 19 commits
Commits
Show all changes
23 commits
Select commit Hold shift + click to select a range
52aaead
refactor: use the definitions!
jamesmckinna Feb 10, 2025
e19936c
add: `LeftMonotonic` and `RightMonotonic`
jamesmckinna Feb 13, 2025
3ac9569
refactor: systematise `Congruence` reasoning
jamesmckinna Feb 13, 2025
a6710a0
fix: knock-on DRY
jamesmckinna Feb 13, 2025
02a2d5a
fix: `CHANGELOG` plus cosmetics
jamesmckinna Feb 13, 2025
d9c1136
refactor: change implicit to explicit
jamesmckinna Feb 14, 2025
c175970
fix: make `CHANGELOG` more detailed
jamesmckinna Feb 14, 2025
84c579c
refactor: rectify the names
jamesmckinna Feb 14, 2025
0ffa707
refactor: use the new definitions
jamesmckinna Feb 14, 2025
54b310d
refactor: use the new definitions some more
jamesmckinna Feb 14, 2025
f18f031
refactor: use the new definitions some more
jamesmckinna Feb 14, 2025
7bc97ce
refactor: unsolved metas make this a pain!
jamesmckinna Feb 14, 2025
0cb6671
refactor: use the new definitions
jamesmckinna Feb 14, 2025
17f626d
refactor: use the new definitions
jamesmckinna Feb 16, 2025
376c6da
refactor: use the new lemma statements; not yet their generic proofs
jamesmckinna Feb 16, 2025
87e0cb1
add: missing redefinitions
jamesmckinna Feb 17, 2025
c9e356e
refactor: use new lemma
jamesmckinna Feb 17, 2025
8528926
refactor: one more use of `Monotonic₁`
jamesmckinna Feb 18, 2025
9acf427
Merge branch 'master' into issue1579
jamesmckinna Feb 18, 2025
d3ed088
refactor: `Congruent` in terms of `Monotonic`
jamesmckinna Feb 19, 2025
7fff276
refactor: de-nest `module Congruence`
jamesmckinna Feb 19, 2025
49c0854
fix: tighten `import`s
jamesmckinna Feb 19, 2025
d1d0617
Merge branch 'master' into issue1579
jamesmckinna Feb 19, 2025
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
30 changes: 30 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -98,6 +98,20 @@ New modules
Additions to existing modules
-----------------------------

* In `Algebra.Consequences.Base`:
```agda
module Congruence (cong : Congruent₂ _≈_ _∙_) (refl : Reflexive _≈_) where
∙-congˡ : LeftCongruent _≈_ _∙_
∙-congʳ : RightCongruent _≈_ _∙_
```

* In `Algebra.Consequences.Setoid`:
```agda
module Congruence (cong : Congruent₂ _≈_ _∙_) where
∙-congˡ : LeftCongruent _≈_ _∙_
∙-congʳ : RightCongruent _≈_ _∙_
```

* In `Algebra.Construct.Pointwise`:
```agda
isNearSemiring : IsNearSemiring _≈_ _+_ _*_ 0# →
Expand Down Expand Up @@ -126,3 +140,19 @@ Additions to existing modules
quasiring : Quasiring c ℓ → Quasiring (a ⊔ c) (a ⊔ ℓ)
commutativeRing : CommutativeRing c ℓ → CommutativeRing (a ⊔ c) (a ⊔ ℓ)
```

* In `Relation.Binary.Consequences`:
```agda
mono₂⇒monoˡ : Reflexive ≤₁ → Monotonic₂ ≤₁ ≤₂ ≤₃ f → LeftMonotonic ≤₂ ≤₃ f
mono₂⇒monoˡ : Reflexive ≤₂ → Monotonic₂ ≤₁ ≤₂ ≤₃ f → RightMonotonic ≤₁ ≤₃ f
monoˡ∧monoʳ⇒mono₂ : Transitive ≤₃ →
LeftMonotonic ≤₂ ≤₃ f → RightMonotonic ≤₁ ≤₃ f →
Monotonic₂ ≤₁ ≤₂ ≤₃ f
```

* In `Relation.Binary.Definitions`:
```agda
LeftMonotonic : Rel B ℓ₁ → Rel C ℓ₂ → (A → B → C) → Set _
RightMonotonic : Rel A ℓ₁ → Rel C ℓ₂ → (A → B → C) → Set _
```

16 changes: 13 additions & 3 deletions src/Algebra/Consequences/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -13,14 +13,24 @@ module Algebra.Consequences.Base
open import Algebra.Core
open import Algebra.Definitions
open import Data.Sum.Base
open import Relation.Binary.Core
open import Relation.Binary.Consequences
using (mono₂⇒monoˡ; mono₂⇒monoʳ)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Definitions using (Reflexive)

module _ {ℓ} {__ : Op₂ A} (_≈_ : Rel A ℓ) where
module _ {ℓ} {__ : Op₂ A} (_≈_ : Rel A ℓ) where

sel⇒idem : Selective _≈_ __ → Idempotent _≈_ __
sel⇒idem : Selective _≈_ __ → Idempotent _≈_ __
sel⇒idem sel x = reduce (sel x x)

module Congruence (cong : Congruent₂ _≈_ _∙_) (refl : Reflexive _≈_) where

∙-congˡ : LeftCongruent _≈_ _∙_
∙-congˡ = mono₂⇒monoˡ {≤₁ = _≈_} {≤₂ = _≈_} {≤₃ = _≈_} refl cong _

∙-congʳ : RightCongruent _≈_ _∙_
∙-congʳ = mono₂⇒monoʳ {≤₁ = _≈_} {≤₂ = _≈_} {≤₃ = _≈_} refl cong _

module _ {ℓ} {f : Op₁ A} (_≈_ : Rel A ℓ) where

reflexive∧selfInverse⇒involutive : Reflexive _≈_ →
Expand Down
87 changes: 54 additions & 33 deletions src/Algebra/Consequences/Setoid.agda
Original file line number Diff line number Diff line change
Expand Up @@ -14,54 +14,65 @@ open import Relation.Binary.Definitions

module Algebra.Consequences.Setoid {a ℓ} (S : Setoid a ℓ) where

open Setoid S renaming (Carrier to A)
open import Algebra.Core
open import Algebra.Definitions _≈_
import Algebra.Consequences.Base as Base
open import Data.Sum.Base using (inj₁; inj₂)
open import Data.Product.Base using (_,_)
open import Function.Base using (_$_; id; _∘_)
open import Function.Definitions
import Relation.Binary.Consequences as Bin
open import Relation.Binary.Reasoning.Setoid S
open import Relation.Unary using (Pred)

open Setoid S renaming (Carrier to A)
open import Algebra.Definitions _≈_
open import Relation.Binary.Reasoning.Setoid S

------------------------------------------------------------------------
-- Re-exports

-- Export base lemmas that don't require the setoid

open import Algebra.Consequences.Base public
open Base public
hiding (module Congruence)

-- Export congruence lemmas using reflexivity

module Congruence {_∙_ : Op₂ A} (cong : Congruent₂ _∙_) where

open Base.Congruence _≈_ cong refl public

------------------------------------------------------------------------
-- MiddleFourExchange

module _ {_∙_ : Op₂ A} (cong : Congruent₂ _∙_) where

open Congruence cong

comm∧assoc⇒middleFour : Commutative _∙_ → Associative _∙_ →
_∙_ MiddleFourExchange _∙_
comm∧assoc⇒middleFour comm assoc w x y z = begin
(w ∙ x) ∙ (y ∙ z) ≈⟨ assoc w x (y ∙ z) ⟩
w ∙ (x ∙ (y ∙ z)) ≈⟨ cong refl (sym (assoc x y z)) ⟩
w ∙ ((x ∙ y) ∙ z) ≈⟨ cong refl (cong (comm x y) refl) ⟩
w ∙ ((y ∙ x) ∙ z) ≈⟨ cong refl (assoc y x z) ⟩
w ∙ (y ∙ (x ∙ z)) ≈⟨ sym (assoc w y (x ∙ z)) ⟩
w ∙ (x ∙ (y ∙ z)) ≈⟨ ∙-congˡ (assoc x y z)
w ∙ ((x ∙ y) ∙ z) ≈⟨ ∙-congˡ (∙-congʳ (comm x y)) ⟩
w ∙ ((y ∙ x) ∙ z) ≈⟨ ∙-congˡ (assoc y x z) ⟩
w ∙ (y ∙ (x ∙ z)) ≈⟨ assoc w y (x ∙ z)
(w ∙ y) ∙ (x ∙ z) ∎

identity∧middleFour⇒assoc : {e : A} → Identity e _∙_ →
_∙_ MiddleFourExchange _∙_ →
Associative _∙_
identity∧middleFour⇒assoc {e} (identityˡ , identityʳ) middleFour x y z = begin
(x ∙ y) ∙ z ≈⟨ cong refl (sym (identityˡ z)) ⟩
(x ∙ y) ∙ z ≈⟨ ∙-congˡ (identityˡ z)
(x ∙ y) ∙ (e ∙ z) ≈⟨ middleFour x y e z ⟩
(x ∙ e) ∙ (y ∙ z) ≈⟨ cong (identityʳ x) refl
(x ∙ e) ∙ (y ∙ z) ≈⟨ ∙-congʳ (identityʳ x) ⟩
x ∙ (y ∙ z) ∎

identity∧middleFour⇒comm : {_+_ : Op₂ A} {e : A} → Identity e _+_ →
_∙_ MiddleFourExchange _+_ →
Commutative _∙_
identity∧middleFour⇒comm {_+_} {e} (identityˡ , identityʳ) middleFour x y
= begin
x ∙ y ≈⟨ sym (cong (identityˡ x) (identityʳ y)) ⟩
x ∙ y ≈⟨ cong (identityˡ x) (identityʳ y)
(e + x) ∙ (y + e) ≈⟨ middleFour e x y e ⟩
(e + y) ∙ (x + e) ≈⟨ cong (identityˡ y) (identityʳ x) ⟩
y ∙ x ∎
Expand Down Expand Up @@ -198,25 +209,27 @@ module _ {_∙_ : Op₂ A} {_⁻¹ : Op₁ A} {e} (comm : Commutative _∙_) whe

module _ {_∙_ : Op₂ A} {_⁻¹ : Op₁ A} {e} (cong : Congruent₂ _∙_) where

open Congruence cong

assoc∧id∧invʳ⇒invˡ-unique : Associative _∙_ →
Identity e _∙_ → RightInverse e _⁻¹ _∙_ →
∀ x y → (x ∙ y) ≈ e → x ≈ (y ⁻¹)
assoc∧id∧invʳ⇒invˡ-unique assoc (idˡ , idʳ) invʳ x y eq = begin
x ≈⟨ sym (idʳ x) ⟩
x ∙ e ≈⟨ cong refl (sym (invʳ y)) ⟩
x ∙ (y ∙ (y ⁻¹)) ≈⟨ sym (assoc x y (y ⁻¹)) ⟩
(x ∙ y) ∙ (y ⁻¹) ≈⟨ cong eq refl
x ≈⟨ idʳ x
x ∙ e ≈⟨ ∙-congˡ (invʳ y)
x ∙ (y ∙ (y ⁻¹)) ≈⟨ assoc x y (y ⁻¹)
(x ∙ y) ∙ (y ⁻¹) ≈⟨ ∙-congʳ eq ⟩
e ∙ (y ⁻¹) ≈⟨ idˡ (y ⁻¹) ⟩
y ⁻¹ ∎

assoc∧id∧invˡ⇒invʳ-unique : Associative _∙_ →
Identity e _∙_ → LeftInverse e _⁻¹ _∙_ →
∀ x y → (x ∙ y) ≈ e → y ≈ (x ⁻¹)
assoc∧id∧invˡ⇒invʳ-unique assoc (idˡ , idʳ) invˡ x y eq = begin
y ≈⟨ sym (idˡ y) ⟩
e ∙ y ≈⟨ cong (sym (invˡ x)) refl ⟩
y ≈⟨ idˡ y
e ∙ y ≈⟨ ∙-congʳ (invˡ x)
((x ⁻¹) ∙ x) ∙ y ≈⟨ assoc (x ⁻¹) x y ⟩
(x ⁻¹) ∙ (x ∙ y) ≈⟨ cong refl eq ⟩
(x ⁻¹) ∙ (x ∙ y) ≈⟨ ∙-congˡ eq ⟩
(x ⁻¹) ∙ e ≈⟨ idʳ (x ⁻¹) ⟩
x ⁻¹ ∎

Expand All @@ -228,6 +241,8 @@ module _ {_∙_ _◦_ : Op₂ A}
(∙-comm : Commutative _∙_)
where

open Congruence ◦-cong renaming (∙-congˡ to ◦-congˡ)

comm∧distrˡ⇒distrʳ : _∙_ DistributesOverˡ _◦_ → _∙_ DistributesOverʳ _◦_
comm∧distrˡ⇒distrʳ distrˡ x y z = begin
(y ◦ z) ∙ x ≈⟨ ∙-comm (y ◦ z) x ⟩
Expand All @@ -250,7 +265,7 @@ module _ {_∙_ _◦_ : Op₂ A}

comm⇒sym[distribˡ] : ∀ x → Symmetric (λ y z → (x ◦ (y ∙ z)) ≈ ((x ◦ y) ∙ (x ◦ z)))
comm⇒sym[distribˡ] x {y} {z} prf = begin
x ◦ (z ∙ y) ≈⟨ ◦-cong refl (∙-comm z y) ⟩
x ◦ (z ∙ y) ≈⟨ ◦-congˡ (∙-comm z y) ⟩
x ◦ (y ∙ z) ≈⟨ prf ⟩
(x ◦ y) ∙ (x ◦ z) ≈⟨ ∙-comm (x ◦ y) (x ◦ z) ⟩
(x ◦ z) ∙ (x ◦ y) ∎
Expand All @@ -262,16 +277,18 @@ module _ {_∙_ _◦_ : Op₂ A}
(◦-comm : Commutative _◦_)
where

open Congruence ∙-cong

distrib∧absorbs⇒distribˡ : _∙_ Absorbs _◦_ →
_◦_ Absorbs _∙_ →
_◦_ DistributesOver _∙_ →
_∙_ DistributesOverˡ _◦_
distrib∧absorbs⇒distribˡ ∙-absorbs-◦ ◦-absorbs-∙ (◦-distribˡ-∙ , ◦-distribʳ-∙) x y z = begin
x ∙ (y ◦ z) ≈⟨ ∙-cong (∙-absorbs-◦ _ _) refl
(x ∙ (x ◦ y)) ∙ (y ◦ z) ≈⟨ ∙-cong (∙-cong refl (◦-comm _ _)) refl
x ∙ (y ◦ z) ≈⟨ ∙-congʳ (∙-absorbs-◦ _ _) ⟨
(x ∙ (x ◦ y)) ∙ (y ◦ z) ≈⟨ ∙-congʳ (∙-congˡ (◦-comm _ _)) ⟩
(x ∙ (y ◦ x)) ∙ (y ◦ z) ≈⟨ ∙-assoc _ _ _ ⟩
x ∙ ((y ◦ x) ∙ (y ◦ z)) ≈⟨ ∙-cong refl (◦-distribˡ-∙ _ _ _) ⟨
x ∙ (y ◦ (x ∙ z)) ≈⟨ ∙-cong (◦-absorbs-∙ _ _) refl
x ∙ ((y ◦ x) ∙ (y ◦ z)) ≈⟨ ∙-congˡ (◦-distribˡ-∙ _ _ _) ⟨
x ∙ (y ◦ (x ∙ z)) ≈⟨ ∙-congʳ (◦-absorbs-∙ _ _) ⟨
(x ◦ (x ∙ z)) ∙ (y ◦ (x ∙ z)) ≈⟨ ◦-distribʳ-∙ _ _ _ ⟨
(x ∙ y) ◦ (x ∙ z) ∎

Expand All @@ -284,27 +301,31 @@ module _ {_+_ _*_ : Op₂ A}
(*-cong : Congruent₂ _*_)
where

open Congruence +-cong renaming (∙-congˡ to +-congˡ; ∙-congʳ to +-congʳ)

open Congruence *-cong renaming (∙-congˡ to *-congˡ; ∙-congʳ to *-congʳ)

assoc∧distribʳ∧idʳ∧invʳ⇒zeˡ : Associative _+_ → _*_ DistributesOverʳ _+_ →
RightIdentity 0# _+_ → RightInverse 0# _⁻¹ _+_ →
LeftZero 0# _*_
assoc∧distribʳ∧idʳ∧invʳ⇒zeˡ +-assoc distribʳ idʳ invʳ x = begin
0# * x ≈⟨ sym (idʳ _) ⟩
(0# * x) + 0# ≈⟨ +-cong refl (sym (invʳ _)) ⟩
(0# * x) + ((0# * x) + ((0# * x)⁻¹)) ≈⟨ sym (+-assoc _ _ _) ⟩
((0# * x) + (0# * x)) + ((0# * x)⁻¹) ≈⟨ +-cong (sym (distribʳ _ _ _)) refl ⟩
((0# + 0#) * x) + ((0# * x)⁻¹) ≈⟨ +-cong (*-cong (idʳ _) refl) refl
0# * x ≈⟨ idʳ _
(0# * x) + 0# ≈⟨ +-congˡ (invʳ _)
(0# * x) + ((0# * x) + ((0# * x)⁻¹)) ≈⟨ +-assoc _ _ _
((0# * x) + (0# * x)) + ((0# * x)⁻¹) ≈⟨ +-congʳ (distribʳ _ _ _)
((0# + 0#) * x) + ((0# * x)⁻¹) ≈⟨ +-congʳ (*-congʳ (idʳ _))
(0# * x) + ((0# * x)⁻¹) ≈⟨ invʳ _ ⟩
0# ∎

assoc∧distribˡ∧idʳ∧invʳ⇒zeʳ : Associative _+_ → _*_ DistributesOverˡ _+_ →
RightIdentity 0# _+_ → RightInverse 0# _⁻¹ _+_ →
RightZero 0# _*_
assoc∧distribˡ∧idʳ∧invʳ⇒zeʳ +-assoc distribˡ idʳ invʳ x = begin
x * 0# ≈⟨ sym (idʳ _) ⟩
(x * 0#) + 0# ≈⟨ +-cong refl (sym (invʳ _)) ⟩
(x * 0#) + ((x * 0#) + ((x * 0#)⁻¹)) ≈⟨ sym (+-assoc _ _ _) ⟩
((x * 0#) + (x * 0#)) + ((x * 0#)⁻¹) ≈⟨ +-cong (sym (distribˡ _ _ _)) refl ⟩
(x * (0# + 0#)) + ((x * 0#)⁻¹) ≈⟨ +-cong (*-cong refl (idʳ _)) refl
x * 0# ≈⟨ idʳ _
(x * 0#) + 0# ≈⟨ +-congˡ (invʳ _)
(x * 0#) + ((x * 0#) + ((x * 0#)⁻¹)) ≈⟨ +-assoc _ _ _
((x * 0#) + (x * 0#)) + ((x * 0#)⁻¹) ≈⟨ +-congʳ (distribˡ _ _ _)
(x * (0# + 0#)) + ((x * 0#)⁻¹) ≈⟨ +-congʳ (*-congˡ (idʳ _)) ⟩
((x * 0#) + ((x * 0#)⁻¹)) ≈⟨ invʳ _ ⟩
0# ∎

Expand Down
4 changes: 2 additions & 2 deletions src/Algebra/Definitions.agda
Original file line number Diff line number Diff line change
Expand Up @@ -37,10 +37,10 @@ Congruent₂ : Op₂ A → Set _
Congruent₂ ∙ = ∙ Preserves₂ _≈_ ⟶ _≈_ ⟶ _≈_

LeftCongruent : Op₂ A → Set _
LeftCongruent _∙_ = ∀ {x} → (x ∙_) Preserves _≈_ ⟶ _≈_
LeftCongruent _∙_ = ∀ {x} → Congruent₁ (x ∙_)

RightCongruent : Op₂ A → Set _
RightCongruent _∙_ = ∀ {x} → (_∙ x) Preserves _≈_ ⟶ _≈_
RightCongruent _∙_ = ∀ {x} → Congruent₁ (_∙ x)

Associative : Op₂ A → Set _
Associative _∙_ = ∀ x y z → ((x ∙ y) ∙ z) ≈ (x ∙ (y ∙ z))
Expand Down
5 changes: 3 additions & 2 deletions src/Algebra/Structures.agda
Original file line number Diff line number Diff line change
Expand Up @@ -54,13 +54,14 @@ record IsMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) where

setoid : Setoid a ℓ
setoid = record { isEquivalence = isEquivalence }

{-
∙-congˡ : LeftCongruent ∙
∙-congˡ y≈z = ∙-cong refl y≈z

∙-congʳ : RightCongruent ∙
∙-congʳ y≈z = ∙-cong y≈z refl

-}
open Consequences.Congruence setoid ∙-cong public

record IsCommutativeMagma (∙ : Op₂ A) : Set (a ⊔ ℓ) where
field
Expand Down
2 changes: 1 addition & 1 deletion src/Data/Fin/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -902,7 +902,7 @@ pinch-surjective _ zero = zero , λ { refl → refl }
pinch-surjective zero (suc j) = suc (suc j) , λ { refl → refl }
pinch-surjective (suc i) (suc j) = map suc (λ {f refl → cong suc (f refl)}) (pinch-surjective i j)

pinch-mono-≤ : ∀ (i : Fin n) → (pinch i) Preserves _≤_ _≤_
pinch-mono-≤ : ∀ (i : Fin n) → Monotonic₁ _≤_ _≤_ (pinch i)
pinch-mono-≤ 0F {0F} {k} 0≤n = z≤n
pinch-mono-≤ 0F {suc j} {suc k} j≤k = ℕ.s≤s⁻¹ j≤k
pinch-mono-≤ (suc i) {0F} {k} 0≤n = z≤n
Expand Down
Loading