forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathRingMonomorphism.agda
161 lines (135 loc) · 7.02 KB
/
RingMonomorphism.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
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
------------------------------------------------------------------------
-- The Agda standard library
--
-- Consequences of a monomorphism between ring-like structures
------------------------------------------------------------------------
-- See Data.Nat.Binary.Properties for examples of how this and similar
-- modules can be used to easily translate properties between types.
{-# OPTIONS --cubical-compatible --safe #-}
open import Algebra.Bundles
open import Algebra.Morphism.Structures
import Algebra.Morphism.GroupMonomorphism as GroupMonomorphism
import Algebra.Morphism.MonoidMonomorphism as MonoidMonomorphism
open import Relation.Binary.Core
module Algebra.Morphism.RingMonomorphism
{a b ℓ₁ ℓ₂} {R₁ : RawRing a ℓ₁} {R₂ : RawRing b ℓ₂} {⟦_⟧}
(isRingMonomorphism : IsRingMonomorphism R₁ R₂ ⟦_⟧)
where
open IsRingMonomorphism isRingMonomorphism
open RawRing R₁ renaming (Carrier to A; _≈_ to _≈₁_)
open RawRing R₂ renaming
( Carrier to B; _≈_ to _≈₂_; _+_ to _⊕_
; _*_ to _⊛_; 1# to 1#₂; 0# to 0#₂; -_ to ⊝_)
open import Algebra.Definitions
open import Algebra.Structures
open import Data.Product.Base using (proj₁; proj₂; _,_)
import Relation.Binary.Reasoning.Setoid as SetoidReasoning
------------------------------------------------------------------------
-- Re-export all properties of group and monoid monomorphisms
open GroupMonomorphism +-isGroupMonomorphism public
renaming
( assoc to +-assoc
; comm to +-comm
; cong to +-cong
; idem to +-idem
; sel to +-sel
; ⁻¹-cong to neg-cong
; identity to +-identity; identityˡ to +-identityˡ; identityʳ to +-identityʳ
; cancel to +-cancel; cancelˡ to +-cancelˡ; cancelʳ to +-cancelʳ
; zero to +-zero; zeroˡ to +-zeroˡ; zeroʳ to +-zeroʳ
; isMagma to +-isMagma
; isSemigroup to +-isSemigroup
; isMonoid to +-isMonoid
; isSelectiveMagma to +-isSelectiveMagma
; isBand to +-isBand
; isCommutativeMonoid to +-isCommutativeMonoid
)
open MonoidMonomorphism *-isMonoidMonomorphism public
renaming
( assoc to *-assoc
; comm to *-comm
; cong to *-cong
; idem to *-idem
; sel to *-sel
; identity to *-identity; identityˡ to *-identityˡ; identityʳ to *-identityʳ
; cancel to *-cancel; cancelˡ to *-cancelˡ; cancelʳ to *-cancelʳ
; zero to *-zero; zeroˡ to *-zeroˡ; zeroʳ to *-zeroʳ
; isMagma to *-isMagma
; isSemigroup to *-isSemigroup
; isMonoid to *-isMonoid
; isSelectiveMagma to *-isSelectiveMagma
; isBand to *-isBand
; isCommutativeMonoid to *-isCommutativeMonoid
)
------------------------------------------------------------------------
-- Properties
module _ (+-isGroup : IsGroup _≈₂_ _⊕_ 0#₂ ⊝_)
(*-isMagma : IsMagma _≈₂_ _⊛_) where
open IsGroup +-isGroup hiding (setoid; refl; sym)
open IsMagma *-isMagma renaming (∙-cong to ◦-cong)
open SetoidReasoning setoid
distribˡ : _DistributesOverˡ_ _≈₂_ _⊛_ _⊕_ → _DistributesOverˡ_ _≈₁_ _*_ _+_
distribˡ distribˡ x y z = injective (begin
⟦ x * (y + z) ⟧ ≈⟨ *-homo x (y + z) ⟩
⟦ x ⟧ ⊛ ⟦ y + z ⟧ ≈⟨ ◦-cong refl (+-homo y z) ⟩
⟦ x ⟧ ⊛ (⟦ y ⟧ ⊕ ⟦ z ⟧) ≈⟨ distribˡ ⟦ x ⟧ ⟦ y ⟧ ⟦ z ⟧ ⟩
⟦ x ⟧ ⊛ ⟦ y ⟧ ⊕ ⟦ x ⟧ ⊛ ⟦ z ⟧ ≈˘⟨ ∙-cong (*-homo x y) (*-homo x z) ⟩
⟦ x * y ⟧ ⊕ ⟦ x * z ⟧ ≈˘⟨ +-homo (x * y) (x * z) ⟩
⟦ x * y + x * z ⟧ ∎)
distribʳ : _DistributesOverʳ_ _≈₂_ _⊛_ _⊕_ → _DistributesOverʳ_ _≈₁_ _*_ _+_
distribʳ distribˡ x y z = injective (begin
⟦ (y + z) * x ⟧ ≈⟨ *-homo (y + z) x ⟩
⟦ y + z ⟧ ⊛ ⟦ x ⟧ ≈⟨ ◦-cong (+-homo y z) refl ⟩
(⟦ y ⟧ ⊕ ⟦ z ⟧) ⊛ ⟦ x ⟧ ≈⟨ distribˡ ⟦ x ⟧ ⟦ y ⟧ ⟦ z ⟧ ⟩
⟦ y ⟧ ⊛ ⟦ x ⟧ ⊕ ⟦ z ⟧ ⊛ ⟦ x ⟧ ≈˘⟨ ∙-cong (*-homo y x) (*-homo z x) ⟩
⟦ y * x ⟧ ⊕ ⟦ z * x ⟧ ≈˘⟨ +-homo (y * x) (z * x) ⟩
⟦ y * x + z * x ⟧ ∎)
distrib : _DistributesOver_ _≈₂_ _⊛_ _⊕_ → _DistributesOver_ _≈₁_ _*_ _+_
distrib distrib = distribˡ (proj₁ distrib) , distribʳ (proj₂ distrib)
zeroˡ : LeftZero _≈₂_ 0#₂ _⊛_ → LeftZero _≈₁_ 0# _*_
zeroˡ zeroˡ x = injective (begin
⟦ 0# * x ⟧ ≈⟨ *-homo 0# x ⟩
⟦ 0# ⟧ ⊛ ⟦ x ⟧ ≈⟨ ◦-cong 0#-homo refl ⟩
0#₂ ⊛ ⟦ x ⟧ ≈⟨ zeroˡ ⟦ x ⟧ ⟩
0#₂ ≈˘⟨ 0#-homo ⟩
⟦ 0# ⟧ ∎)
zeroʳ : RightZero _≈₂_ 0#₂ _⊛_ → RightZero _≈₁_ 0# _*_
zeroʳ zeroʳ x = injective (begin
⟦ x * 0# ⟧ ≈⟨ *-homo x 0# ⟩
⟦ x ⟧ ⊛ ⟦ 0# ⟧ ≈⟨ ◦-cong refl 0#-homo ⟩
⟦ x ⟧ ⊛ 0#₂ ≈⟨ zeroʳ ⟦ x ⟧ ⟩
0#₂ ≈˘⟨ 0#-homo ⟩
⟦ 0# ⟧ ∎)
zero : Zero _≈₂_ 0#₂ _⊛_ → Zero _≈₁_ 0# _*_
zero zero = zeroˡ (proj₁ zero) , zeroʳ (proj₂ zero)
neg-distribˡ-* : (∀ x y → (⊝ (x ⊛ y)) ≈₂ ((⊝ x) ⊛ y)) → (∀ x y → (- (x * y)) ≈₁ ((- x) * y))
neg-distribˡ-* neg-distribˡ-* x y = injective (begin
⟦ - (x * y) ⟧ ≈⟨ -‿homo (x * y) ⟩
⊝ ⟦ x * y ⟧ ≈⟨ ⁻¹-cong (*-homo x y) ⟩
⊝ (⟦ x ⟧ ⊛ ⟦ y ⟧) ≈⟨ neg-distribˡ-* ⟦ x ⟧ ⟦ y ⟧ ⟩
⊝ ⟦ x ⟧ ⊛ ⟦ y ⟧ ≈⟨ ◦-cong (sym (-‿homo x)) refl ⟩
⟦ - x ⟧ ⊛ ⟦ y ⟧ ≈⟨ sym (*-homo (- x) y) ⟩
⟦ - x * y ⟧ ∎)
neg-distribʳ-* : (∀ x y → (⊝ (x ⊛ y)) ≈₂ (x ⊛ (⊝ y))) → (∀ x y → (- (x * y)) ≈₁ (x * (- y)))
neg-distribʳ-* neg-distribʳ-* x y = injective (begin
⟦ - (x * y) ⟧ ≈⟨ -‿homo (x * y) ⟩
⊝ ⟦ x * y ⟧ ≈⟨ ⁻¹-cong (*-homo x y) ⟩
⊝ (⟦ x ⟧ ⊛ ⟦ y ⟧) ≈⟨ neg-distribʳ-* ⟦ x ⟧ ⟦ y ⟧ ⟩
⟦ x ⟧ ⊛ ⊝ ⟦ y ⟧ ≈⟨ ◦-cong refl (sym (-‿homo y)) ⟩
⟦ x ⟧ ⊛ ⟦ - y ⟧ ≈⟨ sym (*-homo x (- y)) ⟩
⟦ x * - y ⟧ ∎)
isRing : IsRing _≈₂_ _⊕_ _⊛_ ⊝_ 0#₂ 1#₂ → IsRing _≈₁_ _+_ _*_ -_ 0# 1#
isRing isRing = record
{ +-isAbelianGroup = isAbelianGroup R.+-isAbelianGroup
; *-cong = *-cong R.*-isMagma
; *-assoc = *-assoc R.*-isMagma R.*-assoc
; *-identity = *-identity R.*-isMagma R.*-identity
; distrib = distrib R.+-isGroup R.*-isMagma R.distrib
; zero = zero R.+-isGroup R.*-isMagma R.zero
} where module R = IsRing isRing
isCommutativeRing : IsCommutativeRing _≈₂_ _⊕_ _⊛_ ⊝_ 0#₂ 1#₂ →
IsCommutativeRing _≈₁_ _+_ _*_ -_ 0# 1#
isCommutativeRing isCommRing = record
{ isRing = isRing C.isRing
; *-comm = *-comm C.*-isMagma C.*-comm
} where module C = IsCommutativeRing isCommRing