Skip to content

Commit c649694

Browse files
jamesmckinnaandreasabel
authored andcommitted
[fixes #1711] Refactoring Data.Nat.Divisibility and Data.Nat.DivMod (#2182)
* added new definitions to `_∣_` * `CHANGELOG` * don't declare `quotient≢0` as an `instance` * replace use of `subst` with one of `trans` * what's sauce for the goose... * switch to a `rewrite`-based solution... * tightened `import`s * simplified dependenciess * simplified dependencies; `CHANGELOG` * removed `module` abstractions * delegated proof of `quotient≢0` to `Data.Nat.Properties` * removed redundant property * cosmetic review changes; others to follow * better proof of `quotient>1` * `where` clause layout * leaving in the flipped equality; moved everything else * new lemmas moved from `Core`; knock-on consequences; lots of tidying up * tidying up; `CHANGELOG` * cosmetic tweaks * reverted to simple version * problems with exporting `quotient` * added last lemma: defining equation for `_/_` * improved `CHANGELOG` * revert: simplified imports * improved `CHANGELOG` * endpoint of simplifying the proof of `*-pres-∣` * simplified the proof of `n/m≡quotient` * simplified the proof of `∣m+n∣m⇒∣n` * simplified the proof of `∣m∸n∣n⇒∣m` * simplified `import`s * simplified a lot of proofs, esp. wrt `divides-refl` and `NonZero` reasoning * simplified more proofs, esp. wrt `divides-refl` reasoning * simplified `import`s * moved `equalityᵒ` proof out of `Core` * `CHANGELOG` * temporary fix: further `NonZero` refactoring advised! * regularised use of instance brackets * further instance simplification * further streamlining * tidied up `CHANGELOG` * simplified `NonZero` pattern matching * regularised use of instance brackets * simplified proof of `/-*-interchange` * simplified proof of `/-*-interchange` * ... permitting the migration of `*-pres-∣` to `Data.Nat.Divisibility` * tweaked proof of `/-*-interchange` * narrowed `import`s * simplified proof; renamed new proofs * Capitalisation * streamlined `import`s; streamlined proof of decidability * spurious duplication after merge * missing symbol import * replaced one use of `1 < m` with `{{NonTrivial m}}` * fixed `CHANGELOG` * removed use of identifier `k` * refactoring: more use of `NonTrivial` instances * knock-on consequences: simplified function * two new lemmas * refactoring: use of `connex` in proofs * new lemmas about `pred` * new lemmas about monus * refactoring: use of the new properties, simplifying pattern analysis * whitespace * questionable? revision after comments on #2221 * silly argument name typo; remove parens * tidied up imports of `Relation.Nullary` * removed spurious `instance` * localised appeals to `Reasoning` * further use of `variable`s * tidied up `record` name in comment * cosmetic * reconciled implicit/explicit arguments in various monus lemmas * fixed knock-on change re monus; reverted change to `m/n<m` * reverted change to `m/n≢0⇒n≤m` * reverted breaking changes involving `NonZero` inference * revised `CHANGELOG` * restored deleted proof * fix whitespace * renaming: `DivMod.nonZeroDivisor` * localised use of `≤-Reasoning` * reverted export; removed anonymous module * revert commit re `yes/no` * renamed flipped equality * tweaked comment * added alias for `equality`
1 parent f5f9727 commit c649694

9 files changed

+407
-300
lines changed

CHANGELOG.md

+40-3
Original file line numberDiff line numberDiff line change
@@ -21,15 +21,20 @@ Deprecated modules
2121
Deprecated names
2222
----------------
2323

24+
* In `Data.Nat.Divisibility.Core`:
25+
```agda
26+
*-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣
27+
```
28+
2429
New modules
2530
-----------
2631

2732
Additions to existing modules
2833
-----------------------------
2934

30-
* In `Data.Maybe.Relation.Binary.Pointwise`:
35+
* In `Data.Fin.Properties`:
3136
```agda
32-
pointwise⊆any : Pointwise R (just x) ⊆ Any (R x)
37+
nonZeroIndex : Fin n → ℕ.NonZero n
3338
```
3439

3540
* In `Data.List.Relation.Unary.All.Properties`:
@@ -39,7 +44,39 @@ Additions to existing modules
3944
```
4045

4146
* In `Data.List.Relation.Unary.AllPairs.Properties`:
42-
```
47+
```agda
4348
catMaybes⁺ : AllPairs (Pointwise R) xs → AllPairs R (catMaybes xs)
4449
tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f)
4550
```
51+
52+
* In `Data.Maybe.Relation.Binary.Pointwise`:
53+
```agda
54+
pointwise⊆any : Pointwise R (just x) ⊆ Any (R x)
55+
```
56+
57+
* In `Data.Nat.Divisibility`:
58+
```agda
59+
quotient≢0 : m ∣ n → .{{NonZero n}} → NonZero quotient
60+
61+
m|n⇒n≡quotient*m : m ∣ n → n ≡ quotient * m
62+
m|n⇒n≡m*quotient : m ∣ n → n ≡ m * quotient
63+
quotient-∣ : m ∣ n → quotient ∣ n
64+
quotient>1 : m ∣ n → m < n → 1 < quotient
65+
quotient-< : m ∣ n → .{{NonTrivial m}} → .{{NonZero n}} → quotient < n
66+
n/m≡quotient : m ∣ n → .{{_ : NonZero m}} → n / m ≡ quotient
67+
68+
m/n≡0⇒m<n : .{{_ : NonZero n}} → m / n ≡ 0 → m < n
69+
m/n≢0⇒n≤m : .{{_ : NonZero n}} → m / n ≢ 0 → n ≤ m
70+
71+
nonZeroDivisor : DivMod dividend divisor → NonZero divisor
72+
```
73+
74+
* Added new proofs in `Data.Nat.Properties`:
75+
```agda
76+
m≤n+o⇒m∸n≤o : ∀ m n {o} → m ≤ n + o → m ∸ n ≤ o
77+
m<n+o⇒m∸n<o : ∀ m n {o} → .{{NonZero o}} → m < n + o → m ∸ n < o
78+
pred-cancel-≤ : pred m ≤ pred n → (m ≡ 1 × n ≡ 0) ⊎ m ≤ n
79+
pred-cancel-< : pred m < pred n → m < n
80+
pred-injective : .{{NonZero m}} → .{{NonZero n}} → pred m ≡ pred n → m ≡ n
81+
pred-cancel-≡ : pred m ≡ pred n → ((m ≡ 0 × n ≡ 1) ⊎ (m ≡ 1 × n ≡ 0)) ⊎ m ≡ n
82+
```

src/Data/Digit.agda

+2-5
Original file line numberDiff line numberDiff line change
@@ -55,11 +55,8 @@ toNatDigits base@(suc (suc _)) n = aux (<-wellFounded-fast n) []
5555
aux : {n : ℕ} Acc _<_ n List ℕ List ℕ
5656
aux {zero} _ xs = (0 ∷ xs)
5757
aux {n@(suc _)} (acc wf) xs with does (0 <? n / base)
58-
... | false = (n % base) ∷ xs
59-
... | true = aux (wf q<n) ((n % base) ∷ xs)
60-
where
61-
q<n : n / base < n
62-
q<n = m/n<m n base (s<s z<s)
58+
... | false = (n % base) ∷ xs -- Could this more simply be n ∷ xs here?
59+
... | true = aux (wf (m/n<m n base sz<ss)) ((n % base) ∷ xs)
6360

6461
------------------------------------------------------------------------
6562
-- Converting between `ℕ` and expansions of `Digit base`

src/Data/Fin/Properties.agda

+3
Original file line numberDiff line numberDiff line change
@@ -65,6 +65,9 @@ private
6565
¬Fin0 : ¬ Fin 0
6666
¬Fin0 ()
6767

68+
nonZeroIndex : Fin n ℕ.NonZero n
69+
nonZeroIndex {n = suc _} _ = _
70+
6871
------------------------------------------------------------------------
6972
-- Bundles
7073

src/Data/Nat/Coprimality.agda

+2-12
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ open import Function.Base using (_∘_)
2121
open import Level using (0ℓ)
2222
open import Relation.Binary.PropositionalEquality.Core as P
2323
using (_≡_; _≢_; refl; trans; cong; subst)
24-
open import Relation.Nullary as Nullary using (¬_; contradiction; yes ; no)
24+
open import Relation.Nullary as Nullary using (¬_; contradiction; map′)
2525
open import Relation.Binary.Core using (Rel)
2626
open import Relation.Binary.Definitions using (Symmetric; Decidable)
2727

@@ -65,18 +65,8 @@ coprime-/gcd m n = GCD≡1⇒coprime (GCD-/gcd m n)
6565
sym : Symmetric Coprime
6666
sym c = c ∘ swap
6767

68-
private
69-
0≢1 : 01
70-
0≢1 ()
71-
72-
2+≢1 : {n} 2+ n ≢ 1
73-
2+≢1 ()
74-
7568
coprime? : Decidable Coprime
76-
coprime? m n with mkGCD m n
77-
... | (0 , g) = no (0≢1 ∘ GCD.unique g ∘ coprime⇒GCD≡1)
78-
... | (1 , g) = yes (GCD≡1⇒coprime g)
79-
... | (2+ _ , g) = no (2+≢1 ∘ GCD.unique g ∘ coprime⇒GCD≡1)
69+
coprime? m n = map′ gcd≡1⇒coprime coprime⇒gcd≡1 (gcd m n ≟ 1)
8070

8171
------------------------------------------------------------------------
8272
-- Other basic properties

0 commit comments

Comments
 (0)