-
Notifications
You must be signed in to change notification settings - Fork 246
/
Copy pathMult.agda
94 lines (72 loc) · 3.12 KB
/
Mult.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
------------------------------------------------------------------------
-- The Agda standard library
--
-- Multiplication over a monoid (i.e. repeated addition)
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
open import Algebra.Bundles using (Monoid)
open import Data.Bool.Base as Bool using (Bool; true; false; if_then_else_)
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; NonZero)
open import Relation.Binary.Core using (_Preserves_⟶_; _Preserves₂_⟶_⟶_)
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
module Algebra.Properties.Monoid.Mult {a ℓ} (M : Monoid a ℓ) where
-- View of the monoid operator as addition
open Monoid M
renaming
( _∙_ to _+_
; ∙-cong to +-cong
; ∙-congʳ to +-congʳ
; ∙-congˡ to +-congˡ
; identityˡ to +-identityˡ
; identityʳ to +-identityʳ
; assoc to +-assoc
; ε to 0#
)
open import Relation.Binary.Reasoning.Setoid setoid
open import Algebra.Definitions _≈_
------------------------------------------------------------------------
-- Definition
open import Algebra.Definitions.RawMonoid rawMonoid public
using (_×_; _∧_; _∧′_∙_)
------------------------------------------------------------------------
-- Properties of _×_
×-congʳ : ∀ n → (n ×_) Preserves _≈_ ⟶ _≈_
×-congʳ 0 x≈x′ = refl
×-congʳ (suc n) x≈x′ = +-cong x≈x′ (×-congʳ n x≈x′)
×-cong : _×_ Preserves₂ _≡_ ⟶ _≈_ ⟶ _≈_
×-cong {n} ≡.refl x≈x′ = ×-congʳ n x≈x′
×-congˡ : ∀ {x} → (_× x) Preserves _≡_ ⟶ _≈_
×-congˡ m≡n = ×-cong m≡n refl
-- _×_ is homomorphic with respect to _ℕ+_/_+_.
×-homo-0 : ∀ x → 0 × x ≈ 0#
×-homo-0 x = refl
×-homo-1 : ∀ x → 1 × x ≈ x
×-homo-1 = +-identityʳ
×-homo-+ : ∀ x m n → (m ℕ.+ n) × x ≈ m × x + n × x
×-homo-+ x 0 n = sym (+-identityˡ (n × x))
×-homo-+ x (suc m) n = begin
x + (m ℕ.+ n) × x ≈⟨ +-congˡ (×-homo-+ x m n) ⟩
x + (m × x + n × x) ≈⟨ sym (+-assoc x (m × x) (n × x)) ⟩
x + m × x + n × x ∎
×-idem : ∀ {c} → _+_ IdempotentOn c →
∀ n → .{{_ : NonZero n}} → n × c ≈ c
×-idem {c} idem (suc zero) = +-identityʳ c
×-idem {c} idem (suc n@(suc _)) = begin
c + (n × c) ≈⟨ +-congˡ (×-idem idem n ) ⟩
c + c ≈⟨ idem ⟩
c ∎
×-assocˡ : ∀ x m n → m × (n × x) ≈ (m ℕ.* n) × x
×-assocˡ x zero n = refl
×-assocˡ x (suc m) n = begin
n × x + m × n × x ≈⟨ +-congˡ (×-assocˡ x m n) ⟩
n × x + (m ℕ.* n) × x ≈⟨ ×-homo-+ x n (m ℕ.* n) ⟨
(suc m ℕ.* n) × x ∎
-- _∧_ is homomorphic with respect to Bool._∧_.
∧-homo-true : ∀ x → true ∧ x ≈ x
∧-homo-true _ = refl
∧-assocˡ : ∀ b b′ x → b ∧ (b′ ∧ x) ≈ (b Bool.∧ b′) ∧ x
∧-assocˡ false _ _ = refl
∧-assocˡ true _ _ = refl
b∧x∙y≈b∧x+y : ∀ b x y → b ∧′ x ∙ y ≈ (b ∧ x) + y
b∧x∙y≈b∧x+y true _ _ = refl
b∧x∙y≈b∧x+y false _ y = sym (+-identityˡ y)