File tree Expand file tree Collapse file tree 2 files changed +43
-0
lines changed
Expand file tree Collapse file tree 2 files changed +43
-0
lines changed Original file line number Diff line number Diff line change @@ -24,5 +24,10 @@ Deprecated names
2424New modules
2525-----------
2626
27+ * Properties of ` IdempotentCommutativeMonoid ` s refactored out from ` Algebra.Solver.IdempotentCommutativeMonoid ` :
28+ ``` agda
29+ Algebra.Properties.IdempotentCommutativeMonoid
30+ ```
31+
2732Additions to existing modules
2833-----------------------------
Original file line number Diff line number Diff line change 1+ ------------------------------------------------------------------------
2+ -- The Agda standard library
3+ --
4+ -- Some derivable properties
5+ ------------------------------------------------------------------------
6+
7+ {-# OPTIONS --cubical-compatible --safe #-}
8+
9+ open import Algebra.Bundles using (IdempotentCommutativeMonoid)
10+
11+ module Algebra.Properties.IdempotentCommutativeMonoid
12+ {c ℓ} (M : IdempotentCommutativeMonoid c ℓ) where
13+
14+ open IdempotentCommutativeMonoid M
15+
16+ open import Algebra.Consequences.Setoid setoid
17+ using (comm∧distrˡ⇒distrʳ; comm∧distrˡ⇒distr)
18+ open import Algebra.Definitions _≈_
19+ using (_DistributesOverˡ_; _DistributesOverʳ_; _DistributesOver_)
20+ open import Algebra.Properties.CommutativeSemigroup commutativeSemigroup
21+ using (interchange)
22+ open import Relation.Binary.Reasoning.Setoid setoid
23+
24+
25+ ------------------------------------------------------------------------
26+ -- Distributivity
27+
28+ ∙-distrˡ-∙ : _∙_ DistributesOverˡ _∙_
29+ ∙-distrˡ-∙ a b c = begin
30+ a ∙ (b ∙ c) ≈⟨ ∙-congʳ (idem a) ⟨
31+ (a ∙ a) ∙ (b ∙ c) ≈⟨ interchange _ _ _ _ ⟩
32+ (a ∙ b) ∙ (a ∙ c) ∎
33+
34+ ∙-distrʳ-∙ : _∙_ DistributesOverʳ _∙_
35+ ∙-distrʳ-∙ = comm∧distrˡ⇒distrʳ ∙-cong comm ∙-distrˡ-∙
36+
37+ ∙-distr-∙ : _∙_ DistributesOver _∙_
38+ ∙-distr-∙ = comm∧distrˡ⇒distr ∙-cong comm ∙-distrˡ-∙
You can’t perform that action at this time.
0 commit comments