Skip to content

Commit

Permalink
fixes issue #2237 (#2238)
Browse files Browse the repository at this point in the history
* fixes issue #2237

* leftover from #2182: subtle naming 'bug'/anomaly
  • Loading branch information
jamesmckinna authored Dec 30, 2023
1 parent 660d983 commit f0b24be
Show file tree
Hide file tree
Showing 3 changed files with 16 additions and 16 deletions.
4 changes: 2 additions & 2 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -58,8 +58,8 @@ Additions to existing modules
```agda
quotient≢0 : m ∣ n → .{{NonZero n}} → NonZero quotient
m|n⇒n≡quotient*m : m ∣ n → n ≡ quotient * m
m|n⇒n≡m*quotient : m ∣ n → n ≡ m * quotient
mn⇒n≡quotient*m : m ∣ n → n ≡ quotient * m
mn⇒n≡m*quotient : m ∣ n → n ≡ m * quotient
quotient-∣ : m ∣ n → quotient ∣ n
quotient>1 : m ∣ n → m < n → 1 < quotient
quotient-< : m ∣ n → .{{NonTrivial m}} → .{{NonZero n}} → quotient < n
Expand Down
26 changes: 13 additions & 13 deletions src/Data/Nat/Divisibility.agda
Original file line number Diff line number Diff line change
Expand Up @@ -42,20 +42,20 @@ open import Data.Nat.Divisibility.Core public hiding (*-pres-∣)
quotient≢0 : (m∣n : m ∣ n) .{{NonZero n}} NonZero (quotient m∣n)
quotient≢0 m∣n rewrite _∣_.equality m∣n = m*n≢0⇒m≢0 (quotient m∣n)

m|n⇒n≡quotient*m : (m∣n : m ∣ n) n ≡ (quotient m∣n) * m
m|n⇒n≡quotient*m m∣n = _∣_.equality m∣n
mn⇒n≡quotient*m : (m∣n : m ∣ n) n ≡ (quotient m∣n) * m
mn⇒n≡quotient*m m∣n = _∣_.equality m∣n

m|n⇒n≡m*quotient : (m∣n : m ∣ n) n ≡ m * (quotient m∣n)
m|n⇒n≡m*quotient {m = m} m∣n rewrite _∣_.equality m∣n = *-comm (quotient m∣n) m
mn⇒n≡m*quotient : (m∣n : m ∣ n) n ≡ m * (quotient m∣n)
mn⇒n≡m*quotient {m = m} m∣n rewrite _∣_.equality m∣n = *-comm (quotient m∣n) m

quotient-∣ : (m∣n : m ∣ n) (quotient m∣n) ∣ n
quotient-∣ {m = m} m∣n = divides m (m|n⇒n≡m*quotient m∣n)
quotient-∣ {m = m} m∣n = divides m (mn⇒n≡m*quotient m∣n)

quotient>1 : (m∣n : m ∣ n) m < n 1 < quotient m∣n
quotient>1 {m} {n} m∣n m<n = *-cancelˡ-< m 1 (quotient m∣n) $ begin-strict
m * 1 ≡⟨ *-identityʳ m ⟩
m <⟨ m<n ⟩
n ≡⟨ m|n⇒n≡m*quotient m∣n ⟩
n ≡⟨ mn⇒n≡m*quotient m∣n ⟩
m * quotient m∣n ∎
where open ≤-Reasoning

Expand Down Expand Up @@ -113,8 +113,8 @@ m%n≡0⇔n∣m m n = mk⇔ (m%n≡0⇒n∣m m n) (n∣m⇒m%n≡0 m n)
divides (q * p) (sym (*-assoc q p _))

∣-antisym : Antisymmetric _≡_ _∣_
∣-antisym {m} {zero} _ q∣p = m|n⇒n≡m*quotient q∣p
∣-antisym {zero} {n} p∣q _ = sym (m|n⇒n≡m*quotient p∣q)
∣-antisym {m} {zero} _ q∣p = mn⇒n≡m*quotient q∣p
∣-antisym {zero} {n} p∣q _ = sym (mn⇒n≡m*quotient p∣q)
∣-antisym {suc m} {suc n} p∣q q∣p = ≤-antisym (∣⇒≤ p∣q) (∣⇒≤ q∣p)

infix 4 _∣?_
Expand Down Expand Up @@ -233,7 +233,7 @@ m*n∣⇒n∣ m n rewrite *-comm m n = m*n∣⇒m∣ n m

*-cancelˡ-∣ : o .{{_ : NonZero o}} o * m ∣ o * n m ∣ n
*-cancelˡ-∣ {m} {n} o o*m∣o*n = divides q $ *-cancelˡ-≡ n (q * m) o $ begin-equality
o * n ≡⟨ m|n⇒n≡m*quotient o*m∣o*n ⟩
o * n ≡⟨ mn⇒n≡m*quotient o*m∣o*n ⟩
o * m * q ≡⟨ *-assoc o m q ⟩
o * (m * q) ≡⟨ cong (o *_) (*-comm q m) ⟨
o * (q * m) ∎
Expand Down Expand Up @@ -349,12 +349,12 @@ hasNonTrivialDivisor-≢ d≢n d∣n

hasNonTrivialDivisor-∣ : m HasNonTrivialDivisorLessThan n m ∣ o
o HasNonTrivialDivisorLessThan n
hasNonTrivialDivisor-∣ (hasNonTrivialDivisor d<n d∣m) n∣o
= hasNonTrivialDivisor d<n (∣-trans d∣m n∣o)
hasNonTrivialDivisor-∣ (hasNonTrivialDivisor d<n d∣m) m∣o
= hasNonTrivialDivisor d<n (∣-trans d∣m m∣o)

-- Monotonicity wrt ≤

hasNonTrivialDivisor-≤ : m HasNonTrivialDivisorLessThan n n ≤ o
m HasNonTrivialDivisorLessThan o
hasNonTrivialDivisor-≤ (hasNonTrivialDivisor d<n d∣m) m≤o
= hasNonTrivialDivisor (<-≤-trans d<n m≤o) d∣m
hasNonTrivialDivisor-≤ (hasNonTrivialDivisor d<n d∣m) n≤o
= hasNonTrivialDivisor (<-≤-trans d<n n≤o) d∣m
2 changes: 1 addition & 1 deletion src/Data/Nat/Primality.agda
Original file line number Diff line number Diff line change
Expand Up @@ -317,7 +317,7 @@ prime⇒¬composite (prime p) = p
where 2≡1⊎2≡0 = irr[0] {2} (divides-refl 0)

irreducible[1] : Irreducible 1
irreducible[1] m|1 = inj₁ (∣1⇒≡1 m|1)
irreducible[1] m1 = inj₁ (∣1⇒≡1 m1)

irreducible[2] : Irreducible 2
irreducible[2] {zero} 0∣2 with () 0∣⇒≡0 0∣2
Expand Down

0 comments on commit f0b24be

Please sign in to comment.