6
6
7
7
{-# OPTIONS --cubical-compatible --safe #-}
8
8
9
- open import Algebra
9
+ open import Algebra.Bundles using (Semiring)
10
10
open import Data.Nat.Base as ℕ using (zero; suc)
11
11
12
12
module Algebra.Properties.Semiring.Mult
13
13
{a ℓ} (S : Semiring a ℓ) where
14
14
15
15
open Semiring S renaming (zero to *-zero)
16
16
open import Relation.Binary.Reasoning.Setoid setoid
17
+ open import Algebra.Definitions _≈_ using (_IdempotentOn_)
17
18
18
19
------------------------------------------------------------------------
19
20
-- Re-export definition from the monoid
@@ -23,21 +24,15 @@ open import Algebra.Properties.Monoid.Mult +-monoid public
23
24
------------------------------------------------------------------------
24
25
-- Properties of _×_
25
26
26
- -- (_× 1# ) is homomorphic with respect to _ℕ.*_/_*_.
27
+ -- (0 ×_ ) is (0# *_)
27
28
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)
36
31
37
- -- (1 ×_) is the identity
32
+ -- (1 ×_) is (1# *_)
38
33
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))
41
36
42
37
-- (n ×_) commutes with _*_
43
38
@@ -60,3 +55,32 @@ open import Algebra.Properties.Monoid.Mult +-monoid public
60
55
x * y + (n × x) * y ≈⟨ +-congˡ (×-assoc-* n _ _) ⟩
61
56
x * y + n × (x * y) ≡⟨⟩
62
57
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