Skip to content

Commit 709da18

Browse files
jamesmckinnaandreasabel
authored andcommitted
lemmas about semiring structure induced by _× x (#2272)
* tidying up proofs of, and using, semiring structure induced by `_× x` * reinstated lemma about `0#` * fixed name clash * added companion lemma for `1#` * new lemma, plus refactoring * removed anonymous `module` * don't inline use of the lemma * revised deprecation warning
1 parent 7c7c030 commit 709da18

File tree

4 files changed

+62
-14
lines changed

4 files changed

+62
-14
lines changed

CHANGELOG.md

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

36+
* In `Algebra.Properties.Semiring.Mult`:
37+
```agda
38+
1×-identityʳ ↦ ×-homo-1
39+
```
40+
3641
* In `Data.Nat.Divisibility.Core`:
3742
```agda
3843
*-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣
@@ -90,6 +95,19 @@ Additions to existing modules
9095
rawModule : RawModule R c ℓ
9196
```
9297

98+
* In `Algebra.Properties.Monoid.Mult`:
99+
```agda
100+
×-homo-0 : ∀ x → 0 × x ≈ 0#
101+
×-homo-1 : ∀ x → 1 × x ≈ x
102+
```
103+
104+
* In `Algebra.Properties.Semiring.Mult`:
105+
```agda
106+
×-homo-0# : ∀ x → 0 × x ≈ 0# * x
107+
×-homo-1# : ∀ x → 1 × x ≈ 1# * x
108+
idem-×-homo-* : (_*_ IdempotentOn x) → (m × x) * (n × x) ≈ (m ℕ.* n) × x
109+
```
110+
93111
* In `Data.Fin.Properties`:
94112
```agda
95113
nonZeroIndex : Fin n → ℕ.NonZero n

src/Algebra/Properties/Monoid/Mult.agda

+6
Original file line numberDiff line numberDiff line change
@@ -51,6 +51,12 @@ open import Algebra.Definitions.RawMonoid rawMonoid public
5151

5252
-- _×_ is homomorphic with respect to _ℕ+_/_+_.
5353

54+
×-homo-0 : x 0 × x ≈ 0#
55+
×-homo-0 x = refl
56+
57+
×-homo-1 : x 1 × x ≈ x
58+
×-homo-1 = +-identityʳ
59+
5460
×-homo-+ : x m n (m ℕ.+ n) × x ≈ m × x + n × x
5561
×-homo-+ x 0 n = sym (+-identityˡ (n × x))
5662
×-homo-+ x (suc m) n = begin

src/Algebra/Properties/Semiring/Binomial.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -216,7 +216,7 @@ term₁+term₂≈term x*y≈y*x n (suc i) with view i
216216
theorem : x * y ≈ y * x n (x + y) ^ n ≈ binomialExpansion n
217217
theorem x*y≈y*x zero = begin
218218
(x + y) ^ 0 ≡⟨⟩
219-
1# ≈⟨ 1×-identityʳ 1# ⟨
219+
1# ≈⟨ ×-homo-1 1# ⟨
220220
1 × 1# ≈⟨ *-identityʳ (1 × 1#) ⟨
221221
(1 × 1#) * 1# ≈⟨ ×-assoc-* 1 1# 1# ⟩
222222
1 × (1# * 1#) ≡⟨⟩

src/Algebra/Properties/Semiring/Mult.agda

+37-13
Original file line numberDiff line numberDiff line change
@@ -6,14 +6,15 @@
66

77
{-# OPTIONS --cubical-compatible --safe #-}
88

9-
open import Algebra
9+
open import Algebra.Bundles using (Semiring)
1010
open import Data.Nat.Base as ℕ using (zero; suc)
1111

1212
module Algebra.Properties.Semiring.Mult
1313
{a ℓ} (S : Semiring a ℓ) where
1414

1515
open Semiring S renaming (zero to *-zero)
1616
open import Relation.Binary.Reasoning.Setoid setoid
17+
open import Algebra.Definitions _≈_ using (_IdempotentOn_)
1718

1819
------------------------------------------------------------------------
1920
-- Re-export definition from the monoid
@@ -23,21 +24,15 @@ open import Algebra.Properties.Monoid.Mult +-monoid public
2324
------------------------------------------------------------------------
2425
-- Properties of _×_
2526

26-
-- (_× 1#) is homomorphic with respect to _ℕ.*_/_*_.
27+
-- (0 ×_) is (0# *_)
2728

28-
×1-homo-* : m n (m ℕ.* n) × 1# ≈ (m × 1#) * (n × 1#)
29-
×1-homo-* 0 n = sym (zeroˡ (n × 1#))
30-
×1-homo-* (suc m) n = begin
31-
(n ℕ.+ m ℕ.* n) × 1# ≈⟨ ×-homo-+ 1# n (m ℕ.* n) ⟩
32-
n × 1# + (m ℕ.* n) × 1# ≈⟨ +-congˡ (×1-homo-* m n) ⟩
33-
n × 1# + (m × 1#) * (n × 1#) ≈⟨ +-congʳ (*-identityˡ _) ⟨
34-
1# * (n × 1#) + (m × 1#) * (n × 1#) ≈⟨ distribʳ (n × 1#) 1# (m × 1#) ⟨
35-
(1# + m × 1#) * (n × 1#) ∎
29+
×-homo-0# : x 0 × x ≈ 0# * x
30+
×-homo-0# x = sym (zeroˡ x)
3631

37-
-- (1 ×_) is the identity
32+
-- (1 ×_) is (1# *_)
3833

39-
1×-identityʳ : x 1 × x ≈ x
40-
1×-identityʳ = +-identityʳ
34+
×-homo-1# : x 1 × x ≈ 1# * x
35+
×-homo-1# x = trans (×-homo-1 x) (sym (*-identityˡ x))
4136

4237
-- (n ×_) commutes with _*_
4338

@@ -60,3 +55,32 @@ open import Algebra.Properties.Monoid.Mult +-monoid public
6055
x * y + (n × x) * y ≈⟨ +-congˡ (×-assoc-* n _ _) ⟩
6156
x * y + n × (x * y) ≡⟨⟩
6257
suc n × (x * y) ∎
58+
59+
-- (_× x) is homomorphic with respect to _ℕ.*_/_*_ for idempotent x.
60+
61+
idem-×-homo-* : m n {x} (_*_ IdempotentOn x) (m × x) * (n × x) ≈ (m ℕ.* n) × x
62+
idem-×-homo-* m n {x} idem = begin
63+
(m × x) * (n × x) ≈⟨ ×-assoc-* m x (n × x) ⟩
64+
m × (x * (n × x)) ≈⟨ ×-congʳ m (×-comm-* n x x) ⟩
65+
m × (n × (x * x)) ≈⟨ ×-assocˡ _ m n ⟩
66+
(m ℕ.* n) × (x * x) ≈⟨ ×-congʳ (m ℕ.* n) idem ⟩
67+
(m ℕ.* n) × x ∎
68+
69+
-- (_× 1#) is homomorphic with respect to _ℕ.*_/_*_.
70+
71+
×1-homo-* : m n (m ℕ.* n) × 1# ≈ (m × 1#) * (n × 1#)
72+
×1-homo-* m n = sym (idem-×-homo-* m n (*-identityʳ 1#))
73+
74+
------------------------------------------------------------------------
75+
-- DEPRECATED NAMES
76+
------------------------------------------------------------------------
77+
-- Please use the new names as continuing support for the old names is
78+
-- not guaranteed.
79+
80+
-- Version 2.1
81+
82+
1×-identityʳ = ×-homo-1
83+
{-# WARNING_ON_USAGE 1×-identityʳ
84+
"Warning: 1×-identityʳ was deprecated in v2.1.
85+
Please use ×-homo-1 instead. "
86+
#-}

0 commit comments

Comments
 (0)