@@ -11,23 +11,98 @@ open import Algebra.Apartness.Bundles using (HeytingCommutativeRing)
1111module Algebra.Apartness.Properties.HeytingCommutativeRing
1212 {c ℓ₁ ℓ₂} (HCR : HeytingCommutativeRing c ℓ₁ ℓ₂) where
1313
14- open import Data.Product.Base using (_,_; proj₂)
15- open import Algebra using (CommutativeRing; RightIdentity)
14+ open import Function.Base using (_∘_)
15+ open import Data.Product.Base using (_,_; proj₁; proj₂)
16+ open import Algebra using (CommutativeRing; RightIdentity; Invertible; LeftInvertible; RightInvertible)
1617
1718open HeytingCommutativeRing HCR
18- open CommutativeRing commutativeRing using (ring)
19+ open CommutativeRing commutativeRing using (ring; *-commutativeMonoid )
1920
20- open import Algebra.Properties.Ring ring using (-0#≈0#)
21+ open import Algebra.Properties.Ring ring
22+ using (-0#≈0#; -‿distribˡ-*; -‿distribʳ-*; -‿anti-homo-+; -‿involutive)
23+ open import Relation.Binary.Definitions using (Symmetric)
24+ import Relation.Binary.Reasoning.Setoid as ReasonSetoid
25+ open import Algebra.Properties.CommutativeMonoid *-commutativeMonoid
2126
27+ private variable
28+ x y z : Carrier
29+
30+ invertibleˡ⇒# : LeftInvertible _≈_ 1# _*_ (x - y) → x # y
31+ invertibleˡ⇒# = invertible⇒# ∘ invertibleˡ⇒invertible
32+
33+ invertibleʳ⇒# : RightInvertible _≈_ 1# _*_ (x - y) → x # y
34+ invertibleʳ⇒# = invertible⇒# ∘ invertibleʳ⇒invertible
2235
2336x-0≈x : RightIdentity _≈_ 0# _-_
2437x-0≈x x = trans (+-congˡ -0#≈0#) (+-identityʳ x)
2538
26391#0 : 1# # 0#
27- 1#0 = invertible ⇒# (1# , 1*[x-0]≈x , [x-0]*1 ≈x)
40+ 1#0 = invertibleˡ ⇒# (1# , 1*[x-0]≈x)
2841 where
29- 1*[x-0]≈x : ∀ {x} → 1# * (x - 0#) ≈ x
42+ 1*[x-0]≈x : 1# * (x - 0#) ≈ x
3043 1*[x-0]≈x {x} = trans (*-identityˡ (x - 0#)) (x-0≈x x)
3144
32- [x-0]*1≈x : ∀ {x} → (x - 0#) * 1# ≈ x
33- [x-0]*1≈x {x} = trans (*-comm (x - 0#) 1#) 1*[x-0]≈x
45+ x#0y#0→xy#0 : x # 0# → y # 0# → x * y # 0#
46+ x#0y#0→xy#0 {x} {y} x#0 y#0 = helper (#⇒invertible x#0) (#⇒invertible y#0)
47+ where
48+
49+ helper : Invertible _≈_ 1# _*_ (x - 0#) → Invertible _≈_ 1# _*_ (y - 0#) → x * y # 0#
50+ helper (x⁻¹ , x⁻¹*x≈1 , x*x⁻¹≈1) (y⁻¹ , y⁻¹*y≈1 , y*y⁻¹≈1)
51+ = invertibleˡ⇒# (y⁻¹ * x⁻¹ , y⁻¹*x⁻¹*x*y≈1)
52+ where
53+ open ReasonSetoid setoid
54+
55+ y⁻¹*x⁻¹*x*y≈1 : y⁻¹ * x⁻¹ * (x * y - 0#) ≈ 1#
56+ y⁻¹*x⁻¹*x*y≈1 = begin
57+ y⁻¹ * x⁻¹ * (x * y - 0#) ≈⟨ *-congˡ (x-0≈x (x * y)) ⟩
58+ y⁻¹ * x⁻¹ * (x * y) ≈⟨ *-assoc y⁻¹ x⁻¹ (x * y) ⟩
59+ y⁻¹ * (x⁻¹ * (x * y)) ≈˘⟨ *-congˡ (*-assoc x⁻¹ x y) ⟩
60+ y⁻¹ * ((x⁻¹ * x) * y) ≈˘⟨ *-congˡ (*-congʳ (*-congˡ (x-0≈x x))) ⟩
61+ y⁻¹ * ((x⁻¹ * (x - 0#)) * y) ≈⟨ *-congˡ (*-congʳ x⁻¹*x≈1) ⟩
62+ y⁻¹ * (1# * y) ≈⟨ *-congˡ (*-identityˡ y) ⟩
63+ y⁻¹ * y ≈˘⟨ *-congˡ (x-0≈x y) ⟩
64+ y⁻¹ * (y - 0#) ≈⟨ y⁻¹*y≈1 ⟩
65+ 1# ∎
66+
67+ #-sym : Symmetric _#_
68+ #-sym {x} {y} x#y = invertibleˡ⇒# (- x-y⁻¹ , x-y⁻¹*y-x≈1)
69+ where
70+ open ReasonSetoid setoid
71+ InvX-Y : Invertible _≈_ 1# _*_ (x - y)
72+ InvX-Y = #⇒invertible x#y
73+
74+ x-y⁻¹ = InvX-Y .proj₁
75+
76+ y-x≈-[x-y] : y - x ≈ - (x - y)
77+ y-x≈-[x-y] = begin
78+ y - x ≈˘⟨ +-congʳ (-‿involutive y) ⟩
79+ - - y - x ≈˘⟨ -‿anti-homo-+ x (- y) ⟩
80+ - (x - y) ∎
81+
82+ x-y⁻¹*y-x≈1 : (- x-y⁻¹) * (y - x) ≈ 1#
83+ x-y⁻¹*y-x≈1 = begin
84+ (- x-y⁻¹) * (y - x) ≈˘⟨ -‿distribˡ-* x-y⁻¹ (y - x) ⟩
85+ - (x-y⁻¹ * (y - x)) ≈⟨ -‿cong (*-congˡ y-x≈-[x-y]) ⟩
86+ - (x-y⁻¹ * - (x - y)) ≈˘⟨ -‿cong (-‿distribʳ-* x-y⁻¹ (x - y)) ⟩
87+ - - (x-y⁻¹ * (x - y)) ≈⟨ -‿involutive (x-y⁻¹ * ((x - y))) ⟩
88+ x-y⁻¹ * (x - y) ≈⟨ InvX-Y .proj₂ .proj₁ ⟩
89+ 1# ∎
90+
91+ #-congʳ : x ≈ y → x # z → y # z
92+ #-congʳ {x} {y} {z} x≈y x#z = helper (#⇒invertible x#z)
93+ where
94+
95+ helper : Invertible _≈_ 1# _*_ (x - z) → y # z
96+ helper (x-z⁻¹ , x-z⁻¹*x-z≈1# , x-z*x-z⁻¹≈1#)
97+ = invertibleˡ⇒# (x-z⁻¹ , x-z⁻¹*y-z≈1)
98+ where
99+ open ReasonSetoid setoid
100+
101+ x-z⁻¹*y-z≈1 : x-z⁻¹ * (y - z) ≈ 1#
102+ x-z⁻¹*y-z≈1 = begin
103+ x-z⁻¹ * (y - z) ≈˘⟨ *-congˡ (+-congʳ x≈y) ⟩
104+ x-z⁻¹ * (x - z) ≈⟨ x-z⁻¹*x-z≈1# ⟩
105+ 1# ∎
106+
107+ #-congˡ : y ≈ z → x # y → x # z
108+ #-congˡ y≈z x#y = #-sym (#-congʳ y≈z (#-sym x#y))
0 commit comments