Skip to content

Commit 03dc09d

Browse files
committed
Attempt at the INT construction
1 parent 320220b commit 03dc09d

File tree

2 files changed

+299
-0
lines changed

2 files changed

+299
-0
lines changed
Lines changed: 50 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,50 @@
1+
{-# OPTIONS --safe --cubical-compatible #-}
2+
3+
module Data.Integer.IntConstruction where
4+
5+
open import Data.Nat.Base as ℕ using (ℕ)
6+
open import Function.Base using (flip)
7+
open import Relation.Binary
8+
open import Relation.Binary.PropositionalEquality
9+
10+
infixl 6 _⊖_
11+
12+
record : Set where
13+
constructor _⊖_
14+
field
15+
minuend :
16+
subtrahend :
17+
18+
_≃_ : Rel ℤ _
19+
(a ⊖ b) ≃ (c ⊖ d) = a ℕ.+ d ≡ c ℕ.+ b
20+
21+
_≤_ : Rel ℤ _
22+
(a ⊖ b) ≤ (c ⊖ d) = a ℕ.+ d ℕ.≤ c ℕ.+ b
23+
24+
_≥_ : Rel ℤ _
25+
_≥_ = flip _≤_
26+
27+
_<_ : Rel ℤ _
28+
(a ⊖ b) < (c ⊖ d) = a ℕ.+ d ℕ.< c ℕ.+ b
29+
30+
_>_ : Rel ℤ _
31+
_>_ = flip _<_
32+
33+
0ℤ :
34+
0ℤ = 00
35+
36+
1ℤ :
37+
1ℤ = 10
38+
39+
_+_ :
40+
(a ⊖ b) + (c ⊖ d) = (a ℕ.+ c) ⊖ (b ℕ.+ d)
41+
42+
_*_ :
43+
(a ⊖ b) * (c ⊖ d) = (a ℕ.* c ℕ.+ b ℕ.* d) ⊖ (a ℕ.* d ℕ.+ b ℕ.* c)
44+
45+
-_ :
46+
- (a ⊖ b) = b ⊖ a
47+
48+
∣_∣ :
49+
∣ a ⊖ b ∣ = ℕ.∣ a - b ∣′
50+
Lines changed: 249 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,249 @@
1+
{-# OPTIONS --safe --cubical-compatible #-}
2+
3+
module Data.Integer.IntConstruction.Properties where
4+
5+
import Algebra.Properties.CommutativeSemigroup as CommSemigroupProps
6+
open import Data.Integer.IntConstruction
7+
open import Data.Nat.Base as ℕ using (ℕ)
8+
import Data.Nat.Properties as ℕ
9+
open import Function.Base
10+
open import Relation.Binary
11+
open import Relation.Binary.PropositionalEquality
12+
open import Relation.Nullary.Decidable
13+
14+
open import Algebra.Definitions _≃_
15+
16+
private
17+
trans-helper
18+
: {r} (R : Rel ℕ r)
19+
a b c d e f
20+
R (a ℕ.+ d ℕ.+ (c ℕ.+ f)) (c ℕ.+ b ℕ.+ (e ℕ.+ d))
21+
R ((a ℕ.+ f) ℕ.+ (d ℕ.+ c)) ((e ℕ.+ b) ℕ.+ (d ℕ.+ c))
22+
trans-helper R a b c d e f mono = subst₂ R preamble postamble mono
23+
where
24+
open ≡-Reasoning
25+
open CommSemigroupProps ℕ.+-commutativeSemigroup
26+
preamble : (a ℕ.+ d) ℕ.+ (c ℕ.+ f) ≡ (a ℕ.+ f) ℕ.+ (d ℕ.+ c)
27+
preamble = begin
28+
(a ℕ.+ d) ℕ.+ (c ℕ.+ f) ≡⟨ cong (a ℕ.+ d ℕ.+_) (ℕ.+-comm c f) ⟩
29+
(a ℕ.+ d) ℕ.+ (f ℕ.+ c) ≡⟨ medial a d f c ⟩
30+
(a ℕ.+ f) ℕ.+ (d ℕ.+ c) ∎
31+
postamble : (c ℕ.+ b) ℕ.+ (e ℕ.+ d) ≡ (e ℕ.+ b) ℕ.+ (d ℕ.+ c)
32+
postamble = begin
33+
(c ℕ.+ b) ℕ.+ (e ℕ.+ d) ≡⟨ ℕ.+-comm (c ℕ.+ b) (e ℕ.+ d) ⟩
34+
(e ℕ.+ d) ℕ.+ (c ℕ.+ b) ≡⟨ cong (e ℕ.+ d ℕ.+_) (ℕ.+-comm c b) ⟩
35+
(e ℕ.+ d) ℕ.+ (b ℕ.+ c) ≡⟨ medial e d b c ⟩
36+
(e ℕ.+ b) ℕ.+ (d ℕ.+ c) ∎
37+
38+
------------------------------------------------------------------------
39+
-- Properties of _≃_
40+
41+
≃-refl : Reflexive _≃_
42+
≃-refl = refl
43+
44+
≃-sym : Symmetric _≃_
45+
≃-sym = sym
46+
47+
≃-trans : Transitive _≃_
48+
≃-trans {a ⊖ b} {c ⊖ d} {e ⊖ f} i≃j j≃k = ℕ.+-cancelʳ-≡ (d ℕ.+ c) (a ℕ.+ f) (e ℕ.+ b) $ trans-helper _≡_ a b c d e f (cong₂ ℕ._+_ i≃j j≃k)
49+
50+
_≃?_ : Decidable _≃_
51+
(a ⊖ b) ≃? (c ⊖ d) = a ℕ.+ d ℕ.≟ c ℕ.+ b
52+
53+
------------------------------------------------------------------------
54+
-- Properties of _≤_
55+
56+
≤-refl : Reflexive _≤_
57+
≤-refl = ℕ.≤-refl
58+
59+
≤-trans : Transitive _≤_
60+
≤-trans {a ⊖ b} {c ⊖ d} {e ⊖ f} i≤j j≤k = ℕ.+-cancelʳ-≤ (d ℕ.+ c) (a ℕ.+ f) (e ℕ.+ b) $ trans-helper ℕ._≤_ a b c d e f (ℕ.+-mono-≤ i≤j j≤k)
61+
62+
≤-antisym : Antisymmetric _≃_ _≤_
63+
≤-antisym i≤j j≤i = ℕ.≤-antisym i≤j j≤i
64+
65+
≤-total : Total _≤_
66+
≤-total (a ⊖ b) (c ⊖ d) = ℕ.≤-total (a ℕ.+ d) (c ℕ.+ b)
67+
68+
_≤?_ : Decidable _≤_
69+
(a ⊖ b) ≤? (c ⊖ d) = a ℕ.+ d ℕ.≤? c ℕ.+ b
70+
71+
------------------------------------------------------------------------
72+
-- Properties of _<_
73+
74+
<-irrefl : Irreflexive _≃_ _<_
75+
<-irrefl = ℕ.<-irrefl
76+
77+
<-trans : Transitive _<_
78+
<-trans {a ⊖ b} {c ⊖ d} {e ⊖ f} i<j j<k = ℕ.+-cancelʳ-< (d ℕ.+ c) (a ℕ.+ f) (e ℕ.+ b) $ trans-helper ℕ._<_ a b c d e f (ℕ.+-mono-< i<j j<k)
79+
80+
<-respˡ-≃ : _<_ Respectsˡ _≃_
81+
<-respˡ-≃ {a ⊖ b} {c ⊖ d} {e ⊖ f} j≃k i>j = ℕ.+-cancelʳ-< (d ℕ.+ c) (e ℕ.+ b) (a ℕ.+ f) $ trans-helper ℕ._>_ a b c d e f (ℕ.+-mono-<-≤ i>j (ℕ.≤-reflexive (sym j≃k)))
82+
83+
<-respʳ-≃ : _<_ Respectsʳ _≃_
84+
<-respʳ-≃ {a ⊖ b} {c ⊖ d} {e ⊖ f} j≃k i<j = ℕ.+-cancelʳ-< (d ℕ.+ c) (a ℕ.+ f) (e ℕ.+ b) $ trans-helper ℕ._<_ a b c d e f (ℕ.+-mono-<-≤ i<j (ℕ.≤-reflexive j≃k))
85+
86+
_<?_ : Decidable _<_
87+
(a ⊖ b) <? (c ⊖ d) = a ℕ.+ d ℕ.<? c ℕ.+ b
88+
89+
<-cmp : Trichotomous _≃_ _<_
90+
<-cmp (a ⊖ b) (c ⊖ d) = ℕ.<-cmp (a ℕ.+ d) (c ℕ.+ b)
91+
92+
------------------------------------------------------------------------
93+
-- Algebraic properties of _+_
94+
95+
+-cong : Congruent₂ _+_
96+
+-cong {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} a+d≡c+b e+h≡g+f = begin
97+
(a ℕ.+ e) ℕ.+ (d ℕ.+ h) ≡⟨ medial a e d h ⟩
98+
(a ℕ.+ d) ℕ.+ (e ℕ.+ h) ≡⟨ cong₂ ℕ._+_ a+d≡c+b e+h≡g+f ⟩
99+
(c ℕ.+ b) ℕ.+ (g ℕ.+ f) ≡⟨ medial c b g f ⟩
100+
(c ℕ.+ g) ℕ.+ (b ℕ.+ f) ∎
101+
where
102+
open ≡-Reasoning
103+
open CommSemigroupProps ℕ.+-commutativeSemigroup
104+
105+
+-assoc : Associative _+_
106+
+-assoc (a ⊖ b) (c ⊖ d) (e ⊖ f) = cong₂ ℕ._+_ (ℕ.+-assoc a c e) (sym (ℕ.+-assoc b d f))
107+
108+
+-identityˡ : LeftIdentity 0ℤ _+_
109+
+-identityˡ (a ⊖ b) = refl
110+
111+
+-identityʳ : RightIdentity 0ℤ _+_
112+
+-identityʳ (a ⊖ b) = cong₂ ℕ._+_ (ℕ.+-identityʳ a) (sym (ℕ.+-identityʳ b))
113+
114+
+-comm : Commutative _+_
115+
+-comm (a ⊖ b) (c ⊖ d) = cong₂ ℕ._+_ (ℕ.+-comm a c) (ℕ.+-comm d b)
116+
117+
------------------------------------------------------------------------
118+
-- Properties of _+_ and _≤_
119+
120+
+-mono-≤ : Monotonic₂ _≤_ _≤_ _≤_ _+_
121+
+-mono-≤ {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} a+d≤c+b e+h≤g+f = begin
122+
(a ℕ.+ e) ℕ.+ (d ℕ.+ h) ≡⟨ medial a e d h ⟩
123+
(a ℕ.+ d) ℕ.+ (e ℕ.+ h) ≤⟨ ℕ.+-mono-≤ a+d≤c+b e+h≤g+f ⟩
124+
(c ℕ.+ b) ℕ.+ (g ℕ.+ f) ≡⟨ medial c b g f ⟩
125+
(c ℕ.+ g) ℕ.+ (b ℕ.+ f) ∎
126+
where
127+
open ℕ.≤-Reasoning
128+
open CommSemigroupProps ℕ.+-commutativeSemigroup
129+
130+
------------------------------------------------------------------------
131+
-- Algebraic properties of -_
132+
133+
-‿cong : Congruent₁ -_
134+
-‿cong {a ⊖ b} {c ⊖ d} a+d≡c+b = begin
135+
b ℕ.+ c ≡⟨ ℕ.+-comm b c ⟩
136+
c ℕ.+ b ≡⟨ a+d≡c+b ⟨
137+
a ℕ.+ d ≡⟨ ℕ.+-comm a d ⟩
138+
d ℕ.+ a ∎
139+
where open ≡-Reasoning
140+
141+
+-inverseˡ : LeftInverse 0ℤ -_ _+_
142+
+-inverseˡ (a ⊖ b) = trans (ℕ.+-identityʳ (b ℕ.+ a)) (ℕ.+-comm b a)
143+
144+
+-inverseʳ : RightInverse 0ℤ -_ _+_
145+
+-inverseʳ (a ⊖ b) = trans (ℕ.+-identityʳ (a ℕ.+ b)) (ℕ.+-comm a b)
146+
147+
------------------------------------------------------------------------
148+
-- Properties of -_ and _≤_
149+
150+
-‿anti-≤ : Antitonic₁ _≤_ _≤_ -_
151+
-‿anti-≤ {a ⊖ b} {c ⊖ d} a+d≥c+b = begin
152+
b ℕ.+ c ≡⟨ ℕ.+-comm b c ⟩
153+
c ℕ.+ b ≤⟨ a+d≥c+b ⟩
154+
a ℕ.+ d ≡⟨ ℕ.+-comm a d ⟩
155+
d ℕ.+ a ∎
156+
where open ℕ.≤-Reasoning
157+
158+
------------------------------------------------------------------------
159+
-- Algebraic properties of _*_
160+
161+
*-cong : Congruent₂ _*_
162+
*-cong {a ⊖ b} {c ⊖ d} {e ⊖ f} {g ⊖ h} a+d≡c+b e+h≡g+f = ℕ.+-cancelʳ-≡ (d ℕ.* e) _ _ $ begin
163+
((a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ℕ.+ d ℕ.* e ≡⟨ ℕ.+-assoc (a ℕ.* e ℕ.+ b ℕ.* f) (c ℕ.* h ℕ.+ d ℕ.* g) (d ℕ.* e) ⟩
164+
(a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ d ℕ.* e) ≡⟨ cong (a ℕ.* e ℕ.+ b ℕ.* f ℕ.+_) (ℕ.+-comm (c ℕ.* h ℕ.+ d ℕ.* g) (d ℕ.* e)) ⟩
165+
(a ℕ.* e ℕ.+ b ℕ.* f) ℕ.+ (d ℕ.* e ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ≡⟨ medial (a ℕ.* e) (b ℕ.* f) (d ℕ.* e) (c ℕ.* h ℕ.+ d ℕ.* g) ⟩
166+
(a ℕ.* e ℕ.+ d ℕ.* e) ℕ.+ (b ℕ.* f ℕ.+ (c ℕ.* h ℕ.+ d ℕ.* g)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ e a d) (ℕ.+-comm (c ℕ.* h ℕ.+ d ℕ.* g) (b ℕ.* f)) ⟨
167+
(a ℕ.+ d) ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (λ k k ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) a+d≡c+b ⟩
168+
(c ℕ.+ b) ℕ.* e ℕ.+ ((c ℕ.* h ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ e c b) (ℕ.+-assoc (c ℕ.* h) (d ℕ.* g) (b ℕ.* f)) ⟩
169+
(c ℕ.* e ℕ.+ b ℕ.* e) ℕ.+ (c ℕ.* h ℕ.+ (d ℕ.* g ℕ.+ b ℕ.* f)) ≡⟨ medial (c ℕ.* e) (b ℕ.* e) (c ℕ.* h) (d ℕ.* g ℕ.+ b ℕ.* f) ⟩
170+
(c ℕ.* e ℕ.+ c ℕ.* h) ℕ.+ (b ℕ.* e ℕ.+ (d ℕ.* g ℕ.+ b ℕ.* f)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ c e h) (ℕ.+-assoc (b ℕ.* e) (d ℕ.* g) (b ℕ.* f)) ⟨
171+
c ℕ.* (e ℕ.+ h) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (λ k c ℕ.* k ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) e+h≡g+f ⟩
172+
c ℕ.* (g ℕ.+ f) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ cong (ℕ._+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f)) (ℕ.*-distribˡ-+ c g f) ⟩
173+
(c ℕ.* g ℕ.+ c ℕ.* f) ℕ.+ ((b ℕ.* e ℕ.+ d ℕ.* g) ℕ.+ b ℕ.* f) ≡⟨ medial (c ℕ.* g) (c ℕ.* f) (b ℕ.* e ℕ.+ d ℕ.* g) (b ℕ.* f) ⟩
174+
(c ℕ.* g ℕ.+ (b ℕ.* e ℕ.+ d ℕ.* g)) ℕ.+ (c ℕ.* f ℕ.+ b ℕ.* f) ≡⟨ cong₂ ℕ._+_ (ℕ.+-assoc (c ℕ.* g) (b ℕ.* e) (d ℕ.* g)) (ℕ.*-distribʳ-+ f c b) ⟨
175+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (c ℕ.+ b) ℕ.* f ≡⟨ cong (λ k ((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ k ℕ.* f) a+d≡c+b ⟨
176+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (a ℕ.+ d) ℕ.* f ≡⟨ cong (c ℕ.* g ℕ.+ b ℕ.* e ℕ.+ d ℕ.* g ℕ.+_) (ℕ.*-distribʳ-+ f a d) ⟩
177+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* g) ℕ.+ (a ℕ.* f ℕ.+ d ℕ.* f) ≡⟨ medial (c ℕ.* g ℕ.+ b ℕ.* e) (d ℕ.* g) (a ℕ.* f) (d ℕ.* f) ⟩
178+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ (d ℕ.* g ℕ.+ d ℕ.* f) ≡⟨ cong (((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+_) (ℕ.*-distribˡ-+ d g f) ⟨
179+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ d ℕ.* (g ℕ.+ f) ≡⟨ cong (λ k c ℕ.* g ℕ.+ b ℕ.* e ℕ.+ a ℕ.* f ℕ.+ d ℕ.* k) e+h≡g+f ⟨
180+
((c ℕ.* g ℕ.+ b ℕ.* e) ℕ.+ a ℕ.* f) ℕ.+ d ℕ.* (e ℕ.+ h) ≡⟨ cong₂ (λ j k j ℕ.+ d ℕ.* k) (ℕ.+-assoc (c ℕ.* g) (b ℕ.* e) (a ℕ.* f)) (ℕ.+-comm e h) ⟩
181+
(c ℕ.* g ℕ.+ (b ℕ.* e ℕ.+ a ℕ.* f)) ℕ.+ d ℕ.* (h ℕ.+ e) ≡⟨ cong₂ (λ j k c ℕ.* g ℕ.+ j ℕ.+ k) (ℕ.+-comm (b ℕ.* e) (a ℕ.* f)) (ℕ.*-distribˡ-+ d h e) ⟩
182+
(c ℕ.* g ℕ.+ (a ℕ.* f ℕ.+ b ℕ.* e)) ℕ.+ (d ℕ.* h ℕ.+ d ℕ.* e) ≡⟨ medial (c ℕ.* g) (a ℕ.* f ℕ.+ b ℕ.* e) (d ℕ.* h) (d ℕ.* e) ⟩
183+
(c ℕ.* g ℕ.+ d ℕ.* h) ℕ.+ ((a ℕ.* f ℕ.+ b ℕ.* e) ℕ.+ d ℕ.* e) ≡⟨ ℕ.+-assoc (c ℕ.* g ℕ.+ d ℕ.* h) (a ℕ.* f ℕ.+ b ℕ.* e) (d ℕ.* e) ⟨
184+
((c ℕ.* g ℕ.+ d ℕ.* h) ℕ.+ (a ℕ.* f ℕ.+ b ℕ.* e)) ℕ.+ d ℕ.* e ∎
185+
where
186+
open ≡-Reasoning
187+
open CommSemigroupProps ℕ.+-commutativeSemigroup
188+
189+
*-assoc : Associative _*_
190+
*-assoc (a ⊖ b) (c ⊖ d) (e ⊖ f) = cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma a b c d f e))
191+
where
192+
open ≡-Reasoning
193+
open CommSemigroupProps ℕ.+-commutativeSemigroup
194+
lemma : u v w x y z (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.* y ℕ.+ (u ℕ.* x ℕ.+ v ℕ.* w) ℕ.* z
195+
≡ u ℕ.* (w ℕ.* y ℕ.+ x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z ℕ.+ x ℕ.* y)
196+
lemma u v w x y z = begin
197+
(u ℕ.* w ℕ.+ v ℕ.* x) ℕ.* y ℕ.+ (u ℕ.* x ℕ.+ v ℕ.* w) ℕ.* z ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ y (u ℕ.* w) (v ℕ.* x)) (ℕ.*-distribʳ-+ z (u ℕ.* x) (v ℕ.* w)) ⟩
198+
((u ℕ.* w) ℕ.* y ℕ.+ (v ℕ.* x) ℕ.* y) ℕ.+ ((u ℕ.* x) ℕ.* z ℕ.+ (v ℕ.* w) ℕ.* z) ≡⟨ cong₂ ℕ._+_ (cong₂ ℕ._+_ (ℕ.*-assoc u w y) (ℕ.*-assoc v x y)) (cong₂ ℕ._+_ (ℕ.*-assoc u x z) (ℕ.*-assoc v w z)) ⟩
199+
(u ℕ.* (w ℕ.* y) ℕ.+ v ℕ.* (x ℕ.* y)) ℕ.+ (u ℕ.* (x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z)) ≡⟨ medial (u ℕ.* (w ℕ.* y)) (v ℕ.* (x ℕ.* y)) (u ℕ.* (x ℕ.* z)) (v ℕ.* (w ℕ.* z)) ⟩
200+
(u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+ (v ℕ.* (x ℕ.* y) ℕ.+ v ℕ.* (w ℕ.* z)) ≡⟨ cong ((u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+_) (ℕ.+-comm (v ℕ.* (x ℕ.* y)) (v ℕ.* (w ℕ.* z))) ⟩
201+
(u ℕ.* (w ℕ.* y) ℕ.+ u ℕ.* (x ℕ.* z)) ℕ.+ (v ℕ.* (w ℕ.* z) ℕ.+ v ℕ.* (x ℕ.* y)) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ u (w ℕ.* y) (x ℕ.* z)) (ℕ.*-distribˡ-+ v (w ℕ.* z) (x ℕ.* y)) ⟨
202+
u ℕ.* (w ℕ.* y ℕ.+ x ℕ.* z) ℕ.+ v ℕ.* (w ℕ.* z ℕ.+ x ℕ.* y) ∎
203+
204+
*-zeroˡ : LeftZero 0ℤ _*_
205+
*-zeroˡ _ = refl
206+
207+
*-zeroʳ : RightZero 0ℤ _*_
208+
*-zeroʳ _ = ℕ.+-identityʳ _
209+
210+
*-identityˡ : LeftIdentity 1ℤ _*_
211+
*-identityˡ (a ⊖ b) = cong₂ ℕ._+_ (lemma a) (sym (lemma b))
212+
where
213+
lemma : n n ℕ.+ 0 ℕ.+ 0 ≡ n
214+
lemma n = trans (ℕ.+-identityʳ (n ℕ.+ 0)) (ℕ.+-identityʳ n)
215+
216+
*-identityʳ : RightIdentity 1ℤ _*_
217+
*-identityʳ (a ⊖ b) = cong₂ ℕ._+_ l (sym r)
218+
where
219+
l : a ℕ.* 1 ℕ.+ b ℕ.* 0 ≡ a
220+
l = trans (cong₂ ℕ._+_ (ℕ.*-identityʳ a) (ℕ.*-zeroʳ b)) (ℕ.+-identityʳ a)
221+
r : a ℕ.* 0 ℕ.+ b ℕ.* 1 ≡ b
222+
r = trans (cong₂ ℕ._+_ (ℕ.*-zeroʳ a) (ℕ.*-identityʳ b)) (ℕ.+-identityˡ b)
223+
224+
*-distribˡ-+ : _*_ DistributesOverˡ _+_
225+
*-distribˡ-+ (a ⊖ b) (c ⊖ d) (e ⊖ f) = cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma a b d c f e))
226+
where
227+
open ≡-Reasoning
228+
open CommSemigroupProps ℕ.+-commutativeSemigroup
229+
lemma : u v w x y z u ℕ.* (w ℕ.+ y) ℕ.+ v ℕ.* (x ℕ.+ z)
230+
≡ (u ℕ.* w ℕ.+ v ℕ.* x) ℕ.+ (u ℕ.* y ℕ.+ v ℕ.* z)
231+
lemma u v w x y z = begin
232+
u ℕ.* (w ℕ.+ y) ℕ.+ v ℕ.* (x ℕ.+ z) ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribˡ-+ u w y) (ℕ.*-distribˡ-+ v x z) ⟩
233+
(u ℕ.* w ℕ.+ u ℕ.* y) ℕ.+ (v ℕ.* x ℕ.+ v ℕ.* z) ≡⟨ medial (u ℕ.* w) (u ℕ.* y) (v ℕ.* x) (v ℕ.* z) ⟩
234+
(u ℕ.* w ℕ.+ v ℕ.* x) ℕ.+ (u ℕ.* y ℕ.+ v ℕ.* z) ∎
235+
236+
*-distribʳ-+ : _*_ DistributesOverʳ _+_
237+
*-distribʳ-+ (a ⊖ b) (c ⊖ d) (e ⊖ f) = cong₂ ℕ._+_ (lemma a b c d e f) (sym (lemma b a c d e f))
238+
where
239+
open ≡-Reasoning
240+
open CommSemigroupProps ℕ.+-commutativeSemigroup
241+
lemma : u v w x y z (w ℕ.+ y) ℕ.* u ℕ.+ (x ℕ.+ z) ℕ.* v
242+
≡ (w ℕ.* u ℕ.+ x ℕ.* v) ℕ.+ (y ℕ.* u ℕ.+ z ℕ.* v)
243+
lemma u v w x y z = begin
244+
(w ℕ.+ y) ℕ.* u ℕ.+ (x ℕ.+ z) ℕ.* v ≡⟨ cong₂ ℕ._+_ (ℕ.*-distribʳ-+ u w y) (ℕ.*-distribʳ-+ v x z) ⟩
245+
(w ℕ.* u ℕ.+ y ℕ.* u) ℕ.+ (x ℕ.* v ℕ.+ z ℕ.* v) ≡⟨ medial (w ℕ.* u) (y ℕ.* u) (x ℕ.* v) (z ℕ.* v) ⟩
246+
(w ℕ.* u ℕ.+ x ℕ.* v) ℕ.+ (y ℕ.* u ℕ.+ z ℕ.* v) ∎
247+
248+
*-comm : Commutative _*_
249+
*-comm (a ⊖ b) (c ⊖ d) = cong₂ ℕ._+_ (cong₂ ℕ._+_ (ℕ.*-comm a c) (ℕ.*-comm b d)) (trans (ℕ.+-comm (c ℕ.* b) (d ℕ.* a)) (cong₂ ℕ._+_ (ℕ.*-comm d a) (ℕ.*-comm c b)))

0 commit comments

Comments
 (0)