Skip to content

Commit 3e64f6c

Browse files
Use NonZero instance arguments through Data.Rational (#1566)
1 parent 48640ff commit 3e64f6c

File tree

5 files changed

+264
-268
lines changed

5 files changed

+264
-268
lines changed

CHANGELOG.md

Lines changed: 16 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -94,7 +94,16 @@ Non-backwards compatible changes
9494
- In `Data.Nat.DivMod`: `_/_`, `_%_`, `_div_`, `_mod_`
9595
- In `Data.Nat.Pseudorandom.LCG`: `Generator`
9696
- In `Data.Integer.DivMod`: `_divℕ_`, `_div_`, `_modℕ_`, `_mod_`
97-
97+
- In `Data.Rational`: `mkℚ+`, `normalize`, `_/_`, `1/_`
98+
- In `Data.Rational.Unnormalised`: `_/_`, `1/_`, `_÷_`
99+
100+
* Consequently the definition of `_≢0` has been removed from `Data.Rational.Unnormalised.Base`
101+
and the following proofs about it have been removed from `Data.Rational.Unnormalised.Properties`:
102+
```
103+
p≄0⇒∣↥p∣≢0 : ∀ p → p ≠ 0ℚᵘ → ℤ.∣ (↥ p) ∣ ≢0
104+
∣↥p∣≢0⇒p≄0 : ∀ p → ℤ.∣ (↥ p) ∣ ≢0 → p ≠ 0ℚᵘ
105+
```
106+
98107
* At the moment, there are 4 different ways such instance arguments can be provided,
99108
listed in order of convenience and clarity:
100109
1. By default there is always an instance of `NonZero (suc n)` for any `n` which
@@ -200,6 +209,12 @@ Deprecated modules
200209
Deprecated names
201210
----------------
202211

212+
* In `Data.Rational.Unnormalised.Properties`:
213+
```
214+
↥[p/q]≡p = ↥[n/d]≡n
215+
↧[p/q]≡q = ↧[n/d]≡d
216+
```
217+
203218
* In `Data.List.Properties`:
204219
```agda
205220
zipWith-identityˡ ↦ zipWith-zeroˡ

src/Data/Rational/Base.agda

Lines changed: 13 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,7 @@ open import Data.Nat.Coprimality as C
1919
using (Coprime; Bézout-coprime; coprime-/gcd; coprime?; ¬0-coprimeTo-2+)
2020
open import Data.Nat.Base as ℕ using (ℕ; zero; suc) hiding (module )
2121
import Data.Nat.DivMod as ℕ
22-
open import Data.Rational.Unnormalised.Base as ℚᵘ using (ℚᵘ; mkℚᵘ; _≢0)
22+
open import Data.Rational.Unnormalised.Base as ℚᵘ using (ℚᵘ; mkℚᵘ)
2323
open import Data.Product
2424
open import Data.Sign using (Sign)
2525
open import Data.Sum.Base using (inj₂)
@@ -59,7 +59,7 @@ open ℚ public using ()
5959
; denominatorℕ to ↧ₙ_
6060
)
6161

62-
mkℚ+ : n d .{d≢0 : d ≢0} .(Coprime n d)
62+
mkℚ+ : n d .{{_ : ℕ.NonZero d}} .(Coprime n d)
6363
mkℚ+ n (suc d) coprime = mkℚ (+ n) d coprime
6464

6565
------------------------------------------------------------------------
@@ -121,21 +121,21 @@ p ≤ᵇ q = (↥ p ℤ.* ↧ q) ℤ.≤ᵇ (↥ q ℤ.* ↧ p)
121121
-- A constructor for ℚ that takes two natural numbers, say 6 and 21,
122122
-- and returns them in a normalized form, e.g. say 2 and 7
123123

124-
normalize : (m n : ℕ) {n≢0 : n ≢0}
125-
normalize m n {n≢0} = mkℚ+ ((m ℕ./ gcd m n) {{g≢0}}) ((n ℕ./ gcd m n) {{g≢0}})
126-
{n/g≢0} (coprime-/gcd m n {{g≢0}})
124+
normalize : (m n : ℕ) .{{_ : ℕ.NonZero n}}
125+
normalize m n {{n≢0}} = mkℚ+ (m ℕ./ gcd m n) (n ℕ./ gcd m n) (coprime-/gcd m n)
127126
where
128-
g≢0 = ℕ.≢-nonZero (gcd[m,n]≢0 m n (inj₂ (toWitnessFalse n≢0)))
129-
n/g≢0 = fromWitnessFalse (n/gcd[m,n]≢0 m n {{ℕ.≢-nonZero (toWitnessFalse n≢0)}} {{g≢0}})
127+
instance
128+
g≢0 = ℕ.≢-nonZero (gcd[m,n]≢0 m n (inj₂ (ℕ.≢-nonZero⁻¹ n≢0)))
129+
n/g≢0 = ℕ.≢-nonZero (n/gcd[m,n]≢0 m n {{n≢0}} {{g≢0}})
130130

131131
-- A constructor for ℚ that (unlike mkℚ) automatically normalises it's
132132
-- arguments. See the constants section below for how to use this operator.
133133

134134
infixl 7 _/_
135135

136-
_/_ : (n : ℤ) (d : ℕ) {d≢0 : d ≢0}
137-
(+ n / d) {d≢0} = normalize n d {d≢0}
138-
(-[1+ n ] / d) {d≢0} = - normalize (suc n) d {d≢0}
136+
_/_ : (n : ℤ) (d : ℕ) .{{_ : ℕ.NonZero d}}
137+
(+ n / d) = normalize n d
138+
(-[1+ n ] / d) = - normalize (suc n) d
139139

140140
------------------------------------------------------------------------
141141
-- Conversion to and from unnormalized rationals
@@ -225,13 +225,13 @@ _-_ : ℚ → ℚ → ℚ
225225
p - q = p + (- q)
226226

227227
-- reciprocal: requires a proof that the numerator is not zero
228-
1/_ : (p : ℚ) .{n≢0 : ℤ.∣ ↥ p ∣ ≢0}
228+
1/_ : (p : ℚ) .{{_ : NonZero p}}
229229
1/ mkℚ +[1+ n ] d prf = mkℚ +[1+ d ] n (C.sym prf)
230230
1/ mkℚ -[1+ n ] d prf = mkℚ -[1+ d ] n (C.sym prf)
231231

232232
-- division: requires a proof that the denominator is not zero
233-
_÷_ : (p q : ℚ) .{n≢0 : ℤ.∣ ↥ q ∣ ≢0}
234-
(p ÷ q) {n≢0} = p * (1/ q) {n≢0}
233+
_÷_ : (p q : ℚ) .{{_ : NonZero q}}
234+
p ÷ q = p * (1/ q)
235235

236236
-- max
237237
_⊔_ : (p q : ℚ)

0 commit comments

Comments
 (0)