diff --git a/CHANGELOG.md b/CHANGELOG.md index 2701cdde9b..bf2d01d6a3 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -21,15 +21,20 @@ Deprecated modules Deprecated names ---------------- +* In `Data.Nat.Divisibility.Core`: + ```agda + *-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣ + ``` + New modules ----------- Additions to existing modules ----------------------------- -* In `Data.Maybe.Relation.Binary.Pointwise`: +* In `Data.Fin.Properties`: ```agda - pointwise⊆any : Pointwise R (just x) ⊆ Any (R x) + nonZeroIndex : Fin n → ℕ.NonZero n ``` * In `Data.List.Relation.Unary.All.Properties`: @@ -39,7 +44,39 @@ Additions to existing modules ``` * In `Data.List.Relation.Unary.AllPairs.Properties`: - ``` + ```agda catMaybes⁺ : AllPairs (Pointwise R) xs → AllPairs R (catMaybes xs) tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f) ``` + +* In `Data.Maybe.Relation.Binary.Pointwise`: + ```agda + pointwise⊆any : Pointwise R (just x) ⊆ Any (R x) + ``` + +* In `Data.Nat.Divisibility`: + ```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 + quotient-∣ : m ∣ n → quotient ∣ n + quotient>1 : m ∣ n → m < n → 1 < quotient + quotient-< : m ∣ n → .{{NonTrivial m}} → .{{NonZero n}} → quotient < n + n/m≡quotient : m ∣ n → .{{_ : NonZero m}} → n / m ≡ quotient + + m/n≡0⇒m0 : ∀ {m n} .{{_ : NonZero n}} → m ≥ n → m / n > 0 m≥n⇒m/n>0 {m@(suc _)} {n@(suc _)} m≥n = begin - 1 ≡⟨ sym (n/n≡1 m) ⟩ + 1 ≡⟨ n/n≡1 m ⟨ m / m ≤⟨ /-monoʳ-≤ m m≥n ⟩ m / n ∎ +m/n≡0⇒m-nonZero (m≥n⇒m/n>0 n≤m) + +m/n≢0⇒n≤m : ∀ {m n} .{{_ : NonZero n}} → m / n ≢ 0 → n ≤ m +m/n≢0⇒n≤m {m} {n@(suc _)} m/n≢0 with <-≤-connex m n +... | inj₁ m0 (≢-nonZero⁻¹ o)) (≮⇒≥ n≮o) + where n∸o0 (≢-nonZero⁻¹ o)) n≥o -m*n/o*n≡m/o : ∀ m n o .⦃ _ : NonZero o ⦄ ⦃ _ : NonZero (o * n) ⦄ → +m*n/o*n≡m/o : ∀ m n o .{{_ : NonZero o}} .{{_ : NonZero (o * n)}} → m * n / (o * n) ≡ m / o -m*n/o*n≡m/o m n o ⦃ _ ⦄ ⦃ o*n≢0 ⦄ = begin-equality +m*n/o*n≡m/o m n o = begin-equality m * n / (o * n) ≡⟨ /-congˡ (*-comm m n) ⟩ n * m / (o * n) ≡⟨ /-congʳ (*-comm o n) ⟩ n * m / (n * o) ≡⟨ m*n/m*o≡n/o n m o ⟩ m / o ∎ - where instance n*o≢0 = subst NonZero (*-comm o n) o*n≢0 - -m1 : (m∣n : m ∣ n) → m < n → 1 < quotient m∣n +quotient>1 {m} {n} m∣n m1 m) ⟩ + quotient m∣n * m ≡⟨ _∣_.equality m∣n ⟨ + n ∎ + where open ≤-Reasoning; instance _ = quotient≢0 m∣n + +------------------------------------------------------------------------ +-- Relating _/_ and quotient + +n/m≡quotient : (m∣n : m ∣ n) .{{_ : NonZero m}} → n / m ≡ quotient m∣n +n/m≡quotient {m = m} (divides-refl q) = m*n/n≡m q m ------------------------------------------------------------------------ -- Relationship with _%_ m%n≡0⇒n∣m : ∀ m n .{{_ : NonZero n}} → m % n ≡ 0 → n ∣ m -m%n≡0⇒n∣m m n eq = divides (m / n) (begin-equality - m ≡⟨ m≡m%n+[m/n]*n m n ⟩ - m % n + m / n * n ≡⟨ cong₂ _+_ eq refl ⟩ - m / n * n ∎) - where open ≤-Reasoning +m%n≡0⇒n∣m m n eq = divides (m / n) $ begin + m ≡⟨ m≡m%n+[m/n]*n m n ⟩ + m % n + [m/n]*n ≡⟨ cong (_+ [m/n]*n) eq ⟩ + [m/n]*n ∎ + where open ≡-Reasoning; [m/n]*n = m / n * n n∣m⇒m%n≡0 : ∀ m n .{{_ : NonZero n}} → n ∣ m → m % n ≡ 0 n∣m⇒m%n≡0 .(q * n) n (divides-refl q) = m*n%n≡0 q n @@ -53,15 +91,11 @@ m%n≡0⇔n∣m m n = mk⇔ (m%n≡0⇒n∣m m n) (n∣m⇒m%n≡0 m n) ------------------------------------------------------------------------ -- Properties of _∣_ and _≤_ -∣⇒≤ : ∀ {m n} .{{_ : NonZero n}} → m ∣ n → m ≤ n -∣⇒≤ {m} {n@(suc _)} (divides (suc q) eq) = begin - m ≤⟨ m≤m+n m (q * m) ⟩ - suc q * m ≡⟨ sym eq ⟩ - n ∎ - where open ≤-Reasoning +∣⇒≤ : .{{_ : NonZero n}} → m ∣ n → m ≤ n +∣⇒≤ {n@.(q * m)} {m} (divides-refl q@(suc p)) = m≤m+n m (p * m) ->⇒∤ : ∀ {m n} .{{_ : NonZero n}} → m > n → m ∤ n ->⇒∤ (s≤s m>n) m∣n = contradiction (∣⇒≤ m∣n) (≤⇒≯ m>n) +>⇒∤ : .{{_ : NonZero n}} → m > n → m ∤ n +>⇒∤ {n@(suc _)} n1⇒m*n>1 m n rewrite *-comm m n = m≢0∧n>1⇒m*n>1 n m -- Other properties of _*_ and _≤_/_<_ *-cancelʳ-≤ : ∀ m n o .{{_ : NonZero o}} → m * o ≤ n * o → m ≤ n -*-cancelʳ-≤ zero _ (suc o) _ = z≤n -*-cancelʳ-≤ (suc m) (suc n) (suc o) le = - s≤s (*-cancelʳ-≤ m n (suc o) (+-cancelˡ-≤ _ _ _ le)) +*-cancelʳ-≤ zero _ _ _ = z≤n +*-cancelʳ-≤ (suc m) (suc n) o@(suc _) le = + s≤s (*-cancelʳ-≤ m n o (+-cancelˡ-≤ _ _ _ le)) *-cancelˡ-≤ : ∀ o .{{_ : NonZero o}} → o * m ≤ o * n → m ≤ n *-cancelˡ-≤ {m} {n} o rewrite *-comm o m | *-comm o n = *-cancelʳ-≤ m n o @@ -969,14 +969,12 @@ n≢0∧m>1⇒m*n>1 m n rewrite *-comm m n = m≢0∧n>1⇒m*n>1 n m *-mono-< (s