From f9889363be370045da674073d3315803142261ee Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 25 Oct 2024 17:24:55 -0400 Subject: [PATCH 01/72] squares of natural numbers --- .../squares-natural-numbers.lagda.md | 82 +++++++++++++++++++ ...strict-inequality-natural-numbers.lagda.md | 15 ++++ 2 files changed, 97 insertions(+) diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index f8afd7063e..1cda66035f 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -13,6 +13,7 @@ open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.coproduct-types open import foundation.decidable-types @@ -67,6 +68,8 @@ square-root-ℕ _ (root , _) = root ### Squares of successors +For any `n` we have `(n + 1)² = (n + 2)n + 1 + ```agda square-succ-ℕ : (n : ℕ) → @@ -95,6 +98,85 @@ square-succ-succ-ℕ n = ( left-distributive-mul-add-ℕ 2 n 2)) ``` +### Any number is less than its own square + +**Proof.** The proof is by induction. Since `0` is the least natural number, be base case is trivial. Now consider a natural number `n` such that `n ≤ n²`. Then we have + +```text + (n + 1 ≤ (n + 1)²) ↔ n + 1 ≤ (n + 2) * n + 1 + ↔ n ≤ n² + n + n. +``` + +The last inequality follows by the following chain of inequalities + +```text + n ≤ n² -- by the induction hypothesis + ≤ n² + n -- since a ≤ a + b for any a,b + ≤ n² + n + n -- since a ≤ a + b for any a,b +``` + +```agda +lower-bound-square-ℕ : + (n : ℕ) → n ≤-ℕ square-ℕ n +lower-bound-square-ℕ zero-ℕ = star +lower-bound-square-ℕ (succ-ℕ n) = + concatenate-leq-eq-ℕ + ( succ-ℕ n) + ( transitive-leq-ℕ + ( n) + ( square-ℕ n) + ( square-ℕ n +ℕ n +ℕ n) + ( transitive-leq-ℕ + ( square-ℕ n) + ( square-ℕ n +ℕ n) + ( square-ℕ n +ℕ n +ℕ n) + ( leq-add-ℕ (square-ℕ n +ℕ n) n) + ( leq-add-ℕ (square-ℕ n) n)) + ( lower-bound-square-ℕ n)) + ( inv (square-succ-ℕ n)) +``` + +### If a number `n` has a square root, then the square root is at most `n` + +```agda +upper-bound-square-root-ℕ : + (n : ℕ) (H : is-square-ℕ n) → square-root-ℕ n H ≤-ℕ n +upper-bound-square-root-ℕ .(square-ℕ x) (x , refl) = + lower-bound-square-ℕ x +``` + +### Any number greater than 1 is strictly less than its square + +```agda +strict-lower-bound-square-ℕ : + (n : ℕ) → 1 <-ℕ n → n <-ℕ square-ℕ n +strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ zero-ℕ)) H = star +strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = + concatenate-le-eq-ℕ + { n +ℕ 3} + ( transitive-le-ℕ + ( n +ℕ 2) + ( square-ℕ (n +ℕ 2)) + ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) + ( strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ n)) star) + ( transitive-le-ℕ + ( square-ℕ (n +ℕ 2)) + ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) + ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) + ( le-add-succ-ℕ (square-ℕ (n +ℕ 2)) (n +ℕ 1)) + ( le-add-succ-ℕ (square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) (n +ℕ 1)))) + ( inv (square-succ-ℕ (succ-ℕ (succ-ℕ n)))) +``` + +### If a number `n` greater than 1 has a square root, then its square root is strictly smaller than `n` + +```agda +strict-upper-bound-square-root-ℕ : + (n : ℕ) → 1 <-ℕ n → (H : is-square-ℕ n) → square-root-ℕ n H <-ℕ n +strict-upper-bound-square-root-ℕ ._ B (succ-ℕ (succ-ℕ x) , refl) = + strict-lower-bound-square-ℕ (x +ℕ 2) star +``` + ### `n > √n` for `n > 1` The idea is to assume `n = m + 2 ≤ sqrt(m + 2)` for some `m : ℕ` and derive a diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index bcc199912a..163cf37f93 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -318,6 +318,21 @@ le-leq-neq-ℕ {succ-ℕ x} {succ-ℕ y} l f = le-leq-neq-ℕ {x} {y} l (λ p → f (ap succ-ℕ p)) ``` +### `x < x + y` for any nonzero natural number `y` + +```agda +le-add-succ-ℕ : + (x y : ℕ) → x <-ℕ x +ℕ (succ-ℕ y) +le-add-succ-ℕ zero-ℕ y = star +le-add-succ-ℕ (succ-ℕ x) y = + concatenate-le-eq-ℕ (le-add-succ-ℕ x y) (inv (left-successor-law-add-ℕ x y)) + +le-add-ℕ : + (x y : ℕ) → is-nonzero-ℕ y → x <-ℕ x +ℕ y +le-add-ℕ x zero-ℕ H = ex-falso (H refl) +le-add-ℕ x (succ-ℕ y) H = le-add-succ-ℕ x y +``` + ### If `1 < x` and `1 < y`, then `1 < xy` ```agda From c66773a81411ecdd9074a00de6408212047778d4 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 25 Oct 2024 17:25:10 -0400 Subject: [PATCH 02/72] add file --- .../irrationality-square-root-of-2.lagda.md | 41 +++++++++++++++++++ 1 file changed, 41 insertions(+) create mode 100644 src/elementary-number-theory/irrationality-square-root-of-2.lagda.md diff --git a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md new file mode 100644 index 0000000000..6cff4056ca --- /dev/null +++ b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md @@ -0,0 +1,41 @@ +# The irrationality of the square root of 2 + +```agda +module elementary-number-theory.irrationality-square-root-of-2 where +``` + +
Imports + +```agda + +``` + +
+ +## Idea + +The {{#concept "irrationality of √2" Agda=irrationality-sqrt-2}} is the theorem that there is no [rational number](elementary-number-theory.rational-numbers.md) whose square is $2$. + +**Proof.** Suppose that there is a rational number $q$ such that $q^2 = 2$. Then we can represent $q$ as a [reduced integer fraction](elementary-number-theory.reduced-integer-fractions.md) $\frac{a}{b}$ such that + +$$ +a^2 = 2b^2. +$$ + +Since $a^2$ is an [even number](elementary-number-theory.parity-natural-numbers.md), it follows that $a$ is an even number. In other words, $a = 2k$ for some number $k$. This implies that $a^2 = 4k^2$, and thus we see that + +$$ +2k^2 = b^2. +$$ + +Now we see that $b^2$ is even, which implies that $b$ is even. The consequence that both $a$ and $b$ are even contradicts the assumption that $\frac{a}{b}$ is a reduced fraction. + +## Properties + +### There is no reduced fraction whose square is $2$ + +```agda +not-two-square-reduced-fraction-ℤ : + ? +not-two-square-reduced-fraction-ℤ = ? +``` From 269e3d0e385e04301b4bed4904c08499e0801ac4 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 25 Oct 2024 17:26:20 -0400 Subject: [PATCH 03/72] make pre-commit --- src/elementary-number-theory.lagda.md | 1 + .../irrationality-square-root-of-2.lagda.md | 18 ++++++++++++++---- .../squares-natural-numbers.lagda.md | 4 +++- 3 files changed, 18 insertions(+), 5 deletions(-) diff --git a/src/elementary-number-theory.lagda.md b/src/elementary-number-theory.lagda.md index 5e914e9155..a479a091f0 100644 --- a/src/elementary-number-theory.lagda.md +++ b/src/elementary-number-theory.lagda.md @@ -76,6 +76,7 @@ open import elementary-number-theory.initial-segments-natural-numbers public open import elementary-number-theory.integer-fractions public open import elementary-number-theory.integer-partitions public open import elementary-number-theory.integers public +open import elementary-number-theory.irrationality-square-root-of-2 public open import elementary-number-theory.jacobi-symbol public open import elementary-number-theory.kolakoski-sequence public open import elementary-number-theory.legendre-symbol public diff --git a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md index 6cff4056ca..0707d600aa 100644 --- a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md +++ b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md @@ -14,21 +14,31 @@ module elementary-number-theory.irrationality-square-root-of-2 where ## Idea -The {{#concept "irrationality of √2" Agda=irrationality-sqrt-2}} is the theorem that there is no [rational number](elementary-number-theory.rational-numbers.md) whose square is $2$. +The {{#concept "irrationality of √2" Agda=irrationality-sqrt-2}} is the theorem +that there is no [rational number](elementary-number-theory.rational-numbers.md) +whose square is $2$. -**Proof.** Suppose that there is a rational number $q$ such that $q^2 = 2$. Then we can represent $q$ as a [reduced integer fraction](elementary-number-theory.reduced-integer-fractions.md) $\frac{a}{b}$ such that +**Proof.** Suppose that there is a rational number $q$ such that $q^2 = 2$. Then +we can represent $q$ as a +[reduced integer fraction](elementary-number-theory.reduced-integer-fractions.md) +$\frac{a}{b}$ such that $$ a^2 = 2b^2. $$ -Since $a^2$ is an [even number](elementary-number-theory.parity-natural-numbers.md), it follows that $a$ is an even number. In other words, $a = 2k$ for some number $k$. This implies that $a^2 = 4k^2$, and thus we see that +Since $a^2$ is an +[even number](elementary-number-theory.parity-natural-numbers.md), it follows +that $a$ is an even number. In other words, $a = 2k$ for some number $k$. This +implies that $a^2 = 4k^2$, and thus we see that $$ 2k^2 = b^2. $$ -Now we see that $b^2$ is even, which implies that $b$ is even. The consequence that both $a$ and $b$ are even contradicts the assumption that $\frac{a}{b}$ is a reduced fraction. +Now we see that $b^2$ is even, which implies that $b$ is even. The consequence +that both $a$ and $b$ are even contradicts the assumption that $\frac{a}{b}$ is +a reduced fraction. ## Properties diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 1cda66035f..388fd85dc0 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -100,7 +100,9 @@ square-succ-succ-ℕ n = ### Any number is less than its own square -**Proof.** The proof is by induction. Since `0` is the least natural number, be base case is trivial. Now consider a natural number `n` such that `n ≤ n²`. Then we have +**Proof.** The proof is by induction. Since `0` is the least natural number, be +base case is trivial. Now consider a natural number `n` such that `n ≤ n²`. Then +we have ```text (n + 1 ≤ (n + 1)²) ↔ n + 1 ≤ (n + 2) * n + 1 From 48b74ecc46c8e0f2abdc9e374a03f907e28fa7b5 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 25 Oct 2024 22:43:57 -0400 Subject: [PATCH 04/72] properties of squares --- .../addition-natural-numbers.lagda.md | 19 ++ .../distance-natural-numbers.lagda.md | 10 +- .../multiplication-natural-numbers.lagda.md | 89 ++++++- .../parity-natural-numbers.lagda.md | 30 +++ .../squares-natural-numbers.lagda.md | 230 +++++++++++++----- ...strict-inequality-natural-numbers.lagda.md | 82 +++++++ 6 files changed, 376 insertions(+), 84 deletions(-) diff --git a/src/elementary-number-theory/addition-natural-numbers.lagda.md b/src/elementary-number-theory/addition-natural-numbers.lagda.md index 94c22a9fc8..1be8622c65 100644 --- a/src/elementary-number-theory/addition-natural-numbers.lagda.md +++ b/src/elementary-number-theory/addition-natural-numbers.lagda.md @@ -138,6 +138,25 @@ abstract associative-add-ℕ ``` +### Swapping iterated addition + +```agda +abstract + right-swap-add-ℕ : + (x y z : ℕ) → (x +ℕ y) +ℕ z = (x +ℕ z) +ℕ y + right-swap-add-ℕ x y z = + associative-add-ℕ x y z ∙ + ap (add-ℕ x) (commutative-add-ℕ y z) ∙ + inv (associative-add-ℕ x z y) + + left-swap-add-ℕ : + (x y z : ℕ) → x +ℕ (y +ℕ z) = y +ℕ (x +ℕ z) + left-swap-add-ℕ x y z = + inv (associative-add-ℕ x y z) ∙ + ap (add-ℕ' z) (commutative-add-ℕ x y) ∙ + associative-add-ℕ y x z +``` + ### Addition by a fixed element on either side is injective ```agda diff --git a/src/elementary-number-theory/distance-natural-numbers.lagda.md b/src/elementary-number-theory/distance-natural-numbers.lagda.md index 1e04905a6a..312f4d7e47 100644 --- a/src/elementary-number-theory/distance-natural-numbers.lagda.md +++ b/src/elementary-number-theory/distance-natural-numbers.lagda.md @@ -350,11 +350,11 @@ left-distributive-mul-dist-ℕ (succ-ℕ m) (succ-ℕ n) (succ-ℕ k) = ( ( ap-dist-ℕ ( right-successor-law-mul-ℕ (succ-ℕ k) m) ( right-successor-law-mul-ℕ (succ-ℕ k) n)) ∙ - ( ( translation-invariant-dist-ℕ - ( succ-ℕ k) - ( (succ-ℕ k) *ℕ m) - ( (succ-ℕ k) *ℕ n)) ∙ - ( inv (left-distributive-mul-dist-ℕ m n (succ-ℕ k))))) + ( translation-invariant-dist-ℕ' + ( succ-ℕ k) + ( succ-ℕ k *ℕ m) + ( succ-ℕ k *ℕ n)) ∙ + ( inv (left-distributive-mul-dist-ℕ m n (succ-ℕ k)))) right-distributive-mul-dist-ℕ : (x y k : ℕ) → (dist-ℕ x y) *ℕ k = dist-ℕ (x *ℕ k) (y *ℕ k) diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index 4333ebb5ea..3a9eb6367b 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -67,6 +67,10 @@ triple-ℕ x = 3 *ℕ x ## Properties +### The zero laws for multiplication + +For any natural number `x` we have `0x = x0 = 0`. + ```agda abstract left-zero-law-mul-ℕ : @@ -78,7 +82,13 @@ abstract right-zero-law-mul-ℕ zero-ℕ = refl right-zero-law-mul-ℕ (succ-ℕ x) = ( right-unit-law-add-ℕ (x *ℕ zero-ℕ)) ∙ (right-zero-law-mul-ℕ x) +``` + +### The unit laws for multiplication + +For any natural number `x` we have `1x = x1 = x`. +```agda abstract right-unit-law-mul-ℕ : (x : ℕ) → x *ℕ 1 = x @@ -89,29 +99,57 @@ abstract (x : ℕ) → 1 *ℕ x = x left-unit-law-mul-ℕ zero-ℕ = refl left-unit-law-mul-ℕ (succ-ℕ x) = ap succ-ℕ (left-unit-law-mul-ℕ x) +``` + +### The successor laws for multiplication +For any two natural numbers `x` and `y` we have: + +```text + (x + 1)y = xy + y -- The left successor law + x(y + 1) = xy + x -- The right successor law + (x + 1)(y + 1) = xy + x + y + 1 -- The double successor law +``` + +```agda abstract left-successor-law-mul-ℕ : - (x y : ℕ) → (succ-ℕ x) *ℕ y = (x *ℕ y) +ℕ y + (x y : ℕ) → succ-ℕ x *ℕ y = x *ℕ y +ℕ y left-successor-law-mul-ℕ x y = refl right-successor-law-mul-ℕ : - (x y : ℕ) → x *ℕ (succ-ℕ y) = x +ℕ (x *ℕ y) + (x y : ℕ) → x *ℕ succ-ℕ y = x *ℕ y +ℕ x right-successor-law-mul-ℕ zero-ℕ y = refl right-successor-law-mul-ℕ (succ-ℕ x) y = - ( ( ap (λ t → succ-ℕ (t +ℕ y)) (right-successor-law-mul-ℕ x y)) ∙ - ( ap succ-ℕ (associative-add-ℕ x (x *ℕ y) y))) ∙ - ( inv (left-successor-law-add-ℕ x ((x *ℕ y) +ℕ y))) + ap + ( succ-ℕ) + ( ap (add-ℕ' y) (right-successor-law-mul-ℕ x y) ∙ + right-swap-add-ℕ (mul-ℕ x y) x y) + + double-successor-law-mul-ℕ : + (x y : ℕ) → succ-ℕ x *ℕ succ-ℕ y = x *ℕ y +ℕ x +ℕ y +ℕ 1 + double-successor-law-mul-ℕ x y = + left-successor-law-mul-ℕ x (succ-ℕ y) ∙ + ap + ( add-ℕ' (succ-ℕ y)) + ( right-successor-law-mul-ℕ x y) +``` + +### Multiplication is commutative +```agda abstract commutative-mul-ℕ : (x y : ℕ) → x *ℕ y = y *ℕ x commutative-mul-ℕ zero-ℕ y = inv (right-zero-law-mul-ℕ y) commutative-mul-ℕ (succ-ℕ x) y = - ( commutative-add-ℕ (x *ℕ y) y) ∙ - ( ( ap (y +ℕ_) (commutative-mul-ℕ x y)) ∙ - ( inv (right-successor-law-mul-ℕ y x))) + ap (add-ℕ' y) (commutative-mul-ℕ x y) ∙ + inv (right-successor-law-mul-ℕ y x) +``` +### The distributive laws of multiplication over addition + +```agda abstract left-distributive-mul-add-ℕ : (x y z : ℕ) → x *ℕ (y +ℕ z) = (x *ℕ y) +ℕ (x *ℕ z) @@ -135,15 +173,23 @@ abstract ( ( left-distributive-mul-add-ℕ z x y) ∙ ( ( ap (_+ℕ (z *ℕ y)) (commutative-mul-ℕ z x)) ∙ ( ap ((x *ℕ z) +ℕ_) (commutative-mul-ℕ z y)))) +``` + +### Multiplication is associative +```agda abstract associative-mul-ℕ : (x y z : ℕ) → (x *ℕ y) *ℕ z = x *ℕ (y *ℕ z) associative-mul-ℕ zero-ℕ y z = refl associative-mul-ℕ (succ-ℕ x) y z = ( right-distributive-mul-add-ℕ (x *ℕ y) y z) ∙ - ( ap (_+ℕ (y *ℕ z)) (associative-mul-ℕ x y z)) + ( ap (_+ℕ y *ℕ z) (associative-mul-ℕ x y z)) +``` + +### For any natural number we have `2x = x2 = x + x` +```agda left-two-law-mul-ℕ : (x : ℕ) → 2 *ℕ x = x +ℕ x left-two-law-mul-ℕ x = @@ -153,16 +199,30 @@ left-two-law-mul-ℕ x = right-two-law-mul-ℕ : (x : ℕ) → x *ℕ 2 = x +ℕ x right-two-law-mul-ℕ x = - ( right-successor-law-mul-ℕ x 1) ∙ - ( ap (x +ℕ_) (right-unit-law-mul-ℕ x)) + right-successor-law-mul-ℕ x 1 ∙ + ap (_+ℕ x) (right-unit-law-mul-ℕ x) +``` + +### The interchange laws for multipliplication interchanging over itself +For any four natural numbers `x`, `y`, `u`, and `v` we have + +```text + (xy)(uv) = (xu)(yv) +``` + +```agda interchange-law-mul-mul-ℕ : interchange-law mul-ℕ mul-ℕ interchange-law-mul-mul-ℕ = interchange-law-commutative-and-associative mul-ℕ commutative-mul-ℕ associative-mul-ℕ +``` + +### Multiplication by a nonzero natural number is injective +```agda is-injective-right-mul-succ-ℕ : (k : ℕ) → is-injective (_*ℕ (succ-ℕ k)) is-injective-right-mul-succ-ℕ k {zero-ℕ} {zero-ℕ} p = refl @@ -200,7 +260,11 @@ is-emb-left-mul-ℕ x H = is-emb-right-mul-ℕ : (x : ℕ) → is-nonzero-ℕ x → is-emb (_*ℕ x) is-emb-right-mul-ℕ x H = is-emb-is-injective is-set-ℕ (is-injective-right-mul-ℕ x H) +``` +### A product of natural numbers is nonzero if and only if both its factors are nonzero + +```agda is-nonzero-mul-ℕ : (x y : ℕ) → is-nonzero-ℕ x → is-nonzero-ℕ y → is-nonzero-ℕ (x *ℕ y) is-nonzero-mul-ℕ x y H K p = @@ -215,7 +279,7 @@ is-nonzero-right-factor-mul-ℕ : is-nonzero-right-factor-mul-ℕ x .zero-ℕ H refl = H (right-zero-law-mul-ℕ x) ``` -We conclude that $y = 1$ if $(x+1)y = x+1$. +### If $(x+1)y = x+1$ then it follows that $y = 1$. ```agda is-one-is-right-unit-mul-ℕ : @@ -253,6 +317,7 @@ neq-mul-ℕ m n p = ( ( m *ℕ (succ-ℕ n)) +ℕ n) ( ( p) ∙ ( ( right-successor-law-mul-ℕ (succ-ℕ m) (succ-ℕ n)) ∙ + ( commutative-add-ℕ _ (succ-ℕ m)) ∙ ( ap ((succ-ℕ m) +ℕ_) (left-successor-law-mul-ℕ m (succ-ℕ n))))) ``` diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 78da36cfb4..1ae412fef7 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -7,6 +7,7 @@ module elementary-number-theory.parity-natural-numbers where
Imports ```agda +open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types @@ -157,3 +158,32 @@ has-odd-expansion-is-odd (succ-ℕ (succ-ℕ n)) p = s : has-odd-expansion n s = has-odd-expansion-is-odd n (is-odd-is-odd-succ-succ-ℕ n p) ``` + +### A number is even if and only if it is divisible by an even number + +```agda +is-even-div-is-even-ℕ : + (n m : ℕ) → is-even-ℕ m → div-ℕ m n → is-even-ℕ n +is-even-div-is-even-ℕ ._ ._ (m' , refl) (k , refl) = + k *ℕ m' , associative-mul-ℕ k m' 2 + +is-even-div-4-ℕ : + (n : ℕ) → div-ℕ 4 n → is-even-ℕ n +is-even-div-4-ℕ n = is-even-div-is-even-ℕ n 4 (2 , refl) +``` + +### If any two out of `x`, `y`, and `x + y` are even, then so is the third + +```agda +is-even-add-ℕ : + (m n : ℕ) → is-even-ℕ m → is-even-ℕ n → is-even-ℕ (add-ℕ m n) +is-even-add-ℕ = div-add-ℕ 2 + +is-even-left-summand-ℕ : + (m n : ℕ) → is-even-ℕ n → is-even-ℕ (add-ℕ m n) → is-even-ℕ m +is-even-left-summand-ℕ = div-left-summand-ℕ 2 + +is-even-right-summand-ℕ : + (m n : ℕ) → is-even-ℕ m → is-even-ℕ (add-ℕ m n) → is-even-ℕ n +is-even-right-summand-ℕ = div-right-summand-ℕ 2 +``` diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 388fd85dc0..8329da6221 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -9,15 +9,19 @@ module elementary-number-theory.squares-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.decidable-types +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.action-on-identifications-functions open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.empty-types open import foundation.identity-types open import foundation.negation open import foundation.unit-type @@ -75,12 +79,17 @@ square-succ-ℕ : (n : ℕ) → square-ℕ (succ-ℕ n) = succ-ℕ ((succ-ℕ (succ-ℕ n)) *ℕ n) square-succ-ℕ n = - ( right-successor-law-mul-ℕ (succ-ℕ n) n) ∙ - ( commutative-add-ℕ (succ-ℕ n) ((succ-ℕ n) *ℕ n)) + right-successor-law-mul-ℕ (succ-ℕ n) n +``` + +### Squares of double successors + +For any `n` we have `(n + 2)² = n² + 4n + 4` +```agda square-succ-succ-ℕ : (n : ℕ) → - square-ℕ (succ-ℕ (succ-ℕ n)) = square-ℕ n +ℕ 2 *ℕ n +ℕ 2 *ℕ n +ℕ 4 + square-ℕ (succ-ℕ (succ-ℕ n)) = square-ℕ n +ℕ 4 *ℕ n +ℕ 4 square-succ-succ-ℕ n = equational-reasoning square-ℕ (n +ℕ 2) @@ -96,9 +105,71 @@ square-succ-succ-ℕ n = ( ap-add-ℕ {square-ℕ n +ℕ 2 *ℕ n} {2 *ℕ (n +ℕ 2)} ( refl) ( left-distributive-mul-add-ℕ 2 n 2)) + = square-ℕ n +ℕ (2 *ℕ n +ℕ 2 *ℕ n +ℕ 4) + by + associative-add-ℕ (square-ℕ n) (2 *ℕ n) (2 *ℕ n +ℕ 4) + = square-ℕ n +ℕ 4 *ℕ n +ℕ 4 + by + ap + ( add-ℕ (square-ℕ n)) + ( ap + ( add-ℕ' 4) + ( ( inv (associative-add-ℕ (2 *ℕ n) (0 +ℕ n) n)) ∙ + ( inv (ap (add-ℕ' n) (associative-add-ℕ (2 *ℕ n) 0 n))))) +``` + +### The square function is order preserving + +```agda +preserves-order-square-ℕ : + (m n : ℕ) → m ≤-ℕ n → square-ℕ m ≤-ℕ square-ℕ n +preserves-order-square-ℕ m n H = + preserves-leq-mul-ℕ m n m n H H + +preserves-strict-order-square-ℕ : + (m n : ℕ) → m <-ℕ n → square-ℕ m <-ℕ square-ℕ n +preserves-strict-order-square-ℕ m n H = + preserves-strict-order-mul-ℕ m n m n H H ``` -### Any number is less than its own square +### The square function reflects the order on the natural numbers + +For any two natural numbers `m` and `n` we have `m² ≤ n² → m ≤ n` and `m² < n² → m < n`. + +```agda +reflects-order-square-ℕ : + (m n : ℕ) → square-ℕ m ≤-ℕ square-ℕ n → m ≤-ℕ n +reflects-order-square-ℕ m n H = + leq-not-le-ℕ n m + ( λ K → + contradiction-le-ℕ + ( square-ℕ n) + ( square-ℕ m) + ( preserves-strict-order-square-ℕ n m K) + ( H)) + +reflects-strict-order-square-ℕ : + (m n : ℕ) → square-ℕ m <-ℕ square-ℕ n → m <-ℕ n +reflects-strict-order-square-ℕ m n H = + le-not-leq-ℕ n m + ( λ K → + contradiction-le-ℕ + ( square-ℕ m) + ( square-ℕ n) + ( H) + ( preserves-order-square-ℕ n m K)) +``` + +### Squares distribute over multiplication + +```agda +distributive-square-mul-ℕ : + (m n : ℕ) → square-ℕ (m *ℕ n) = square-ℕ m *ℕ square-ℕ n +distributive-square-mul-ℕ m n = + interchange-law-mul-mul-ℕ m n m n +``` + +### Any number is less than or equal to its own square **Proof.** The proof is by induction. Since `0` is the least natural number, be base case is trivial. Now consider a natural number `n` such that `n ≤ n²`. Then @@ -179,82 +250,107 @@ strict-upper-bound-square-root-ℕ ._ B (succ-ℕ (succ-ℕ x) , refl) = strict-lower-bound-square-ℕ (x +ℕ 2) star ``` -### `n > √n` for `n > 1` +### If `n² ≤ n`, then `n ≤ 1` + +```agda +leq-one-leq-square-ℕ : + (n : ℕ) → square-ℕ n ≤-ℕ n → n ≤-ℕ 1 +leq-one-leq-square-ℕ zero-ℕ H = star +leq-one-leq-square-ℕ (succ-ℕ zero-ℕ) H = star +leq-one-leq-square-ℕ (succ-ℕ (succ-ℕ n)) H = + ex-falso + ( contradiction-le-ℕ + ( n +ℕ 2) + ( square-ℕ (n +ℕ 2)) + ( strict-lower-bound-square-ℕ (n +ℕ 2) star) + ( H)) +``` + +### If the square root of a perfect square `n` is greater than or equal to `n`, then `n ≤ 1` + +```agda +leq-one-leq-square-root-ℕ : + (n : ℕ) (H : is-square-ℕ n) → n ≤-ℕ square-root-ℕ n H → n ≤-ℕ 1 +leq-one-leq-square-root-ℕ ._ (x , refl) H = + leq-one-leq-square-ℕ (square-ℕ x) (preserves-order-square-ℕ (square-ℕ x) x H) +``` -The idea is to assume `n = m + 2 ≤ sqrt(m + 2)` for some `m : ℕ` and derive a -contradiction by squaring both sides of the inequality +### The strict inequality `n² < n` never holds ```agda -greater-than-square-root-ℕ : - (n root : ℕ) → ¬ ((n +ℕ 2 ≤-ℕ root) × (n +ℕ 2 = square-ℕ root)) -greater-than-square-root-ℕ n root (pf-leq , pf-eq) = - reflects-leq-left-add-ℕ - ( square-ℕ root) - ( square-ℕ n +ℕ 2 *ℕ n +ℕ n +ℕ 2) - ( 0) - ( tr - ( leq-ℕ (square-ℕ n +ℕ 2 *ℕ n +ℕ n +ℕ 2 +ℕ square-ℕ root)) - ( inv (left-unit-law-add-ℕ (square-ℕ root))) - ( concatenate-eq-leq-ℕ - ( square-ℕ root) - ( inv - ( lemma ∙ - ( ap-add-ℕ {square-ℕ n +ℕ 2 *ℕ n +ℕ n +ℕ 2} {n +ℕ 2} - ( refl) - ( pf-eq)))) - ( preserves-leq-mul-ℕ (n +ℕ 2) root (n +ℕ 2) root pf-leq pf-leq))) - where - lemma : square-ℕ (n +ℕ 2) = square-ℕ n +ℕ 2 *ℕ n +ℕ n +ℕ 2 +ℕ n +ℕ 2 - lemma = - equational-reasoning - square-ℕ (n +ℕ 2) - = square-ℕ n +ℕ 2 *ℕ n +ℕ 2 *ℕ n +ℕ 4 - by (square-succ-succ-ℕ n) - = square-ℕ n +ℕ 2 *ℕ n +ℕ (n +ℕ 2 +ℕ n +ℕ 2) - by - ( ap-add-ℕ {square-ℕ n +ℕ 2 *ℕ n} {2 *ℕ n +ℕ 4} - ( refl) - ( equational-reasoning - 2 *ℕ n +ℕ 4 - = n +ℕ n +ℕ 2 +ℕ 2 - by - ( ap-add-ℕ {2 *ℕ n} {4} - ( left-two-law-mul-ℕ n) - ( refl)) - = n +ℕ 2 +ℕ 2 +ℕ n - by (commutative-add-ℕ n (n +ℕ 2 +ℕ 2)) - = n +ℕ 2 +ℕ (2 +ℕ n) - by (associative-add-ℕ (n +ℕ 2) 2 n) - = n +ℕ 2 +ℕ (n +ℕ 2) - by - ( ap-add-ℕ {n +ℕ 2} {2 +ℕ n} - ( refl) - ( commutative-add-ℕ 2 n)))) - = square-ℕ n +ℕ 2 *ℕ n +ℕ n +ℕ 2 +ℕ n +ℕ 2 - by - ( inv - ( associative-add-ℕ - ( square-ℕ n +ℕ 2 *ℕ n) - ( n +ℕ 2) - ( n +ℕ 2))) +not-le-square-ℕ : + (n : ℕ) → ¬ (square-ℕ n <-ℕ n) +not-le-square-ℕ n H = + contradiction-le-ℕ + ( square-ℕ n) + ( n) + ( H) + ( lower-bound-square-ℕ n) ``` -### Squareness in ℕ is decidable +### Being a square natural number is decidable ```agda -is-decidable-big-root : - (n : ℕ) → is-decidable (Σ ℕ (λ root → (n ≤-ℕ root) × (n = square-ℕ root))) -is-decidable-big-root 0 = inl (0 , star , refl) -is-decidable-big-root 1 = inl (1 , star , refl) -is-decidable-big-root (succ-ℕ (succ-ℕ n)) = - inr (λ H → greater-than-square-root-ℕ n (pr1 H) (pr2 H)) +has-greater-root-ℕ : (n : ℕ) → UU lzero +has-greater-root-ℕ n = Σ ℕ (λ x → (n ≤-ℕ x) × (n = square-ℕ x)) + +is-decidable-has-greater-root-ℕ : + (n : ℕ) → is-decidable (has-greater-root-ℕ n) +is-decidable-has-greater-root-ℕ 0 = inl (0 , star , refl) +is-decidable-has-greater-root-ℕ 1 = inl (1 , star , refl) +is-decidable-has-greater-root-ℕ (succ-ℕ (succ-ℕ n)) = + inr ( λ (x , b , p) → leq-one-leq-square-root-ℕ (n +ℕ 2) (x , p) b) is-decidable-is-square-ℕ : (n : ℕ) → is-decidable (is-square-ℕ n) is-decidable-is-square-ℕ n = is-decidable-Σ-ℕ n ( λ x → n = square-ℕ x) ( λ x → has-decidable-equality-ℕ n (square-ℕ x)) - ( is-decidable-big-root n) + ( is-decidable-has-greater-root-ℕ n) +``` + +### Equivalent characterizations of a number being even in terms of its square + +Consider a natural number `n`. The following are equivalent: +- The number `n` is even. +- Its square is even. +- Its square is divisible by 4. + +```agda +div-four-square-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → div-ℕ 4 (square-ℕ n) +pr1 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = + square-ℕ m +pr2 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = + inv (distributive-square-mul-ℕ m 2) + +is-even-square-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → is-even-ℕ (square-ℕ n) +is-even-square-is-even-ℕ .(m *ℕ 2) (m , refl) = + is-even-div-4-ℕ _ (div-four-square-is-even-ℕ (m *ℕ 2) (m , refl)) + +is-even-is-even-square-ℕ : + (n : ℕ) → is-even-ℕ (square-ℕ n) → is-even-ℕ n +is-even-is-even-square-ℕ zero-ℕ H = is-even-zero-ℕ +is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (zero-ℕ , ()) +is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (succ-ℕ k , ()) +is-even-is-even-square-ℕ (succ-ℕ (succ-ℕ n)) (m , p) = + is-even-succ-succ-is-even-ℕ n + ( is-even-is-even-square-ℕ n + ( is-even-left-summand-ℕ + ( square-ℕ n) + ( 4 *ℕ n) + ( is-even-div-4-ℕ (4 *ℕ n) (n , commutative-mul-ℕ n 4)) + ( is-even-left-summand-ℕ + ( square-ℕ n +ℕ 4 *ℕ n) + ( 4) + ( 2 , refl) + ( m , p ∙ square-succ-succ-ℕ n)))) + +is-even-div-four-square-ℕ : + (n : ℕ) → div-ℕ 4 (square-ℕ n) → is-even-ℕ n +is-even-div-four-square-ℕ n H = + is-even-is-even-square-ℕ n (is-even-div-4-ℕ (square-ℕ n) H) ``` ## See also diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 163cf37f93..d2d0f0019f 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -251,6 +251,16 @@ leq-not-le-ℕ (succ-ℕ m) zero-ℕ H = star leq-not-le-ℕ (succ-ℕ m) (succ-ℕ n) H = leq-not-le-ℕ m n H ``` +### If `m ≰ n` then `n < m` + +```agda +le-not-leq-ℕ : (m n : ℕ) → ¬ (m ≤-ℕ n) → n <-ℕ m +le-not-leq-ℕ zero-ℕ zero-ℕ H = ex-falso (H star) +le-not-leq-ℕ zero-ℕ (succ-ℕ n) H = ex-falso (H star) +le-not-leq-ℕ (succ-ℕ m) zero-ℕ H = star +le-not-leq-ℕ (succ-ℕ m) (succ-ℕ n) H = le-not-leq-ℕ m n H +``` + ### If `x < y` then `x ≤ y` ```agda @@ -339,3 +349,75 @@ le-add-ℕ x (succ-ℕ y) H = le-add-succ-ℕ x y le-one-mul-ℕ : (x y : ℕ) → 1 <-ℕ x → 1 <-ℕ y → 1 <-ℕ (x *ℕ y) le-one-mul-ℕ (succ-ℕ (succ-ℕ x)) (succ-ℕ (succ-ℕ y)) star star = star ``` + +### Addition is strictly order preserving + +```agda +preserves-strict-order-left-add-ℕ : + (a b c : ℕ) → a <-ℕ b → a +ℕ c <-ℕ b +ℕ c +preserves-strict-order-left-add-ℕ zero-ℕ (succ-ℕ b) c H = + concatenate-eq-le-eq-ℕ + ( left-unit-law-add-ℕ c) + ( le-add-succ-ℕ c b) + ( commutative-add-ℕ c (succ-ℕ b)) +preserves-strict-order-left-add-ℕ (succ-ℕ a) (succ-ℕ b) c H = + concatenate-eq-le-eq-ℕ + ( left-successor-law-add-ℕ a c) + ( preserves-strict-order-left-add-ℕ a b c H) + ( inv (left-successor-law-add-ℕ b c)) + +preserves-strict-order-right-add-ℕ : + (a c d : ℕ) → c <-ℕ d → a +ℕ c <-ℕ a +ℕ d +preserves-strict-order-right-add-ℕ a c d H = + concatenate-eq-le-eq-ℕ + ( commutative-add-ℕ a c) + ( preserves-strict-order-left-add-ℕ c d a H) + ( commutative-add-ℕ d a) + +preserves-strict-order-add-ℕ : + (a b c d : ℕ) → a <-ℕ b → c <-ℕ d → a +ℕ c <-ℕ b +ℕ d +preserves-strict-order-add-ℕ a b c d H K = + transitive-le-ℕ + ( a +ℕ c) + ( b +ℕ c) + ( b +ℕ d) + ( preserves-strict-order-left-add-ℕ a b c H) + ( preserves-strict-order-right-add-ℕ b c d K) +``` + +### Multiplication is strictly order preserving + +```agda +preserves-strict-order-mul-ℕ : + (a b c d : ℕ) → a <-ℕ b → c <-ℕ d → a *ℕ c <-ℕ b *ℕ d +preserves-strict-order-mul-ℕ zero-ℕ (succ-ℕ b) zero-ℕ (succ-ℕ d) H K = star +preserves-strict-order-mul-ℕ zero-ℕ (succ-ℕ b) (succ-ℕ c) (succ-ℕ d) H K = star +preserves-strict-order-mul-ℕ (succ-ℕ a) (succ-ℕ b) zero-ℕ (succ-ℕ d) H K = + concatenate-eq-le-ℕ + { a *ℕ 0} + { 0} + { succ-ℕ b *ℕ succ-ℕ d} + ( right-zero-law-mul-ℕ a) + ( star) +preserves-strict-order-mul-ℕ (succ-ℕ a) (succ-ℕ b) (succ-ℕ c) (succ-ℕ d) H K = + concatenate-eq-le-eq-ℕ + { succ-ℕ a *ℕ succ-ℕ c} + { a *ℕ c +ℕ a +ℕ c +ℕ 1} + { b *ℕ d +ℕ b +ℕ d +ℕ 1} + { succ-ℕ b *ℕ succ-ℕ d} + ( double-successor-law-mul-ℕ a c) + ( preserves-strict-order-add-ℕ + ( a *ℕ c +ℕ a) + ( b *ℕ d +ℕ b) + ( c) + ( d) + ( preserves-strict-order-add-ℕ + ( a *ℕ c) + ( b *ℕ d) + ( a) + ( b) + ( preserves-strict-order-mul-ℕ a b c d H K) + ( H)) + ( K)) + ( inv (double-successor-law-mul-ℕ b d)) +``` From 7952fc4a721d61b4b4bf3d809f00ec2955cb6338 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 25 Oct 2024 23:48:49 -0400 Subject: [PATCH 05/72] some work on integer fractions --- .../integer-fractions.lagda.md | 80 ++++++++----------- .../irrationality-square-root-of-2.lagda.md | 14 +++- .../multiplication-integer-fractions.lagda.md | 43 +++++++--- .../multiplication-integers.lagda.md | 18 +++++ .../reduced-integer-fractions.lagda.md | 8 +- 5 files changed, 99 insertions(+), 64 deletions(-) diff --git a/src/elementary-number-theory/integer-fractions.lagda.md b/src/elementary-number-theory/integer-fractions.lagda.md index dea0da34e6..d3d34c5340 100644 --- a/src/elementary-number-theory/integer-fractions.lagda.md +++ b/src/elementary-number-theory/integer-fractions.lagda.md @@ -10,6 +10,7 @@ module elementary-number-theory.integer-fractions where open import elementary-number-theory.greatest-common-divisor-integers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.natural-numbers open import elementary-number-theory.nonzero-integers open import elementary-number-theory.positive-and-negative-integers open import elementary-number-theory.positive-integers @@ -64,6 +65,10 @@ denominator-fraction-ℤ x = pr1 (positive-denominator-fraction-ℤ x) is-positive-denominator-fraction-ℤ : (x : fraction-ℤ) → is-positive-ℤ (denominator-fraction-ℤ x) is-positive-denominator-fraction-ℤ x = pr2 (positive-denominator-fraction-ℤ x) + +nat-denominator-fraction-ℤ : fraction-ℤ → ℕ +nat-denominator-fraction-ℤ x = + nat-positive-ℤ (positive-denominator-fraction-ℤ x) ``` ### Inclusion of the integers @@ -129,12 +134,32 @@ fraction-ℤ-Set : Set lzero fraction-ℤ-Set = fraction-ℤ , is-set-fraction-ℤ ``` +### The standard equivalence relation on integer fractions + +Two integer fractions `a/b` and `c/d` are said to be equivalent if they satisfy + +```text + ad = cb. +``` + +This relation is obviously reflexive and symmetric. To see that it is transitive, suppose that `a/b ~ c/d` and that `c/d ~ e/f`. Since `d` is positive, we have the implication + +```text + (af)d = (eb)d → af = eb. +``` + +Therefore it suffices to show that `d(af) = d(eb)`. To see this, we calculate + +```text + (af)d = (ad)f = (cb)f = (cf)b = (ed)b = (eb)d. +``` + ```agda sim-fraction-ℤ-Prop : fraction-ℤ → fraction-ℤ → Prop lzero sim-fraction-ℤ-Prop x y = Id-Prop ℤ-Set - ((numerator-fraction-ℤ x) *ℤ (denominator-fraction-ℤ y)) - ((numerator-fraction-ℤ y) *ℤ (denominator-fraction-ℤ x)) + ( numerator-fraction-ℤ x *ℤ denominator-fraction-ℤ y) + ( numerator-fraction-ℤ y *ℤ denominator-fraction-ℤ x) sim-fraction-ℤ : fraction-ℤ → fraction-ℤ → UU lzero sim-fraction-ℤ x y = type-Prop (sim-fraction-ℤ-Prop x y) @@ -150,54 +175,15 @@ symmetric-sim-fraction-ℤ x y r = inv r abstract transitive-sim-fraction-ℤ : is-transitive sim-fraction-ℤ - transitive-sim-fraction-ℤ x y z s r = + transitive-sim-fraction-ℤ (a , b , pb) y@(c , d , pd) (e , f , pf) s r = is-injective-right-mul-ℤ ( denominator-fraction-ℤ y) ( is-nonzero-denominator-fraction-ℤ y) - ( ( associative-mul-ℤ - ( numerator-fraction-ℤ x) - ( denominator-fraction-ℤ z) - ( denominator-fraction-ℤ y)) ∙ - ( ap - ( (numerator-fraction-ℤ x) *ℤ_) - ( commutative-mul-ℤ - ( denominator-fraction-ℤ z) - ( denominator-fraction-ℤ y))) ∙ - ( inv - ( associative-mul-ℤ - ( numerator-fraction-ℤ x) - ( denominator-fraction-ℤ y) - ( denominator-fraction-ℤ z))) ∙ - ( ap ( _*ℤ (denominator-fraction-ℤ z)) r) ∙ - ( associative-mul-ℤ - ( numerator-fraction-ℤ y) - ( denominator-fraction-ℤ x) - ( denominator-fraction-ℤ z)) ∙ - ( ap - ( (numerator-fraction-ℤ y) *ℤ_) - ( commutative-mul-ℤ - ( denominator-fraction-ℤ x) - ( denominator-fraction-ℤ z))) ∙ - ( inv - ( associative-mul-ℤ - ( numerator-fraction-ℤ y) - ( denominator-fraction-ℤ z) - ( denominator-fraction-ℤ x))) ∙ - ( ap (_*ℤ (denominator-fraction-ℤ x)) s) ∙ - ( associative-mul-ℤ - ( numerator-fraction-ℤ z) - ( denominator-fraction-ℤ y) - ( denominator-fraction-ℤ x)) ∙ - ( ap - ( (numerator-fraction-ℤ z) *ℤ_) - ( commutative-mul-ℤ - ( denominator-fraction-ℤ y) - ( denominator-fraction-ℤ x))) ∙ - ( inv - ( associative-mul-ℤ - ( numerator-fraction-ℤ z) - ( denominator-fraction-ℤ x) - ( denominator-fraction-ℤ y)))) + ( right-swap-mul-ℤ a f d ∙ + ap (_*ℤ f) r ∙ + right-swap-mul-ℤ c b f ∙ + ap (_*ℤ b) s ∙ + right-swap-mul-ℤ e d b) equivalence-relation-sim-fraction-ℤ : equivalence-relation lzero fraction-ℤ pr1 equivalence-relation-sim-fraction-ℤ = sim-fraction-ℤ-Prop diff --git a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md index 0707d600aa..9651b2f1d0 100644 --- a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md +++ b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md @@ -7,7 +7,14 @@ module elementary-number-theory.irrationality-square-root-of-2 where
Imports ```agda - +open import elementary-number-theory.integer-fractions +open import elementary-number-theory.integers +open import elementary-number-theory.multiplication-integer-fractions +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.reduced-integer-fractions + +open import foundation.identity-types +open import foundation.negation ```
@@ -46,6 +53,7 @@ a reduced fraction. ```agda not-two-square-reduced-fraction-ℤ : - ? -not-two-square-reduced-fraction-ℤ = ? + (x : fraction-ℤ) (H : is-reduced-fraction-ℤ x) → + ¬ (sim-fraction-ℤ (mul-fraction-ℤ x x) (in-fraction-ℤ (int-ℕ 2))) +not-two-square-reduced-fraction-ℤ x H K = {!!} ``` diff --git a/src/elementary-number-theory/multiplication-integer-fractions.lagda.md b/src/elementary-number-theory/multiplication-integer-fractions.lagda.md index 33ca87b0b6..e280b9c755 100644 --- a/src/elementary-number-theory/multiplication-integer-fractions.lagda.md +++ b/src/elementary-number-theory/multiplication-integer-fractions.lagda.md @@ -13,6 +13,7 @@ open import elementary-number-theory.integer-fractions open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers open import elementary-number-theory.multiplication-positive-and-negative-integers +open import elementary-number-theory.positive-integers open import foundation.action-on-identifications-binary-functions open import foundation.action-on-identifications-functions @@ -31,16 +32,36 @@ on the [integers](elementary-number-theory.integers.md) to the basic properties of multiplication on integer fraction only hold up to fraction similarity. -## Definition +## Definitions + +### Multiplication of integer fractions ```agda +numerator-mul-fraction-ℤ : + (x y : fraction-ℤ) → ℤ +numerator-mul-fraction-ℤ x y = + numerator-fraction-ℤ x *ℤ numerator-fraction-ℤ y + +positive-denominator-mul-fraction-ℤ : + (x y : fraction-ℤ) → positive-ℤ +positive-denominator-mul-fraction-ℤ x y = + mul-positive-ℤ + ( positive-denominator-fraction-ℤ x) + ( positive-denominator-fraction-ℤ y) + +denominator-mul-fraction-ℤ : + (x y : fraction-ℤ) → ℤ +denominator-mul-fraction-ℤ x y = + int-positive-ℤ (positive-denominator-mul-fraction-ℤ x y) + +is-positive-denominator-mul-fraction-ℤ : + (x y : fraction-ℤ) → is-positive-ℤ (denominator-mul-fraction-ℤ x y) +is-positive-denominator-mul-fraction-ℤ x y = + is-positive-int-positive-ℤ (positive-denominator-mul-fraction-ℤ x y) + mul-fraction-ℤ : fraction-ℤ → fraction-ℤ → fraction-ℤ -pr1 (mul-fraction-ℤ (m , n , n-pos) (m' , n' , n'-pos)) = - m *ℤ m' -pr1 (pr2 (mul-fraction-ℤ (m , n , n-pos) (m' , n' , n'-pos))) = - n *ℤ n' -pr2 (pr2 (mul-fraction-ℤ (m , n , n-pos) (m' , n' , n'-pos))) = - is-positive-mul-ℤ n-pos n'-pos +pr1 (mul-fraction-ℤ x y) = numerator-mul-fraction-ℤ x y +pr2 (mul-fraction-ℤ x y) = positive-denominator-mul-fraction-ℤ x y mul-fraction-ℤ' : fraction-ℤ → fraction-ℤ → fraction-ℤ mul-fraction-ℤ' x y = mul-fraction-ℤ y x @@ -98,8 +119,8 @@ right-unit-law-mul-fraction-ℤ (n , d , p) = associative-mul-fraction-ℤ : (x y z : fraction-ℤ) → sim-fraction-ℤ - (mul-fraction-ℤ (mul-fraction-ℤ x y) z) - (mul-fraction-ℤ x (mul-fraction-ℤ y z)) + ( mul-fraction-ℤ (mul-fraction-ℤ x y) z) + ( mul-fraction-ℤ x (mul-fraction-ℤ y z)) associative-mul-fraction-ℤ (nx , dx , dxp) (ny , dy , dyp) (nz , dz , dzp) = ap-mul-ℤ (associative-mul-ℤ nx ny nz) (inv (associative-mul-ℤ dx dy dz)) ``` @@ -119,8 +140,8 @@ commutative-mul-fraction-ℤ (nx , dx , dxp) (ny , dy , dyp) = left-distributive-mul-add-fraction-ℤ : (x y z : fraction-ℤ) → sim-fraction-ℤ - (mul-fraction-ℤ x (add-fraction-ℤ y z)) - (add-fraction-ℤ (mul-fraction-ℤ x y) (mul-fraction-ℤ x z)) + ( mul-fraction-ℤ x (add-fraction-ℤ y z)) + ( add-fraction-ℤ (mul-fraction-ℤ x y) (mul-fraction-ℤ x z)) left-distributive-mul-add-fraction-ℤ (nx , dx , dxp) (ny , dy , dyp) (nz , dz , dzp) = ( ap diff --git a/src/elementary-number-theory/multiplication-integers.lagda.md b/src/elementary-number-theory/multiplication-integers.lagda.md index 5306775315..201fd3f103 100644 --- a/src/elementary-number-theory/multiplication-integers.lagda.md +++ b/src/elementary-number-theory/multiplication-integers.lagda.md @@ -378,6 +378,24 @@ interchange-law-mul-mul-ℤ = associative-mul-ℤ ``` +### Swapping the order of multiplication from one side + +```agda +right-swap-mul-ℤ : + (x y z : ℤ) → (x *ℤ y) *ℤ z = (x *ℤ z) *ℤ y +right-swap-mul-ℤ x y z = + associative-mul-ℤ x y z ∙ + ap (x *ℤ_) (commutative-mul-ℤ y z) ∙ + inv (associative-mul-ℤ x z y) + +left-swap-mul-ℤ : + (x y z : ℤ) → x *ℤ (y *ℤ z) = y *ℤ (x *ℤ z) +left-swap-mul-ℤ x y z = + inv (associative-mul-ℤ x y z) ∙ + ap (_*ℤ z) (commutative-mul-ℤ x y) ∙ + associative-mul-ℤ y x z +``` + ### Computing multiplication of integers that come from natural numbers ```agda diff --git a/src/elementary-number-theory/reduced-integer-fractions.lagda.md b/src/elementary-number-theory/reduced-integer-fractions.lagda.md index b74d316dd3..c5e38528cf 100644 --- a/src/elementary-number-theory/reduced-integer-fractions.lagda.md +++ b/src/elementary-number-theory/reduced-integer-fractions.lagda.md @@ -38,12 +38,14 @@ open import foundation.universe-levels ## Idea -A reduced fraction is a fraction in which its numerator and denominator are -coprime. +A {{#concept "reduced fraction" Agda=reduced-fraction-ℤ}} is a +[fraction](elementary-number-theory.integer-fractions.md) in which its +numerator and denominator are +[coprime](elementary-number-theory.relatively-prime-integers.md). ## Definitions -### Reduced fraction +### The predicate of being a reduced fraction ```agda is-reduced-fraction-ℤ : fraction-ℤ → UU lzero From 156dd9943b9af0b16bbb37855fe295df507bda56 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 26 Oct 2024 14:08:05 -0400 Subject: [PATCH 06/72] squares of integers --- .../absolute-value-integers.lagda.md | 32 ++-- .../addition-integers.lagda.md | 19 ++ .../bezouts-lemma-integers.lagda.md | 4 +- .../finitary-natural-numbers.lagda.md | 15 +- .../integer-fractions.lagda.md | 14 +- .../irrationality-square-root-of-2.lagda.md | 12 +- .../multiplication-integers.lagda.md | 58 ++++++ .../squares-integers.lagda.md | 173 ++++++++++++++++-- .../squares-natural-numbers.lagda.md | 26 ++- ...it-elements-standard-finite-types.lagda.md | 2 +- 10 files changed, 290 insertions(+), 65 deletions(-) diff --git a/src/elementary-number-theory/absolute-value-integers.lagda.md b/src/elementary-number-theory/absolute-value-integers.lagda.md index 59455c73a5..9ff83babef 100644 --- a/src/elementary-number-theory/absolute-value-integers.lagda.md +++ b/src/elementary-number-theory/absolute-value-integers.lagda.md @@ -178,36 +178,36 @@ is-nonzero-abs-ℤ : is-nonzero-abs-ℤ (inr (inr x)) H = is-nonzero-succ-ℕ x ``` -### The absolute value function is multiplicative +### The absolute value distributes over multiplication ```agda -multiplicative-abs-ℤ' : - (x y : ℤ) → abs-ℤ (explicit-mul-ℤ x y) = (abs-ℤ x) *ℕ (abs-ℤ y) -multiplicative-abs-ℤ' (inl x) (inl y) = +distributive-abs-mul-ℤ' : + (x y : ℤ) → abs-ℤ (explicit-mul-ℤ x y) = abs-ℤ x *ℕ abs-ℤ y +distributive-abs-mul-ℤ' (inl x) (inl y) = abs-int-ℕ _ -multiplicative-abs-ℤ' (inl x) (inr (inl star)) = +distributive-abs-mul-ℤ' (inl x) (inr (inl star)) = inv (right-zero-law-mul-ℕ x) -multiplicative-abs-ℤ' (inl x) (inr (inr y)) = +distributive-abs-mul-ℤ' (inl x) (inr (inr y)) = ( abs-neg-ℤ (inl ((x *ℕ (succ-ℕ y)) +ℕ y))) ∙ ( abs-int-ℕ ((succ-ℕ x) *ℕ (succ-ℕ y))) -multiplicative-abs-ℤ' (inr (inl star)) (inl y) = +distributive-abs-mul-ℤ' (inr (inl star)) (inl y) = refl -multiplicative-abs-ℤ' (inr (inr x)) (inl y) = +distributive-abs-mul-ℤ' (inr (inr x)) (inl y) = ( abs-neg-ℤ (inl ((x *ℕ (succ-ℕ y)) +ℕ y))) ∙ ( abs-int-ℕ (succ-ℕ ((x *ℕ (succ-ℕ y)) +ℕ y))) -multiplicative-abs-ℤ' (inr (inl star)) (inr (inl star)) = +distributive-abs-mul-ℤ' (inr (inl star)) (inr (inl star)) = refl -multiplicative-abs-ℤ' (inr (inl star)) (inr (inr x)) = +distributive-abs-mul-ℤ' (inr (inl star)) (inr (inr x)) = refl -multiplicative-abs-ℤ' (inr (inr x)) (inr (inl star)) = +distributive-abs-mul-ℤ' (inr (inr x)) (inr (inl star)) = inv (right-zero-law-mul-ℕ (succ-ℕ x)) -multiplicative-abs-ℤ' (inr (inr x)) (inr (inr y)) = +distributive-abs-mul-ℤ' (inr (inr x)) (inr (inr y)) = abs-int-ℕ _ -multiplicative-abs-ℤ : - (x y : ℤ) → abs-ℤ (x *ℤ y) = (abs-ℤ x) *ℕ (abs-ℤ y) -multiplicative-abs-ℤ x y = - ap abs-ℤ (compute-mul-ℤ x y) ∙ multiplicative-abs-ℤ' x y +distributive-abs-mul-ℤ : + (x y : ℤ) → abs-ℤ (x *ℤ y) = abs-ℤ x *ℕ abs-ℤ y +distributive-abs-mul-ℤ x y = + ap abs-ℤ (compute-mul-ℤ x y) ∙ distributive-abs-mul-ℤ' x y ``` ### `|(-x)y| = |xy|` diff --git a/src/elementary-number-theory/addition-integers.lagda.md b/src/elementary-number-theory/addition-integers.lagda.md index 0074401613..614bd337cf 100644 --- a/src/elementary-number-theory/addition-integers.lagda.md +++ b/src/elementary-number-theory/addition-integers.lagda.md @@ -341,6 +341,25 @@ abstract associative-add-ℤ ``` +### Swapping the order of iterated addition + +```agda +abstract + right-swap-add-ℤ : + (x y z : ℤ) → (x +ℤ y) +ℤ z = (x +ℤ z) +ℤ y + right-swap-add-ℤ x y z = + associative-add-ℤ x y z ∙ + ap (x +ℤ_) (commutative-add-ℤ y z) ∙ + inv (associative-add-ℤ x z y) + + left-swap-add-ℤ : + (x y z : ℤ) → x +ℤ (y +ℤ z) = y +ℤ (x +ℤ z) + left-swap-add-ℤ x y z = + inv (associative-add-ℤ x y z) ∙ + ap (_+ℤ z) (commutative-add-ℤ x y) ∙ + associative-add-ℤ y x z +``` + ### Addition by `x` is a binary equivalence ```agda diff --git a/src/elementary-number-theory/bezouts-lemma-integers.lagda.md b/src/elementary-number-theory/bezouts-lemma-integers.lagda.md index 2347b1a20e..a3e59d7eb9 100644 --- a/src/elementary-number-theory/bezouts-lemma-integers.lagda.md +++ b/src/elementary-number-theory/bezouts-lemma-integers.lagda.md @@ -117,7 +117,7 @@ bezouts-lemma-eqn-to-int x y H = ( minimal-positive-distance-y-coeff (abs-ℤ x) (abs-ℤ y) H))) ( abs-ℤ y))) ( inv - ( multiplicative-abs-ℤ + ( distributive-abs-mul-ℤ ( int-ℕ (minimal-positive-distance-x-coeff (abs-ℤ x) (abs-ℤ y) H)) ( x)))) = dist-ℕ @@ -138,7 +138,7 @@ bezouts-lemma-eqn-to-int x y H = ( minimal-positive-distance-x-coeff (abs-ℤ x) (abs-ℤ y) H)) ( x)))) ( inv - ( multiplicative-abs-ℤ + ( distributive-abs-mul-ℤ ( int-ℕ (minimal-positive-distance-y-coeff (abs-ℤ x) (abs-ℤ y) H)) ( y)))) diff --git a/src/elementary-number-theory/finitary-natural-numbers.lagda.md b/src/elementary-number-theory/finitary-natural-numbers.lagda.md index 57acf9d9c0..6e6f189d42 100644 --- a/src/elementary-number-theory/finitary-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finitary-natural-numbers.lagda.md @@ -206,15 +206,12 @@ convert-based-succ-based-ℕ ( succ-ℕ (convert-based-ℕ (succ-ℕ k) (succ-based-ℕ (succ-ℕ k) n)))) +ℕ_) ( is-zero-nat-zero-Fin {k})) ∙ - ( ( ap - ( ((succ-ℕ k) *ℕ_) ∘ succ-ℕ) - ( convert-based-succ-based-ℕ (succ-ℕ k) n)) ∙ - ( ( right-successor-law-mul-ℕ - ( succ-ℕ k) - ( succ-ℕ (convert-based-ℕ (succ-ℕ k) n))) ∙ - ( commutative-add-ℕ - ( succ-ℕ k) - ( (succ-ℕ k) *ℕ (succ-ℕ (convert-based-ℕ (succ-ℕ k) n)))))) + ( ap + ( ((succ-ℕ k) *ℕ_) ∘ succ-ℕ) + ( convert-based-succ-based-ℕ (succ-ℕ k) n)) ∙ + ( ( right-successor-law-mul-ℕ + ( succ-ℕ k) + ( succ-ℕ (convert-based-ℕ (succ-ℕ k) n)))) is-section-inv-convert-based-ℕ : (k n : ℕ) → convert-based-ℕ (succ-ℕ k) (inv-convert-based-ℕ k n) = n diff --git a/src/elementary-number-theory/integer-fractions.lagda.md b/src/elementary-number-theory/integer-fractions.lagda.md index d3d34c5340..729a763f39 100644 --- a/src/elementary-number-theory/integer-fractions.lagda.md +++ b/src/elementary-number-theory/integer-fractions.lagda.md @@ -136,7 +136,7 @@ fraction-ℤ-Set = fraction-ℤ , is-set-fraction-ℤ ### The standard equivalence relation on integer fractions -Two integer fractions `a/b` and `c/d` are said to be equivalent if they satisfy +Two integer fractions `a/b` and `c/d` are said to be equivalent if they satisfy the equation ```text ad = cb. @@ -148,7 +148,7 @@ This relation is obviously reflexive and symmetric. To see that it is transitive (af)d = (eb)d → af = eb. ``` -Therefore it suffices to show that `d(af) = d(eb)`. To see this, we calculate +Therefore it suffices to show that `(af)d = (eb)d`. To see this, we calculate ```text (af)d = (ad)f = (cb)f = (cf)b = (ed)b = (eb)d. @@ -176,14 +176,10 @@ symmetric-sim-fraction-ℤ x y r = inv r abstract transitive-sim-fraction-ℤ : is-transitive sim-fraction-ℤ transitive-sim-fraction-ℤ (a , b , pb) y@(c , d , pd) (e , f , pf) s r = - is-injective-right-mul-ℤ - ( denominator-fraction-ℤ y) + is-injective-right-mul-ℤ d ( is-nonzero-denominator-fraction-ℤ y) - ( right-swap-mul-ℤ a f d ∙ - ap (_*ℤ f) r ∙ - right-swap-mul-ℤ c b f ∙ - ap (_*ℤ b) s ∙ - right-swap-mul-ℤ e d b) + ( right-swap-mul-ℤ a f d ∙ ap (_*ℤ f) r ∙ + right-swap-mul-ℤ c b f ∙ ap (_*ℤ b) s ∙ right-swap-mul-ℤ e d b) equivalence-relation-sim-fraction-ℤ : equivalence-relation lzero fraction-ℤ pr1 equivalence-relation-sim-fraction-ℤ = sim-fraction-ℤ-Prop diff --git a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md index 9651b2f1d0..cf9d371555 100644 --- a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md +++ b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md @@ -7,12 +7,16 @@ module elementary-number-theory.irrationality-square-root-of-2 where
Imports ```agda +open import elementary-number-theory.addition-integers open import elementary-number-theory.integer-fractions open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integer-fractions +open import elementary-number-theory.multiplication-integers open import elementary-number-theory.natural-numbers open import elementary-number-theory.reduced-integer-fractions +open import foundation.coproduct-types +open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.negation ``` @@ -51,9 +55,13 @@ a reduced fraction. ### There is no reduced fraction whose square is $2$ -```agda +```text not-two-square-reduced-fraction-ℤ : (x : fraction-ℤ) (H : is-reduced-fraction-ℤ x) → ¬ (sim-fraction-ℤ (mul-fraction-ℤ x x) (in-fraction-ℤ (int-ℕ 2))) -not-two-square-reduced-fraction-ℤ x H K = {!!} +not-two-square-reduced-fraction-ℤ x@(a , b , pb) H K = + {!!} ``` + +Id (mul-ℤ (mul-ℤ a a) (inr (inr 0))) +(add-ℤ (mul-ℤ b b) (mul-ℤ b b)) diff --git a/src/elementary-number-theory/multiplication-integers.lagda.md b/src/elementary-number-theory/multiplication-integers.lagda.md index 201fd3f103..5223aecf0e 100644 --- a/src/elementary-number-theory/multiplication-integers.lagda.md +++ b/src/elementary-number-theory/multiplication-integers.lagda.md @@ -160,6 +160,11 @@ left-successor-law-mul-ℤ (inr (inl star)) l = inv (right-unit-law-add-ℤ l) left-successor-law-mul-ℤ (inr (inr n)) l = refl +left-successor-law-mul-ℤ' : + (k l : ℤ) → succ-ℤ k *ℤ l = k *ℤ l +ℤ l +left-successor-law-mul-ℤ' k l = + left-successor-law-mul-ℤ k l ∙ commutative-add-ℤ l (k *ℤ l) + left-predecessor-law-mul-ℤ : (k l : ℤ) → (pred-ℤ k) *ℤ l = (neg-ℤ l) +ℤ (k *ℤ l) left-predecessor-law-mul-ℤ (inl n) l = refl @@ -174,6 +179,11 @@ left-predecessor-law-mul-ℤ (inr (inr (succ-ℕ x))) l = ( inv (left-inverse-law-add-ℤ l))) ∙ ( associative-add-ℤ (neg-ℤ l) l ((in-pos-ℤ x) *ℤ l)) +left-predecessor-law-mul-ℤ' : + (k l : ℤ) → pred-ℤ k *ℤ l = k *ℤ l -ℤ l +left-predecessor-law-mul-ℤ' k l = + left-predecessor-law-mul-ℤ k l ∙ commutative-add-ℤ (neg-ℤ l) (k *ℤ l) + right-successor-law-mul-ℤ : (k l : ℤ) → k *ℤ (succ-ℤ l) = k +ℤ (k *ℤ l) right-successor-law-mul-ℤ (inl zero-ℕ) l = inv (pred-neg-ℤ l) @@ -212,6 +222,11 @@ right-successor-law-mul-ℤ (inr (inr (succ-ℕ n))) l = ( inv (left-successor-law-add-ℤ (in-pos-ℤ n) l))))) ∙ ( associative-add-ℤ (inr (inr (succ-ℕ n))) l ((inr (inr n)) *ℤ l))))) +right-successor-law-mul-ℤ' : + (k l : ℤ) → k *ℤ succ-ℤ l = k *ℤ l +ℤ k +right-successor-law-mul-ℤ' k l = + right-successor-law-mul-ℤ k l ∙ commutative-add-ℤ k (k *ℤ l) + right-predecessor-law-mul-ℤ : (k l : ℤ) → k *ℤ (pred-ℤ l) = (neg-ℤ k) +ℤ (k *ℤ l) right-predecessor-law-mul-ℤ (inl zero-ℕ) l = @@ -245,6 +260,30 @@ right-predecessor-law-mul-ℤ (inr (inr (succ-ℕ n))) l = ( ( ap pred-ℤ (commutative-add-ℤ l (inl n))) ∙ ( inv (left-predecessor-law-add-ℤ (inl n) l))))) ∙ ( associative-add-ℤ (inl (succ-ℕ n)) l ((inr (inr n)) *ℤ l))))) + +right-predecessor-law-mul-ℤ' : + (k l : ℤ) → k *ℤ pred-ℤ l = k *ℤ l -ℤ k +right-predecessor-law-mul-ℤ' k l = + right-predecessor-law-mul-ℤ k l ∙ commutative-add-ℤ (neg-ℤ k) (k *ℤ l) + +double-successor-law-mul-ℤ : + (k l : ℤ) → succ-ℤ k *ℤ succ-ℤ l = k *ℤ l +ℤ k +ℤ l +ℤ int-ℕ 1 +double-successor-law-mul-ℤ k l = + left-successor-law-mul-ℤ' k (succ-ℤ l) ∙ + ap-add-ℤ + ( right-successor-law-mul-ℤ' k l) + ( inv (right-add-one-ℤ l)) ∙ + inv (associative-add-ℤ (k *ℤ l +ℤ k) l (int-ℕ 1)) + +double-predecessor-law-mul-ℤ : + (k l : ℤ) → pred-ℤ k *ℤ pred-ℤ l = ((k *ℤ l -ℤ k) -ℤ l) +ℤ int-ℕ 1 +double-predecessor-law-mul-ℤ k l = + left-predecessor-law-mul-ℤ' k (pred-ℤ l) ∙ + ap-add-ℤ + ( right-predecessor-law-mul-ℤ' k l) + ( ap neg-ℤ (inv (right-add-neg-one-ℤ l)) ∙ + distributive-neg-add-ℤ l neg-one-ℤ) ∙ + inv (associative-add-ℤ (k *ℤ l -ℤ k) (neg-ℤ l) (int-ℕ 1)) ``` ### Multiplication on the integers distributes on the right over addition @@ -340,6 +379,25 @@ left-distributive-mul-add-ℤ m k l = ( ap-add-ℤ (commutative-mul-ℤ k m) (commutative-mul-ℤ l m))) ``` +### Distributivity of multiplication over addition on both sides + +For any four integers `a`, `b`, `c`, and `d` we have + +```text + (a + b)(c + d) = (ac + ad) + (bc + bd). +``` + +```agda +double-distributive-mul-add-ℤ : + (a b c d : ℤ) → + (a +ℤ b) *ℤ (c +ℤ d) = (a *ℤ c +ℤ a *ℤ d) +ℤ (b *ℤ c +ℤ b *ℤ d) +double-distributive-mul-add-ℤ a b c d = + right-distributive-mul-add-ℤ a b (c +ℤ d) ∙ + ap-add-ℤ + ( left-distributive-mul-add-ℤ a c d) + ( left-distributive-mul-add-ℤ b c d) +``` + ### Right multiplication by the negative of an integer is the negative of the multiplication ```agda diff --git a/src/elementary-number-theory/squares-integers.lagda.md b/src/elementary-number-theory/squares-integers.lagda.md index 2497b17b33..67849e77ae 100644 --- a/src/elementary-number-theory/squares-integers.lagda.md +++ b/src/elementary-number-theory/squares-integers.lagda.md @@ -8,8 +8,11 @@ module elementary-number-theory.squares-integers where ```agda open import elementary-number-theory.absolute-value-integers +open import elementary-number-theory.addition-integers +open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.multiplication-positive-and-negative-integers open import elementary-number-theory.natural-numbers open import elementary-number-theory.nonnegative-integers @@ -24,6 +27,7 @@ open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.logical-equivalences open import foundation.negation +open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types @@ -32,24 +36,63 @@ open import foundation-core.transport-along-identifications
-## Definition +## Definitions + +### The square of an integer ```agda square-ℤ : ℤ → ℤ square-ℤ a = mul-ℤ a a +``` + +### The cube of an integer +```agda cube-ℤ : ℤ → ℤ cube-ℤ a = mul-ℤ (square-ℤ a) a +``` +### The predicate on integers of being square + +```agda is-square-ℤ : ℤ → UU lzero is-square-ℤ a = Σ ℤ (λ x → a = square-ℤ x) +``` +### The square root of a square integer + +```agda square-root-ℤ : (a : ℤ) → is-square-ℤ a → ℤ square-root-ℤ _ (root , _) = root ``` ## Properties +### The integer square of a natural number is the integer of its square + +```agda +compute-square-int-ℕ : (n : ℕ) → square-ℤ (int-ℕ n) = int-ℕ (square-ℕ n) +compute-square-int-ℕ n = mul-int-ℕ n n +``` + +### The square of an integer is the square of its absolute value + +```agda +compute-square-abs-ℤ : (a : ℤ) → square-ℤ a = int-ℕ (square-ℕ (abs-ℤ a)) +compute-square-abs-ℤ (inl x) = compute-mul-ℤ (inl x) (inl x) +compute-square-abs-ℤ (inr (inl star)) = refl +compute-square-abs-ℤ (inr (inr x)) = mul-int-ℕ (succ-ℕ x) (succ-ℕ x) + +compute-square-abs-ℤ' : + (a : ℤ) → square-ℤ a = square-ℤ (int-abs-ℤ a) +compute-square-abs-ℤ' (inl x) = + compute-mul-ℤ (inl x) (inl x) ∙ inv (mul-int-ℕ (succ-ℕ x) (succ-ℕ x)) +compute-square-abs-ℤ' (inr (inl star)) = + refl +compute-square-abs-ℤ' (inr (inr x)) = + refl +``` + ### Squares in ℤ are nonnegative ```agda @@ -59,22 +102,21 @@ is-nonnegative-square-ℤ a = ( λ H → is-nonnegative-is-positive-ℤ (is-positive-mul-negative-ℤ H H)) ( λ H → is-nonnegative-mul-ℤ H H) ( decide-is-negative-is-nonnegative-ℤ {a}) + +is-nonnegative-is-square-ℤ : {a : ℤ} → is-square-ℤ a → is-nonnegative-ℤ a +is-nonnegative-is-square-ℤ (r , refl) = is-nonnegative-square-ℤ r ``` ### The squares in ℤ are exactly the squares in ℕ ```agda is-square-int-is-square-nat : {n : ℕ} → is-square-ℕ n → is-square-ℤ (int-ℕ n) -is-square-int-is-square-nat (root , pf-square) = - ( ( int-ℕ root) , - ( ( ap int-ℕ pf-square) ∙ - ( inv (mul-int-ℕ root root)))) +is-square-int-is-square-nat (r , refl) = + ( int-ℕ r , inv (compute-square-int-ℕ r)) is-square-nat-is-square-int : {a : ℤ} → is-square-ℤ a → is-square-ℕ (abs-ℤ a) -is-square-nat-is-square-int (root , pf-square) = - ( ( abs-ℤ root) , - ( ( ap abs-ℤ pf-square) ∙ - ( multiplicative-abs-ℤ root root))) +is-square-nat-is-square-int (r , refl) = + ( abs-ℤ r , distributive-abs-mul-ℤ r r) iff-is-square-int-is-square-nat : (n : ℕ) → is-square-ℕ n ↔ is-square-ℤ (int-ℕ n) @@ -84,17 +126,16 @@ pr2 (iff-is-square-int-is-square-nat n) H = iff-is-nonneg-square-nat-is-square-int : (a : ℤ) → is-square-ℤ a ↔ is-nonnegative-ℤ a × is-square-ℕ (abs-ℤ a) -pr1 (iff-is-nonneg-square-nat-is-square-int a) (root , pf-square) = - ( ( tr is-nonnegative-ℤ (inv pf-square) (is-nonnegative-square-ℤ root)) , - ( is-square-nat-is-square-int (root , pf-square))) -pr2 - ( iff-is-nonneg-square-nat-is-square-int a) (pf-nonneg , (root , pf-square)) = - ( ( int-ℕ root) , - ( ( inv (int-abs-is-nonnegative-ℤ a pf-nonneg)) ∙ - ( pr2 (is-square-int-is-square-nat (root , pf-square))))) +pr1 (iff-is-nonneg-square-nat-is-square-int a) H = + ( is-nonnegative-is-square-ℤ H , is-square-nat-is-square-int H) +pr2 (iff-is-nonneg-square-nat-is-square-int a) (H , r , p) = + ( int-ℕ r , + inv (int-abs-is-nonnegative-ℤ a H) ∙ + ap int-ℕ p ∙ + inv (compute-square-int-ℕ r)) ``` -### Squareness in ℤ is decidable +### Being a square integer is decidable ```agda is-decidable-is-square-ℤ : (a : ℤ) → is-decidable (is-square-ℤ a) @@ -107,3 +148,99 @@ is-decidable-is-square-ℤ (inr (inr n)) = ( is-square-nat-is-square-int) ( is-decidable-is-square-ℕ (succ-ℕ n)) ``` + +### Squares of successors + +For any integer `a` we have the equation + +```text + (a + 1)² = a² + 2a + 1. +``` + +```agda +square-succ-ℤ : + (a : ℤ) → square-ℤ (succ-ℤ a) = square-ℤ a +ℤ int-ℕ 2 *ℤ a +ℤ int-ℕ 1 +square-succ-ℤ a = + double-successor-law-mul-ℤ a a ∙ + ap (_+ℤ int-ℕ 1) (associative-add-ℤ (square-ℤ a) a a) +``` + +### Squares of double successors + +For any `n` we have `(n + 2)² = n² + 4n + 4` + +```agda +square-succ-succ-ℤ : + (a : ℤ) → + square-ℤ (succ-ℤ (succ-ℤ a)) = square-ℤ a +ℤ int-ℕ 4 *ℤ a +ℤ int-ℕ 4 +square-succ-succ-ℤ a = + ap + ( square-ℤ) + ( ap succ-ℤ (inv (right-add-one-ℤ a)) ∙ + inv (right-add-one-ℤ (a +ℤ one-ℤ)) ∙ + associative-add-ℤ a one-ℤ one-ℤ) ∙ + double-distributive-mul-add-ℤ a (int-ℕ 2) a (int-ℕ 2) ∙ + inv + ( associative-add-ℤ (square-ℤ a +ℤ a *ℤ int-ℕ 2) (int-ℕ 2 *ℤ a) (int-ℕ 4)) ∙ + ap + ( _+ℤ int-ℕ 4) + ( associative-add-ℤ (square-ℤ a) (a *ℤ int-ℕ 2) (int-ℕ 2 *ℤ a) ∙ + ap + ( square-ℤ a +ℤ_) + ( ap (_+ℤ int-ℕ 2 *ℤ a) (commutative-mul-ℤ a (int-ℕ 2)) ∙ + inv (right-distributive-mul-add-ℤ (int-ℕ 2) (int-ℕ 2) a))) +``` + +### Squares distribute over multiplication + +```agda +distributive-square-mul-ℤ : + (a b : ℤ) → square-ℤ (a *ℤ b) = square-ℤ a *ℤ square-ℤ b +distributive-square-mul-ℤ a b = + interchange-law-mul-mul-ℤ a b a b +``` + +### Equivalent characterizations of a number being even in terms of its square + +Consider a integer `n`. The following are equivalent: + +- The number `n` is even. +- Its square is even. +- Its square is divisible by 4. + +```text +div-four-square-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → div-ℕ 4 (square-ℕ n) +pr1 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = + square-ℕ m +pr2 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = + inv (distributive-square-mul-ℕ m 2) + +is-even-square-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → is-even-ℕ (square-ℕ n) +is-even-square-is-even-ℕ .(m *ℕ 2) (m , refl) = + is-even-div-4-ℕ _ (div-four-square-is-even-ℕ (m *ℕ 2) (m , refl)) + +is-even-is-even-square-ℕ : + (n : ℕ) → is-even-ℕ (square-ℕ n) → is-even-ℕ n +is-even-is-even-square-ℕ zero-ℕ H = is-even-zero-ℕ +is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (zero-ℕ , ()) +is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (succ-ℕ k , ()) +is-even-is-even-square-ℕ (succ-ℕ (succ-ℕ n)) (m , p) = + is-even-succ-succ-is-even-ℕ n + ( is-even-is-even-square-ℕ n + ( is-even-left-summand-ℕ + ( square-ℕ n) + ( 4 *ℕ n) + ( is-even-div-4-ℕ (4 *ℕ n) (n , commutative-mul-ℕ n 4)) + ( is-even-left-summand-ℕ + ( square-ℕ n +ℕ 4 *ℕ n) + ( 4) + ( 2 , refl) + ( m , p ∙ square-succ-succ-ℕ n)))) + +is-even-div-four-square-ℕ : + (n : ℕ) → div-ℕ 4 (square-ℕ n) → is-even-ℕ n +is-even-div-four-square-ℕ n H = + is-even-is-even-square-ℕ n (is-even-div-4-ℕ (square-ℕ n) H) +``` diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 8329da6221..d78891eed1 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -72,14 +72,22 @@ square-root-ℕ _ (root , _) = root ### Squares of successors -For any `n` we have `(n + 1)² = (n + 2)n + 1 +For any `n` we have the equations + +```text + (n + 1)² = (n + 2)n + 1 + (n + 1)² = n² + 2n + 1. +``` ```agda -square-succ-ℕ : - (n : ℕ) → - square-ℕ (succ-ℕ n) = succ-ℕ ((succ-ℕ (succ-ℕ n)) *ℕ n) -square-succ-ℕ n = +square-succ-ℕ' : + (n : ℕ) → square-ℕ (succ-ℕ n) = (n +ℕ 2) *ℕ n +ℕ 1 +square-succ-ℕ' n = right-successor-law-mul-ℕ (succ-ℕ n) n + +square-succ-ℕ : (n : ℕ) → square-ℕ (succ-ℕ n) = square-ℕ n +ℕ 2 *ℕ n +ℕ 1 +square-succ-ℕ n = + square-succ-ℕ' n ∙ ap succ-ℕ (right-distributive-mul-add-ℕ n 2 n) ``` ### Squares of double successors @@ -134,7 +142,8 @@ preserves-strict-order-square-ℕ m n H = ### The square function reflects the order on the natural numbers -For any two natural numbers `m` and `n` we have `m² ≤ n² → m ≤ n` and `m² < n² → m < n`. +For any two natural numbers `m` and `n` we have `m² ≤ n² → m ≤ n` and +`m² < n² → m < n`. ```agda reflects-order-square-ℕ : @@ -206,7 +215,7 @@ lower-bound-square-ℕ (succ-ℕ n) = ( leq-add-ℕ (square-ℕ n +ℕ n) n) ( leq-add-ℕ (square-ℕ n) n)) ( lower-bound-square-ℕ n)) - ( inv (square-succ-ℕ n)) + ( inv (square-succ-ℕ' n)) ``` ### If a number `n` has a square root, then the square root is at most `n` @@ -238,7 +247,7 @@ strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) ( le-add-succ-ℕ (square-ℕ (n +ℕ 2)) (n +ℕ 1)) ( le-add-succ-ℕ (square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) (n +ℕ 1)))) - ( inv (square-succ-ℕ (succ-ℕ (succ-ℕ n)))) + ( inv (square-succ-ℕ' (succ-ℕ (succ-ℕ n)))) ``` ### If a number `n` greater than 1 has a square root, then its square root is strictly smaller than `n` @@ -312,6 +321,7 @@ is-decidable-is-square-ℕ n = ### Equivalent characterizations of a number being even in terms of its square Consider a natural number `n`. The following are equivalent: + - The number `n` is even. - Its square is even. - Its square is divisible by 4. diff --git a/src/elementary-number-theory/unit-elements-standard-finite-types.lagda.md b/src/elementary-number-theory/unit-elements-standard-finite-types.lagda.md index 4b6f36cc3b..84cbfc3a49 100644 --- a/src/elementary-number-theory/unit-elements-standard-finite-types.lagda.md +++ b/src/elementary-number-theory/unit-elements-standard-finite-types.lagda.md @@ -68,7 +68,7 @@ pr2 (is-unit-neg-one-Fin {succ-ℕ k}) = ( concatenate-eq-cong-ℕ ( succ-ℕ (succ-ℕ k)) { x3 = 1} - ( square-succ-ℕ k) + ( square-succ-ℕ' k) ( pair k ( ( commutative-mul-ℕ k (succ-ℕ (succ-ℕ k))) ∙ ( inv (right-unit-law-dist-ℕ ((succ-ℕ (succ-ℕ k)) *ℕ k)))))) From defae26555f941ed89bad672ced3bba0e257e8c9 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 26 Oct 2024 14:12:24 -0400 Subject: [PATCH 07/72] make pre-commit --- src/elementary-number-theory/integer-fractions.lagda.md | 7 +++++-- .../irrationality-square-root-of-2.lagda.md | 3 --- .../multiplication-natural-numbers.lagda.md | 2 +- .../reduced-integer-fractions.lagda.md | 4 ++-- 4 files changed, 8 insertions(+), 8 deletions(-) diff --git a/src/elementary-number-theory/integer-fractions.lagda.md b/src/elementary-number-theory/integer-fractions.lagda.md index 729a763f39..21c5a49459 100644 --- a/src/elementary-number-theory/integer-fractions.lagda.md +++ b/src/elementary-number-theory/integer-fractions.lagda.md @@ -136,13 +136,16 @@ fraction-ℤ-Set = fraction-ℤ , is-set-fraction-ℤ ### The standard equivalence relation on integer fractions -Two integer fractions `a/b` and `c/d` are said to be equivalent if they satisfy the equation +Two integer fractions `a/b` and `c/d` are said to be equivalent if they satisfy +the equation ```text ad = cb. ``` -This relation is obviously reflexive and symmetric. To see that it is transitive, suppose that `a/b ~ c/d` and that `c/d ~ e/f`. Since `d` is positive, we have the implication +This relation is obviously reflexive and symmetric. To see that it is +transitive, suppose that `a/b ~ c/d` and that `c/d ~ e/f`. Since `d` is +positive, we have the implication ```text (af)d = (eb)d → af = eb. diff --git a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md index cf9d371555..8e93d5d5fd 100644 --- a/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md +++ b/src/elementary-number-theory/irrationality-square-root-of-2.lagda.md @@ -62,6 +62,3 @@ not-two-square-reduced-fraction-ℤ : not-two-square-reduced-fraction-ℤ x@(a , b , pb) H K = {!!} ``` - -Id (mul-ℤ (mul-ℤ a a) (inr (inr 0))) -(add-ℤ (mul-ℤ b b) (mul-ℤ b b)) diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index 3a9eb6367b..d28741d96b 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -279,7 +279,7 @@ is-nonzero-right-factor-mul-ℕ : is-nonzero-right-factor-mul-ℕ x .zero-ℕ H refl = H (right-zero-law-mul-ℕ x) ``` -### If $(x+1)y = x+1$ then it follows that $y = 1$. +### If $(x+1)y = x+1$ then it follows that $y = 1$ ```agda is-one-is-right-unit-mul-ℕ : diff --git a/src/elementary-number-theory/reduced-integer-fractions.lagda.md b/src/elementary-number-theory/reduced-integer-fractions.lagda.md index c5e38528cf..5882c3391d 100644 --- a/src/elementary-number-theory/reduced-integer-fractions.lagda.md +++ b/src/elementary-number-theory/reduced-integer-fractions.lagda.md @@ -39,8 +39,8 @@ open import foundation.universe-levels ## Idea A {{#concept "reduced fraction" Agda=reduced-fraction-ℤ}} is a -[fraction](elementary-number-theory.integer-fractions.md) in which its -numerator and denominator are +[fraction](elementary-number-theory.integer-fractions.md) in which its numerator +and denominator are [coprime](elementary-number-theory.relatively-prime-integers.md). ## Definitions From 9b917e721325e71af9cdbc63424f2a485b36956b Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 26 Oct 2024 16:55:42 -0400 Subject: [PATCH 08/72] parity integers --- src/elementary-number-theory.lagda.md | 1 + .../parity-integers.lagda.md | 188 ++++++++++++++++++ .../parity-natural-numbers.lagda.md | 13 +- .../reduced-integer-fractions.lagda.md | 11 +- .../squares-integers.lagda.md | 67 ++++--- 5 files changed, 244 insertions(+), 36 deletions(-) create mode 100644 src/elementary-number-theory/parity-integers.lagda.md diff --git a/src/elementary-number-theory.lagda.md b/src/elementary-number-theory.lagda.md index a479a091f0..edc4629d37 100644 --- a/src/elementary-number-theory.lagda.md +++ b/src/elementary-number-theory.lagda.md @@ -113,6 +113,7 @@ open import elementary-number-theory.nonzero-integers public open import elementary-number-theory.nonzero-natural-numbers public open import elementary-number-theory.nonzero-rational-numbers public open import elementary-number-theory.ordinal-induction-natural-numbers public +open import elementary-number-theory.parity-integers public open import elementary-number-theory.parity-natural-numbers public open import elementary-number-theory.peano-arithmetic public open import elementary-number-theory.pisano-periods public diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md new file mode 100644 index 0000000000..cd08a09f2a --- /dev/null +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -0,0 +1,188 @@ +# Parity of the integers + +```agda +module elementary-number-theory.parity-integers where +``` + +
Imports + +```agda +open import elementary-number-theory.divisibility-integers +open import elementary-number-theory.integers + +open import foundation.negation +open import foundation.universe-levels +``` + +
+ +## Idea + +{{#concept "Parity" WD="parity" WDID=Q141160}} +[partitions](foundation.partitions.md) the +[integers](elementary-number-theory.integers.md) into two +[classes](foundation.equivalence-relations.md): the +{{#concept "even" Disambiguation="integer" Agda=is-even-ℤ WD="even number" WDID=Q13366104}} +and the +{{#concept "odd" Disambiguation="integer" Agda=is-odd-ℤ WD="odd number" WDID=Q13366129}} +integers. Even integers are those that are +[divisible](elementary-number-theory.divisibility-integers.md) by two, and odd +integers are those that aren't. + +## Definitions + +### Even and odd integers + +```agda +is-even-ℤ : ℤ → UU lzero +is-even-ℤ a = div-ℤ (int-ℕ 2) a + +is-odd-ℤ : ℤ → UU lzero +is-odd-ℤ a = ¬ (is-even-ℤ a) +``` + +## Properties + +### Being even or odd is decidable + +```text +is-decidable-is-even-ℕ : (x : ℕ) → is-decidable (is-even-ℕ x) +is-decidable-is-even-ℕ x = is-decidable-div-ℕ 2 x + +is-decidable-is-odd-ℕ : (x : ℕ) → is-decidable (is-odd-ℕ x) +is-decidable-is-odd-ℕ x = is-decidable-neg (is-decidable-is-even-ℕ x) +``` + +### `0` is an even natural number + +```text +is-even-zero-ℕ : is-even-ℕ 0 +is-even-zero-ℕ = div-zero-ℕ 2 + +is-odd-one-ℕ : is-odd-ℕ 1 +is-odd-one-ℕ H = Eq-eq-ℕ (is-one-div-one-ℕ 2 H) +``` + +### A natural number `x` is even if and only if `x + 2` is even + +```text +is-even-is-even-succ-succ-ℕ : + (n : ℕ) → is-even-ℕ (succ-ℕ (succ-ℕ n)) → is-even-ℕ n +pr1 (is-even-is-even-succ-succ-ℕ n (succ-ℕ d , p)) = d +pr2 (is-even-is-even-succ-succ-ℕ n (succ-ℕ d , p)) = + is-injective-succ-ℕ (is-injective-succ-ℕ p) + +is-even-succ-succ-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → is-even-ℕ (succ-ℕ (succ-ℕ n)) +pr1 (is-even-succ-succ-is-even-ℕ n (d , p)) = succ-ℕ d +pr2 (is-even-succ-succ-is-even-ℕ n (d , p)) = ap (succ-ℕ ∘ succ-ℕ) p +``` + +### A natural number `x` is odd if and only if `x + 2` is odd + +```text +is-odd-is-odd-succ-succ-ℕ : + (n : ℕ) → is-odd-ℕ (succ-ℕ (succ-ℕ n)) → is-odd-ℕ n +is-odd-is-odd-succ-succ-ℕ n = map-neg (is-even-succ-succ-is-even-ℕ n) + +is-odd-succ-succ-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → is-odd-ℕ (succ-ℕ (succ-ℕ n)) +is-odd-succ-succ-is-odd-ℕ n = map-neg (is-even-is-even-succ-succ-ℕ n) +``` + +### If a natural number `x` is odd, then `x + 1` is even + +```text +is-even-succ-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → is-even-ℕ (succ-ℕ n) +is-even-succ-is-odd-ℕ zero-ℕ p = ex-falso (p is-even-zero-ℕ) +is-even-succ-is-odd-ℕ (succ-ℕ zero-ℕ) p = (1 , refl) +is-even-succ-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = + is-even-succ-succ-is-even-ℕ + ( succ-ℕ n) + ( is-even-succ-is-odd-ℕ n (is-odd-is-odd-succ-succ-ℕ n p)) +``` + +### If a natural number `x` is even, then `x + 1` is odd + +```text +is-odd-succ-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → is-odd-ℕ (succ-ℕ n) +is-odd-succ-is-even-ℕ zero-ℕ p = is-odd-one-ℕ +is-odd-succ-is-even-ℕ (succ-ℕ zero-ℕ) p = ex-falso (is-odd-one-ℕ p) +is-odd-succ-is-even-ℕ (succ-ℕ (succ-ℕ n)) p = + is-odd-succ-succ-is-odd-ℕ + ( succ-ℕ n) + ( is-odd-succ-is-even-ℕ n (is-even-is-even-succ-succ-ℕ n p)) +``` + +### If a natural number `x + 1` is odd, then `x` is even + +```text +is-even-is-odd-succ-ℕ : + (n : ℕ) → is-odd-ℕ (succ-ℕ n) → is-even-ℕ n +is-even-is-odd-succ-ℕ n p = + is-even-is-even-succ-succ-ℕ + ( n) + ( is-even-succ-is-odd-ℕ (succ-ℕ n) p) +``` + +### If a natural number `x + 1` is even, then `x` is odd + +```text +is-odd-is-even-succ-ℕ : + (n : ℕ) → is-even-ℕ (succ-ℕ n) → is-odd-ℕ n +is-odd-is-even-succ-ℕ n p = + is-odd-is-odd-succ-succ-ℕ + ( n) + ( is-odd-succ-is-even-ℕ (succ-ℕ n) p) +``` + +### A natural number `x` is odd if and only if there is a natural number `y` such that `succ-ℕ (y *ℕ 2) = x` + +```text +has-odd-expansion : ℕ → UU lzero +has-odd-expansion x = Σ ℕ (λ y → (succ-ℕ (y *ℕ 2)) = x) + +is-odd-has-odd-expansion : (n : ℕ) → has-odd-expansion n → is-odd-ℕ n +is-odd-has-odd-expansion .(succ-ℕ (m *ℕ 2)) (m , refl) = + is-odd-succ-is-even-ℕ (m *ℕ 2) (m , refl) + +has-odd-expansion-is-odd : (n : ℕ) → is-odd-ℕ n → has-odd-expansion n +has-odd-expansion-is-odd zero-ℕ p = ex-falso (p is-even-zero-ℕ) +has-odd-expansion-is-odd (succ-ℕ zero-ℕ) p = 0 , refl +has-odd-expansion-is-odd (succ-ℕ (succ-ℕ n)) p = + ( succ-ℕ (pr1 s)) , ap (succ-ℕ ∘ succ-ℕ) (pr2 s) + where + s : has-odd-expansion n + s = has-odd-expansion-is-odd n (is-odd-is-odd-succ-succ-ℕ n p) +``` + +### A number is even if and only if it is divisible by an even number + +```text +is-even-div-is-even-ℕ : + (n m : ℕ) → is-even-ℕ m → div-ℕ m n → is-even-ℕ n +is-even-div-is-even-ℕ ._ ._ (m' , refl) (k , refl) = + k *ℕ m' , associative-mul-ℕ k m' 2 + +is-even-div-4-ℕ : + (n : ℕ) → div-ℕ 4 n → is-even-ℕ n +is-even-div-4-ℕ n = is-even-div-is-even-ℕ n 4 (2 , refl) +``` + +### If any two out of `x`, `y`, and `x + y` are even, then so is the third + +```text +is-even-add-ℕ : + (m n : ℕ) → is-even-ℕ m → is-even-ℕ n → is-even-ℕ (add-ℕ m n) +is-even-add-ℕ = div-add-ℕ 2 + +is-even-left-summand-ℕ : + (m n : ℕ) → is-even-ℕ n → is-even-ℕ (add-ℕ m n) → is-even-ℕ m +is-even-left-summand-ℕ = div-left-summand-ℕ 2 + +is-even-right-summand-ℕ : + (m n : ℕ) → is-even-ℕ m → is-even-ℕ (add-ℕ m n) → is-even-ℕ n +is-even-right-summand-ℕ = div-right-summand-ℕ 2 +``` diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 1ae412fef7..7e144ed5ff 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -28,9 +28,16 @@ open import foundation.universe-levels ## Idea -Parity partitions the natural numbers into two classes: the even and the odd -natural numbers. Even natural numbers are those that are divisible by two, and -odd natural numbers are those that aren't. +{{#concept "Parity" WD="parity" WDID=Q141160}} +[partitions](foundation.partitions.md) the [natural +numbers](elementary-number-theory.natural numbers.md) into two +[classes](foundation.equivalence-relations.md): the +{{#concept "even" Disambiguation="natural number" Agda=is-even-ℤ WD="even number" WDID=Q13366104}} +and the +{{#concept "odd" Disambiguation="natural number" Agda=is-odd-ℤ WD="odd number" WDID=Q13366129}} +natural numbers. Even natural numbers are those that are +[divisible](elementary-number-theory.divisibility-natural numbers.md) by two, +and odd natural numbers are those that aren't. ## Definition diff --git a/src/elementary-number-theory/reduced-integer-fractions.lagda.md b/src/elementary-number-theory/reduced-integer-fractions.lagda.md index 5882c3391d..aa543cd985 100644 --- a/src/elementary-number-theory/reduced-integer-fractions.lagda.md +++ b/src/elementary-number-theory/reduced-integer-fractions.lagda.md @@ -38,10 +38,13 @@ open import foundation.universe-levels ## Idea -A {{#concept "reduced fraction" Agda=reduced-fraction-ℤ}} is a +A {{#concept "reduced fraction" Agda=is-reduced-fraction-ℤ}} is a [fraction](elementary-number-theory.integer-fractions.md) in which its numerator and denominator are -[coprime](elementary-number-theory.relatively-prime-integers.md). +[coprime](elementary-number-theory.relatively-prime-integers.md). In the module +[`elementary-number-theory.rational-numbers`](elementary-number-theory.rational-numbers.md) +we will define the type `ℚ` of rational numbers to be the type of reduced +integer fractions. ## Definitions @@ -582,3 +585,7 @@ unique-reduce-fraction-ℤ x y H = ( eq-is-prop ( is-prop-is-positive-ℤ (int-reduce-denominator-fraction-ℤ y)))))) ``` + +## See also + +- [The type of rational numbers](elementary-number-theory.rational-numbers.md) diff --git a/src/elementary-number-theory/squares-integers.lagda.md b/src/elementary-number-theory/squares-integers.lagda.md index 67849e77ae..989d98a3a6 100644 --- a/src/elementary-number-theory/squares-integers.lagda.md +++ b/src/elementary-number-theory/squares-integers.lagda.md @@ -107,6 +107,14 @@ is-nonnegative-is-square-ℤ : {a : ℤ} → is-square-ℤ a → is-nonnegative- is-nonnegative-is-square-ℤ (r , refl) = is-nonnegative-square-ℤ r ``` +### The square of an integer is equal to the square of its negative + +```agda +compute-square-neg-ℤ : + (a : ℤ) → square-ℤ a = square-ℤ (neg-ℤ a) +compute-square-neg-ℤ a = inv (double-negative-law-mul-ℤ a a) +``` + ### The squares in ℤ are exactly the squares in ℕ ```agda @@ -209,38 +217,35 @@ Consider a integer `n`. The following are equivalent: - Its square is divisible by 4. ```text -div-four-square-is-even-ℕ : - (n : ℕ) → is-even-ℕ n → div-ℕ 4 (square-ℕ n) -pr1 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = - square-ℕ m -pr2 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = - inv (distributive-square-mul-ℕ m 2) - -is-even-square-is-even-ℕ : - (n : ℕ) → is-even-ℕ n → is-even-ℕ (square-ℕ n) -is-even-square-is-even-ℕ .(m *ℕ 2) (m , refl) = - is-even-div-4-ℕ _ (div-four-square-is-even-ℕ (m *ℕ 2) (m , refl)) - -is-even-is-even-square-ℕ : - (n : ℕ) → is-even-ℕ (square-ℕ n) → is-even-ℕ n -is-even-is-even-square-ℕ zero-ℕ H = is-even-zero-ℕ -is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (zero-ℕ , ()) -is-even-is-even-square-ℕ (succ-ℕ zero-ℕ) (succ-ℕ k , ()) -is-even-is-even-square-ℕ (succ-ℕ (succ-ℕ n)) (m , p) = - is-even-succ-succ-is-even-ℕ n - ( is-even-is-even-square-ℕ n - ( is-even-left-summand-ℕ - ( square-ℕ n) - ( 4 *ℕ n) - ( is-even-div-4-ℕ (4 *ℕ n) (n , commutative-mul-ℕ n 4)) - ( is-even-left-summand-ℕ - ( square-ℕ n +ℕ 4 *ℕ n) +div-four-square-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → div-ℤ 4 (square-ℤ a) +dif-four-square-is-even-ℤ = ? + +is-even-square-is-even-ℤ : + (n : ℤ) → is-even-ℤ n → is-even-ℤ (square-ℤ n) +is-even-square-is-even-ℤ .(m *ℤ 2) (m , refl) = + is-even-div-4-ℤ _ (div-four-square-is-even-ℤ (m *ℤ 2) (m , refl)) + +is-even-is-even-square-ℤ : + (n : ℤ) → is-even-ℤ (square-ℤ n) → is-even-ℤ n +is-even-is-even-square-ℤ zero-ℤ H = is-even-zero-ℤ +is-even-is-even-square-ℤ (succ-ℤ zero-ℤ) (zero-ℤ , ()) +is-even-is-even-square-ℤ (succ-ℤ zero-ℤ) (succ-ℤ k , ()) +is-even-is-even-square-ℤ (succ-ℤ (succ-ℤ n)) (m , p) = + is-even-succ-succ-is-even-ℤ n + ( is-even-is-even-square-ℤ n + ( is-even-left-summand-ℤ + ( square-ℤ n) + ( 4 *ℤ n) + ( is-even-div-4-ℤ (4 *ℤ n) (n , commutative-mul-ℤ n 4)) + ( is-even-left-summand-ℤ + ( square-ℤ n +ℤ 4 *ℤ n) ( 4) ( 2 , refl) - ( m , p ∙ square-succ-succ-ℕ n)))) + ( m , p ∙ square-succ-succ-ℤ n)))) -is-even-div-four-square-ℕ : - (n : ℕ) → div-ℕ 4 (square-ℕ n) → is-even-ℕ n -is-even-div-four-square-ℕ n H = - is-even-is-even-square-ℕ n (is-even-div-4-ℕ (square-ℕ n) H) +is-even-div-four-square-ℤ : + (n : ℤ) → div-ℤ 4 (square-ℤ n) → is-even-ℤ n +is-even-div-four-square-ℤ n H = + is-even-is-even-square-ℤ n (is-even-div-4-ℤ (square-ℤ n) H) ``` From 01b1c33b87b264fdd77fc96cf4faf6bf7e188a7d Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 26 Oct 2024 18:03:59 -0400 Subject: [PATCH 09/72] parity integers --- .../divisibility-integers.lagda.md | 44 ++++-- .../parity-integers.lagda.md | 147 +++++++++++++----- .../parity-natural-numbers.lagda.md | 4 + 3 files changed, 149 insertions(+), 46 deletions(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index bed59408c3..4b1fd94b37 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -202,16 +202,6 @@ is-zero-is-zero-div-ℤ : (x k : ℤ) → div-ℤ k x → is-zero-ℤ k → is-z is-zero-is-zero-div-ℤ x .zero-ℤ k-div-x refl = is-zero-div-zero-ℤ x k-div-x ``` -### If `x` divides both `y` and `z`, then it divides `y + z` - -```agda -div-add-ℤ : (x y z : ℤ) → div-ℤ x y → div-ℤ x z → div-ℤ x (y +ℤ z) -pr1 (div-add-ℤ x y z (pair d p) (pair e q)) = d +ℤ e -pr2 (div-add-ℤ x y z (pair d p) (pair e q)) = - ( right-distributive-mul-add-ℤ d e x) ∙ - ( ap-add-ℤ p q) -``` - ### If `x` divides `y` then `x` divides any multiple of `y` ```agda @@ -246,6 +236,40 @@ pr2 (neg-div-ℤ x y (pair d p)) = by p ``` +### If `x` divides both `y` and `z`, then it divides `y + z` + +```agda +div-add-ℤ : (x y z : ℤ) → div-ℤ x y → div-ℤ x z → div-ℤ x (y +ℤ z) +pr1 (div-add-ℤ x y z (pair d p) (pair e q)) = d +ℤ e +pr2 (div-add-ℤ x y z (pair d p) (pair e q)) = + ( right-distributive-mul-add-ℤ d e x) ∙ + ( ap-add-ℤ p q) + +div-right-summand-ℤ : + (x y z : ℤ) → div-ℤ x y → div-ℤ x (y +ℤ z) → div-ℤ x z +div-right-summand-ℤ x y z H K = + tr + ( div-ℤ x) + ( is-retraction-left-add-neg-ℤ y z) + ( div-add-ℤ x + ( neg-ℤ y) + ( y +ℤ z) + ( div-neg-ℤ x y H) + ( K)) + +div-left-summand-ℤ : + (x y z : ℤ) → div-ℤ x z → div-ℤ x (y +ℤ z) → div-ℤ x y +div-left-summand-ℤ x y z H K = + tr + ( div-ℤ x) + ( is-retraction-right-add-neg-ℤ z y) + ( div-add-ℤ x + ( y +ℤ z) + ( neg-ℤ z) + ( K) + ( div-neg-ℤ x z H)) +``` + ### Multiplication preserves divisibility ```agda diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index cd08a09f2a..ca96303564 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -7,9 +7,14 @@ module elementary-number-theory.parity-integers where
Imports ```agda +open import elementary-number-theory.absolute-value-integers +open import elementary-number-theory.addition-integers open import elementary-number-theory.divisibility-integers open import elementary-number-theory.integers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers +open import foundation.decidable-types open import foundation.negation open import foundation.universe-levels ``` @@ -43,54 +48,124 @@ is-odd-ℤ a = ¬ (is-even-ℤ a) ## Properties -### Being even or odd is decidable +### A natural number is even if and only if it is even as an integer -```text -is-decidable-is-even-ℕ : (x : ℕ) → is-decidable (is-even-ℕ x) -is-decidable-is-even-ℕ x = is-decidable-div-ℕ 2 x +```agda +is-even-int-is-even-ℕ : (n : ℕ) → is-even-ℕ n → is-even-ℤ (int-ℕ n) +is-even-int-is-even-ℕ n H = div-int-div-ℕ H -is-decidable-is-odd-ℕ : (x : ℕ) → is-decidable (is-odd-ℕ x) -is-decidable-is-odd-ℕ x = is-decidable-neg (is-decidable-is-even-ℕ x) +is-even-is-even-int-ℕ : (n : ℕ) → is-even-ℤ (int-ℕ n) → is-even-ℕ n +is-even-is-even-int-ℕ n H = div-div-int-ℕ H ``` -### `0` is an even natural number +### A natural number is odd if and only if it is odd as an integer -```text -is-even-zero-ℕ : is-even-ℕ 0 -is-even-zero-ℕ = div-zero-ℕ 2 +```agda +is-odd-int-is-odd-ℕ : (n : ℕ) → is-odd-ℕ n → is-odd-ℤ (int-ℕ n) +is-odd-int-is-odd-ℕ n H K = H (is-even-is-even-int-ℕ n K) -is-odd-one-ℕ : is-odd-ℕ 1 -is-odd-one-ℕ H = Eq-eq-ℕ (is-one-div-one-ℕ 2 H) +is-odd-is-odd-int-ℕ : (n : ℕ) → is-odd-ℤ (int-ℕ n) → is-odd-ℕ n +is-odd-is-odd-int-ℕ n H K = H (is-even-int-is-even-ℕ n K) ``` -### A natural number `x` is even if and only if `x + 2` is even +### An integer is even if and only if its absolute value is an even integer -```text -is-even-is-even-succ-succ-ℕ : - (n : ℕ) → is-even-ℕ (succ-ℕ (succ-ℕ n)) → is-even-ℕ n -pr1 (is-even-is-even-succ-succ-ℕ n (succ-ℕ d , p)) = d -pr2 (is-even-is-even-succ-succ-ℕ n (succ-ℕ d , p)) = - is-injective-succ-ℕ (is-injective-succ-ℕ p) +```agda +is-even-int-abs-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → is-even-ℤ (int-abs-ℤ a) +is-even-int-abs-is-even-ℤ a = + div-sim-unit-ℤ + ( refl-sim-unit-ℤ (int-ℕ 2)) + ( symmetric-sim-unit-ℤ (int-abs-ℤ a) a (sim-unit-abs-ℤ a)) + +is-even-is-even-int-abs-ℤ : + (a : ℤ) → is-even-ℤ (int-abs-ℤ a) → is-even-ℤ a +is-even-is-even-int-abs-ℤ a = + div-sim-unit-ℤ + ( refl-sim-unit-ℤ (int-ℕ 2)) + ( sim-unit-abs-ℤ a) + +is-even-abs-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → is-even-ℕ (abs-ℤ a) +is-even-abs-is-even-ℤ a H = + is-even-is-even-int-ℕ (abs-ℤ a) (is-even-int-abs-is-even-ℤ a H) + +is-even-is-even-abs-ℤ : + (a : ℤ) → is-even-ℕ (abs-ℤ a) → is-even-ℤ a +is-even-is-even-abs-ℤ a H = + is-even-is-even-int-abs-ℤ a (is-even-int-is-even-ℕ (abs-ℤ a) H) +``` -is-even-succ-succ-is-even-ℕ : - (n : ℕ) → is-even-ℕ n → is-even-ℕ (succ-ℕ (succ-ℕ n)) -pr1 (is-even-succ-succ-is-even-ℕ n (d , p)) = succ-ℕ d -pr2 (is-even-succ-succ-is-even-ℕ n (d , p)) = ap (succ-ℕ ∘ succ-ℕ) p +### An integer is odd if and only if its absolute value is an odd integer + +```agda +is-odd-int-abs-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → is-odd-ℤ (int-abs-ℤ a) +is-odd-int-abs-is-odd-ℤ a H K = H (is-even-is-even-int-abs-ℤ a K) + +is-odd-is-odd-int-abs-ℤ : (a : ℤ) → is-odd-ℤ (int-abs-ℤ a) → is-odd-ℤ a +is-odd-is-odd-int-abs-ℤ a H K = H (is-even-int-abs-is-even-ℤ a K) + +is-odd-abs-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → is-odd-ℕ (abs-ℤ a) +is-odd-abs-is-odd-ℤ a H K = H (is-even-is-even-abs-ℤ a K) + +is-odd-is-odd-abs-ℤ : (a : ℤ) → is-odd-ℕ (abs-ℤ a) → is-odd-ℤ a +is-odd-is-odd-abs-ℤ a H K = H (is-even-abs-is-even-ℤ a K) ``` -### A natural number `x` is odd if and only if `x + 2` is odd +### Being even or odd is decidable -```text -is-odd-is-odd-succ-succ-ℕ : - (n : ℕ) → is-odd-ℕ (succ-ℕ (succ-ℕ n)) → is-odd-ℕ n -is-odd-is-odd-succ-succ-ℕ n = map-neg (is-even-succ-succ-is-even-ℕ n) +```agda +is-decidable-is-even-ℤ : + (a : ℤ) → is-decidable (is-even-ℤ a) +is-decidable-is-even-ℤ a = + is-decidable-iff + ( is-even-is-even-abs-ℤ a) + ( is-even-abs-is-even-ℤ a) + ( is-decidable-is-even-ℕ (abs-ℤ a)) + +is-decidable-is-odd-ℤ : (a : ℤ) → is-decidable (is-odd-ℤ a) +is-decidable-is-odd-ℤ a = is-decidable-neg (is-decidable-is-even-ℤ a) +``` + +### `0` is an even integer + +```agda +is-even-zero-ℤ : is-even-ℤ (int-ℕ 0) +is-even-zero-ℤ = is-even-int-is-even-ℕ 0 is-even-zero-ℕ +``` + +### `1` is an odd integer + +```agda +is-odd-one-ℤ : is-odd-ℤ (int-ℕ 1) +is-odd-one-ℤ = is-odd-int-is-odd-ℕ 1 is-odd-one-ℕ +``` + +### A integer `x` is even if and only if `x + 2` is even + +```agda +is-even-is-even-add-two-ℤ : + (a : ℤ) → is-even-ℤ (a +ℤ int-ℕ 2) → is-even-ℤ a +is-even-is-even-add-two-ℤ a = + div-left-summand-ℤ (int-ℕ 2) a (int-ℕ 2) (refl-div-ℤ (int-ℕ 2)) + +is-even-add-two-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → is-even-ℤ (a +ℤ int-ℕ 2) +is-even-add-two-is-even-ℤ a H = + div-add-ℤ (int-ℕ 2) a (int-ℕ 2) H (refl-div-ℤ (int-ℕ 2)) +``` + +### A integer `x` is odd if and only if `x + 2` is odd + +```agda +is-odd-is-odd-add-two-ℤ : (a : ℤ) → is-odd-ℤ (a +ℤ int-ℕ 2) → is-odd-ℤ a +is-odd-is-odd-add-two-ℤ a H K = H (is-even-add-two-is-even-ℤ a K) -is-odd-succ-succ-is-odd-ℕ : - (n : ℕ) → is-odd-ℕ n → is-odd-ℕ (succ-ℕ (succ-ℕ n)) -is-odd-succ-succ-is-odd-ℕ n = map-neg (is-even-is-even-succ-succ-ℕ n) +is-odd-add-two-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → is-odd-ℤ (a +ℤ int-ℕ 2) +is-odd-add-two-is-odd-ℤ a H K = H (is-even-is-even-add-two-ℤ a K) ``` -### If a natural number `x` is odd, then `x + 1` is even +### If a integer `x` is odd, then `x + 1` is even ```text is-even-succ-is-odd-ℕ : @@ -103,7 +178,7 @@ is-even-succ-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = ( is-even-succ-is-odd-ℕ n (is-odd-is-odd-succ-succ-ℕ n p)) ``` -### If a natural number `x` is even, then `x + 1` is odd +### If a integer `x` is even, then `x + 1` is odd ```text is-odd-succ-is-even-ℕ : @@ -116,7 +191,7 @@ is-odd-succ-is-even-ℕ (succ-ℕ (succ-ℕ n)) p = ( is-odd-succ-is-even-ℕ n (is-even-is-even-succ-succ-ℕ n p)) ``` -### If a natural number `x + 1` is odd, then `x` is even +### If a integer `x + 1` is odd, then `x` is even ```text is-even-is-odd-succ-ℕ : @@ -127,7 +202,7 @@ is-even-is-odd-succ-ℕ n p = ( is-even-succ-is-odd-ℕ (succ-ℕ n) p) ``` -### If a natural number `x + 1` is even, then `x` is odd +### If a integer `x + 1` is even, then `x` is odd ```text is-odd-is-even-succ-ℕ : @@ -138,7 +213,7 @@ is-odd-is-even-succ-ℕ n p = ( is-odd-succ-is-even-ℕ (succ-ℕ n) p) ``` -### A natural number `x` is odd if and only if there is a natural number `y` such that `succ-ℕ (y *ℕ 2) = x` +### A integer `x` is odd if and only if there is a integer `y` such that `succ-ℕ (y *ℕ 2) = x` ```text has-odd-expansion : ℕ → UU lzero diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 7e144ed5ff..e2e0de1b25 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -66,7 +66,11 @@ is-decidable-is-odd-ℕ x = is-decidable-neg (is-decidable-is-even-ℕ x) ```agda is-even-zero-ℕ : is-even-ℕ 0 is-even-zero-ℕ = div-zero-ℕ 2 +``` + +### `1` is an odd natural number +```agda is-odd-one-ℕ : is-odd-ℕ 1 is-odd-one-ℕ H = Eq-eq-ℕ (is-one-div-one-ℕ 2 H) ``` From 70edf996e689dcd08b21f7eea1c7ca162c9b1211 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 27 Oct 2024 15:32:32 -0400 Subject: [PATCH 10/72] work --- src/elementary-number-theory/divisibility-integers.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 4b1fd94b37..d00f822417 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -236,7 +236,7 @@ pr2 (neg-div-ℤ x y (pair d p)) = by p ``` -### If `x` divides both `y` and `z`, then it divides `y + z` +### If `x` divides two out of the three integers `y`, `z`, and `y + z`, then it divides the third ```agda div-add-ℤ : (x y z : ℤ) → div-ℤ x y → div-ℤ x z → div-ℤ x (y +ℤ z) From e55cb69ed6cbdf825733ac288cfb9c7ee12b4add Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 28 Oct 2024 09:32:04 -0400 Subject: [PATCH 11/72] work --- src/elementary-number-theory/absolute-value-integers.lagda.md | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/elementary-number-theory/absolute-value-integers.lagda.md b/src/elementary-number-theory/absolute-value-integers.lagda.md index 9ff83babef..949bcc95db 100644 --- a/src/elementary-number-theory/absolute-value-integers.lagda.md +++ b/src/elementary-number-theory/absolute-value-integers.lagda.md @@ -31,7 +31,9 @@ open import foundation.unit-type The {{#concept "absolute value" Disambiguation="of an integer" Agda=abs-ℤ}} of an integer is the natural number with the same distance from `0`. -## Definition +## Definitions + +### The absolute value of an integer ```agda abs-ℤ : ℤ → ℕ From a12e652cd07d161742af9b3e2f5055e95802d804 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 29 Oct 2024 17:26:11 -0400 Subject: [PATCH 12/72] generalizations to commutative rings --- ...groups-of-units-commutative-rings.lagda.md | 30 ++- ...rtible-elements-commutative-rings.lagda.md | 59 +++++- .../divisibility-integers.lagda.md | 162 ++-------------- .../greatest-common-divisor-integers.lagda.md | 8 +- .../modular-arithmetic.lagda.md | 4 +- .../unit-integers.lagda.md | 178 ++++++++++++++++++ .../unit-similarity-integers.lagda.md | 165 ++++++++++++++++ .../invertible-elements-monoids.lagda.md | 79 ++++++-- .../groups-of-units-rings.lagda.md | 35 +++- .../invertible-elements-rings.lagda.md | 53 +++++- 10 files changed, 594 insertions(+), 179 deletions(-) create mode 100644 src/elementary-number-theory/unit-integers.lagda.md create mode 100644 src/elementary-number-theory/unit-similarity-integers.lagda.md diff --git a/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md b/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md index f3fcd5aac1..78457c72b9 100644 --- a/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md +++ b/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md @@ -11,6 +11,7 @@ open import category-theory.functors-large-precategories open import commutative-algebra.commutative-rings open import commutative-algebra.homomorphisms-commutative-rings +open import commutative-algebra.invertible-elements-commutative-rings open import commutative-algebra.precategory-of-commutative-rings open import foundation.dependent-pair-types @@ -35,7 +36,7 @@ open import ring-theory.groups-of-units-rings ## Idea -The **group of units** of a +The {{#concept "group of units" Disambiguation="commutative ring" WD="unit" WDID=Q118084}} of a [commutative ring](commutative-algebra.commutative-rings.md) `A` is the [abelian group](group-theory.abelian-groups.md) consisting of all the [invertible elements](commutative-algebra.invertible-elements-commutative-rings.md) @@ -177,6 +178,14 @@ module _ inclusion-group-of-units-Commutative-Ring = inclusion-group-of-units-Ring (ring-Commutative-Ring A) + is-invertible-element-inclusion-group-of-units-Commutative-Ring : + (x : type-group-of-units-Commutative-Ring) → + is-invertible-element-Commutative-Ring A + ( inclusion-group-of-units-Commutative-Ring x) + is-invertible-element-inclusion-group-of-units-Commutative-Ring = + is-invertible-element-inclusion-group-of-units-Ring + ( ring-Commutative-Ring A) + preserves-mul-inclusion-group-of-units-Commutative-Ring : {x y : type-group-of-units-Commutative-Ring} → inclusion-group-of-units-Commutative-Ring @@ -324,3 +333,22 @@ preserves-id-functor-Large-Precategory group-of-units-commutative-ring-functor-Large-Precategory {X = A} = preserves-id-hom-group-of-units-hom-Commutative-Ring A ``` + +### Negatives of units + +```agda +module _ + {l : Level} (A : Commutative-Ring l) + where + + neg-group-of-units-Commutative-Ring : + type-group-of-units-Commutative-Ring A → + type-group-of-units-Commutative-Ring A + neg-group-of-units-Commutative-Ring = + neg-group-of-units-Ring (ring-Commutative-Ring A) + + neg-unit-group-of-units-Commutative-Ring : + type-group-of-units-Commutative-Ring A + neg-unit-group-of-units-Commutative-Ring = + neg-unit-group-of-units-Ring (ring-Commutative-Ring A) +``` diff --git a/src/commutative-algebra/invertible-elements-commutative-rings.lagda.md b/src/commutative-algebra/invertible-elements-commutative-rings.lagda.md index ebc4116027..ab8e95e1fd 100644 --- a/src/commutative-algebra/invertible-elements-commutative-rings.lagda.md +++ b/src/commutative-algebra/invertible-elements-commutative-rings.lagda.md @@ -9,6 +9,7 @@ module commutative-algebra.invertible-elements-commutative-rings where ```agda open import commutative-algebra.commutative-rings +open import foundation.action-on-identifications-functions open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.identity-types @@ -70,7 +71,7 @@ module _ ( ring-Commutative-Ring A) ``` -### Aight invertible elements of commutative rings +### Right invertible elements of commutative rings ```agda module _ @@ -241,7 +242,7 @@ module _ is-invertible-element-inverse-Ring (ring-Commutative-Ring A) ``` -### Any invertible element of a monoid has a contractible type of right inverses +### Any invertible element of a commutative ring has a contractible type of right inverses ```agda module _ @@ -255,7 +256,7 @@ module _ is-contr-is-right-invertible-element-Ring (ring-Commutative-Ring A) ``` -### Any invertible element of a monoid has a contractible type of left inverses +### Any invertible element of a commutative ring has a contractible type of left inverses ```agda module _ @@ -269,7 +270,7 @@ module _ is-contr-is-left-invertible-Ring (ring-Commutative-Ring A) ``` -### The unit of a monoid is invertible +### The unit of a commutative ring is invertible ```agda module _ @@ -290,9 +291,17 @@ module _ is-invertible-element-Commutative-Ring A (one-Commutative-Ring A) is-invertible-element-one-Commutative-Ring = is-invertible-element-one-Ring (ring-Commutative-Ring A) + + is-invertible-element-is-one-Commutative-Ring : + (x : type-Commutative-Ring A) → one-Commutative-Ring A = x → + is-invertible-element-Commutative-Ring A x + is-invertible-element-is-one-Commutative-Ring = + is-invertible-element-is-one-Ring (ring-Commutative-Ring A) ``` -### Invertible elements are closed under multiplication +### A product `xy` is invertible if and only if both `x` and `y` are invertible + +#### Invertible elements are closed under multiplication ```agda module _ @@ -324,6 +333,46 @@ module _ is-invertible-element-mul-Ring (ring-Commutative-Ring A) ``` +#### If `xy` is invertible then so is `x` + +```agda +module _ + {l : Level} (A : Commutative-Ring l) (x y : type-Commutative-Ring A) + where + + is-invertible-element-left-factor-Commutative-Ring : + is-invertible-element-Commutative-Ring A (mul-Commutative-Ring A x y) → + is-invertible-element-Commutative-Ring A x + pr1 (is-invertible-element-left-factor-Commutative-Ring (u , p , q)) = + mul-Commutative-Ring A y u + pr1 (pr2 (is-invertible-element-left-factor-Commutative-Ring (u , p , q))) = + inv (associative-mul-Commutative-Ring A x y u) ∙ p + pr2 (pr2 (is-invertible-element-left-factor-Commutative-Ring (u , p , q))) = + right-swap-mul-Commutative-Ring A y u x ∙ + ap (mul-Commutative-Ring' A u) (commutative-mul-Commutative-Ring A y x) ∙ + p +``` + +#### If `xy` is invertible then so is `y` + +```agda +module _ + {l : Level} (A : Commutative-Ring l) (x y : type-Commutative-Ring A) + where + + is-invertible-element-right-factor-Commutative-Ring : + is-invertible-element-Commutative-Ring A (mul-Commutative-Ring A x y) → + is-invertible-element-Commutative-Ring A y + pr1 (is-invertible-element-right-factor-Commutative-Ring (u , p , q)) = + mul-Commutative-Ring A u x + pr1 (pr2 (is-invertible-element-right-factor-Commutative-Ring (u , p , q))) = + left-swap-mul-Commutative-Ring A y u x ∙ + ap (mul-Commutative-Ring A u) (commutative-mul-Commutative-Ring A y x) ∙ + q + pr2 (pr2 (is-invertible-element-right-factor-Commutative-Ring (u , p , q))) = + associative-mul-Commutative-Ring A u x y ∙ q +``` + ### The inverse of an invertible element is invertible ```agda diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index d00f822417..7d06065f9d 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -95,21 +95,11 @@ concatenate-eq-div-eq-ℤ : concatenate-eq-div-eq-ℤ refl p refl = p ``` -### Unit integers - -```agda -is-unit-ℤ : ℤ → UU lzero -is-unit-ℤ x = div-ℤ x one-ℤ - -unit-ℤ : UU lzero -unit-ℤ = Σ ℤ is-unit-ℤ - -int-unit-ℤ : unit-ℤ → ℤ -int-unit-ℤ = pr1 +## Properties -is-unit-int-unit-ℤ : (x : unit-ℤ) → is-unit-ℤ (int-unit-ℤ x) -is-unit-int-unit-ℤ = pr2 +### A unit integer divides every integer +```agda div-is-unit-ℤ : (x y : ℤ) → is-unit-ℤ x → div-ℤ x y pr1 (div-is-unit-ℤ x y (pair d p)) = y *ℤ d @@ -117,21 +107,6 @@ pr2 (div-is-unit-ℤ x y (pair d p)) = associative-mul-ℤ y d x ∙ (ap (y *ℤ_) p ∙ right-unit-law-mul-ℤ y) ``` -### The equivalence relation `sim-unit-ℤ` - -We define the equivalence relation `sim-unit-ℤ` in such a way that -`sim-unit-ℤ x y` is always a proposition. - -```agda -presim-unit-ℤ : ℤ → ℤ → UU lzero -presim-unit-ℤ x y = Σ unit-ℤ (λ u → (pr1 u) *ℤ x = y) - -sim-unit-ℤ : ℤ → ℤ → UU lzero -sim-unit-ℤ x y = ¬ (is-zero-ℤ x × is-zero-ℤ y) → presim-unit-ℤ x y -``` - -## Properties - ### Divisibility by a nonzero integer is a property ```agda @@ -465,121 +440,6 @@ is-unit-right-factor-ℤ x y (pair d p) = ( pair d (ap (d *ℤ_) (commutative-mul-ℤ y x) ∙ p)) ``` -### The relations `presim-unit-ℤ` and `sim-unit-ℤ` are logically equivalent - -```agda -sim-unit-presim-unit-ℤ : - {x y : ℤ} → presim-unit-ℤ x y → sim-unit-ℤ x y -sim-unit-presim-unit-ℤ {x} {y} H f = H - -presim-unit-sim-unit-ℤ : - {x y : ℤ} → sim-unit-ℤ x y → presim-unit-ℤ x y -presim-unit-sim-unit-ℤ {inl x} {inl y} H = H (λ t → Eq-eq-ℤ (pr1 t)) -presim-unit-sim-unit-ℤ {inl x} {inr y} H = H (λ t → Eq-eq-ℤ (pr1 t)) -presim-unit-sim-unit-ℤ {inr x} {inl y} H = H (λ t → Eq-eq-ℤ (pr2 t)) -pr1 (presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inl star)} H) = one-unit-ℤ -pr2 (presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inl star)} H) = refl -presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inr y)} H = - H (λ t → Eq-eq-ℤ (pr2 t)) -presim-unit-sim-unit-ℤ {inr (inr x)} {inr (inl star)} H = - H (λ t → Eq-eq-ℤ (pr1 t)) -presim-unit-sim-unit-ℤ {inr (inr x)} {inr (inr y)} H = - H (λ t → Eq-eq-ℤ (pr1 t)) -``` - -### The relations `presim-unit-ℤ` and `sim-unit-ℤ` relate `zero-ℤ` only to itself - -```agda -is-nonzero-presim-unit-ℤ : - {x y : ℤ} → presim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y -is-nonzero-presim-unit-ℤ {x} {y} (pair (pair v (pair u α)) β) f p = - Eq-eq-ℤ (ap (_*ℤ u) (inv q) ∙ (commutative-mul-ℤ v u ∙ α)) - where - q : is-zero-ℤ v - q = is-injective-right-mul-ℤ x f {v} {zero-ℤ} (β ∙ p) - -is-nonzero-sim-unit-ℤ : - {x y : ℤ} → sim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y -is-nonzero-sim-unit-ℤ H f = - is-nonzero-presim-unit-ℤ (H (f ∘ pr1)) f - -is-zero-sim-unit-ℤ : - {x y : ℤ} → sim-unit-ℤ x y → is-zero-ℤ x → is-zero-ℤ y -is-zero-sim-unit-ℤ {x} {y} H p = - double-negation-elim-is-decidable - ( has-decidable-equality-ℤ y zero-ℤ) - ( λ g → g (inv (β g) ∙ (ap ((u g) *ℤ_) p ∙ right-zero-law-mul-ℤ (u g)))) - where - K : is-nonzero-ℤ y → presim-unit-ℤ x y - K g = H (λ (u , v) → g v) - u : is-nonzero-ℤ y → ℤ - u g = pr1 (pr1 (K g)) - v : is-nonzero-ℤ y → ℤ - v g = pr1 (pr2 (pr1 (K g))) - β : (g : is-nonzero-ℤ y) → (u g) *ℤ x = y - β g = pr2 (K g) -``` - -### The relations `presim-unit-ℤ` and `sim-unit-ℤ` are equivalence relations - -```agda -refl-presim-unit-ℤ : is-reflexive presim-unit-ℤ -pr1 (refl-presim-unit-ℤ x) = one-unit-ℤ -pr2 (refl-presim-unit-ℤ x) = left-unit-law-mul-ℤ x - -refl-sim-unit-ℤ : is-reflexive sim-unit-ℤ -refl-sim-unit-ℤ x f = refl-presim-unit-ℤ x - -presim-unit-eq-ℤ : {x y : ℤ} → x = y → presim-unit-ℤ x y -presim-unit-eq-ℤ {x} refl = refl-presim-unit-ℤ x - -sim-unit-eq-ℤ : {x y : ℤ} → x = y → sim-unit-ℤ x y -sim-unit-eq-ℤ {x} refl = refl-sim-unit-ℤ x - -symmetric-presim-unit-ℤ : is-symmetric presim-unit-ℤ -symmetric-presim-unit-ℤ x y (pair (pair u H) p) = - f (is-one-or-neg-one-is-unit-ℤ u H) - where - f : is-one-or-neg-one-ℤ u → presim-unit-ℤ y x - pr1 (f (inl refl)) = one-unit-ℤ - pr2 (f (inl refl)) = inv p - pr1 (f (inr refl)) = neg-one-unit-ℤ - pr2 (f (inr refl)) = inv (inv (neg-neg-ℤ x) ∙ ap (neg-one-ℤ *ℤ_) p) - -symmetric-sim-unit-ℤ : is-symmetric sim-unit-ℤ -symmetric-sim-unit-ℤ x y H f = - symmetric-presim-unit-ℤ x y (H (λ p → f (pair (pr2 p) (pr1 p)))) - -is-nonzero-sim-unit-ℤ' : - {x y : ℤ} → sim-unit-ℤ x y → is-nonzero-ℤ y → is-nonzero-ℤ x -is-nonzero-sim-unit-ℤ' {x} {y} H = - is-nonzero-sim-unit-ℤ (symmetric-sim-unit-ℤ x y H) - -is-zero-sim-unit-ℤ' : - {x y : ℤ} → sim-unit-ℤ x y → is-zero-ℤ y → is-zero-ℤ x -is-zero-sim-unit-ℤ' {x} {y} H = is-zero-sim-unit-ℤ (symmetric-sim-unit-ℤ x y H) - -transitive-presim-unit-ℤ : is-transitive presim-unit-ℤ -transitive-presim-unit-ℤ x y z (pair (pair v K) q) (pair (pair u H) p) = - f (is-one-or-neg-one-is-unit-ℤ u H) (is-one-or-neg-one-is-unit-ℤ v K) - where - f : is-one-or-neg-one-ℤ u → is-one-or-neg-one-ℤ v → presim-unit-ℤ x z - pr1 (f (inl refl) (inl refl)) = one-unit-ℤ - pr2 (f (inl refl) (inl refl)) = p ∙ q - pr1 (f (inl refl) (inr refl)) = neg-one-unit-ℤ - pr2 (f (inl refl) (inr refl)) = ap neg-ℤ p ∙ q - pr1 (f (inr refl) (inl refl)) = neg-one-unit-ℤ - pr2 (f (inr refl) (inl refl)) = p ∙ q - pr1 (f (inr refl) (inr refl)) = one-unit-ℤ - pr2 (f (inr refl) (inr refl)) = inv (neg-neg-ℤ x) ∙ (ap neg-ℤ p ∙ q) - -transitive-sim-unit-ℤ : is-transitive sim-unit-ℤ -transitive-sim-unit-ℤ x y z K H f = - transitive-presim-unit-ℤ x y z - ( K (λ (p , q) → f (is-zero-sim-unit-ℤ' H p , q))) - ( H (λ (p , q) → f (p , is-zero-sim-unit-ℤ K q))) -``` - ### `sim-unit-ℤ x y` holds if and only if `x|y` and `y|x` ```agda @@ -624,7 +484,7 @@ pr2 (div-presim-unit-ℤ {x} {y} {x'} {y'} (pair u q) (pair v r) (pair d p)) = ( ( ap ( ((int-unit-ℤ v) *ℤ d) *ℤ_) ( ( inv (associative-mul-ℤ (int-unit-ℤ u) (int-unit-ℤ u) x)) ∙ - ( ap (_*ℤ x) (idempotent-is-unit-ℤ (is-unit-int-unit-ℤ u))))) ∙ + ( ap (_*ℤ x) (idempotent-is-unit-ℤ (is-unit-unit-ℤ u))))) ∙ ( ( associative-mul-ℤ (int-unit-ℤ v) d x) ∙ ( ( ap ((int-unit-ℤ v) *ℤ_) p) ∙ ( r))))) @@ -634,17 +494,21 @@ div-sim-unit-ℤ : div-ℤ x y → div-ℤ x' y' div-sim-unit-ℤ {x} {y} {x'} {y'} H K = div-presim-unit-ℤ (presim-unit-sim-unit-ℤ H) (presim-unit-sim-unit-ℤ K) +``` -div-int-abs-div-ℤ : +### An integer `x` divides an integer `y` if and only if its absolute value `|x|` divides `y` + +```agda +div-int-abs-divisor-div-ℤ : {x y : ℤ} → div-ℤ x y → div-ℤ (int-abs-ℤ x) y -div-int-abs-div-ℤ {x} {y} = +div-int-abs-divisor-div-ℤ {x} {y} = div-sim-unit-ℤ ( symmetric-sim-unit-ℤ (int-abs-ℤ x) x (sim-unit-abs-ℤ x)) ( refl-sim-unit-ℤ y) -div-div-int-abs-ℤ : +div-div-int-abs-divisor-ℤ : {x y : ℤ} → div-ℤ (int-abs-ℤ x) y → div-ℤ x y -div-div-int-abs-ℤ {x} {y} = +div-div-int-abs-divisor-ℤ {x} {y} = div-sim-unit-ℤ (sim-unit-abs-ℤ x) (refl-sim-unit-ℤ y) ``` @@ -660,7 +524,7 @@ is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz with ( is-one-or-neg-one-is-unit-ℤ ( int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) - ( is-unit-int-unit-ℤ (pr1 (H (λ u → nz (pr1 u)))))) + ( is-unit-unit-ℤ (pr1 (H (λ u → nz (pr1 u)))))) is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz | inl pos = inl ( equational-reasoning diff --git a/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md b/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md index fc8abf413d..42462305bf 100644 --- a/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md +++ b/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md @@ -127,13 +127,13 @@ is-common-divisor-int-abs-is-common-divisor-ℤ : {x y d : ℤ} → is-common-divisor-ℤ x y d → is-common-divisor-ℤ x y (int-abs-ℤ d) is-common-divisor-int-abs-is-common-divisor-ℤ = - map-product div-int-abs-div-ℤ div-int-abs-div-ℤ + map-product div-int-abs-divisor-div-ℤ div-int-abs-divisor-div-ℤ is-common-divisor-is-common-divisor-int-abs-ℤ : {x y d : ℤ} → is-common-divisor-ℤ x y (int-abs-ℤ d) → is-common-divisor-ℤ x y d is-common-divisor-is-common-divisor-int-abs-ℤ = - map-product div-div-int-abs-ℤ div-div-int-abs-ℤ + map-product div-div-int-abs-divisor-ℤ div-div-int-abs-divisor-ℤ is-common-divisor-is-gcd-ℤ : (a b d : ℤ) → is-gcd-ℤ a b d → is-common-divisor-ℤ a b d @@ -147,7 +147,7 @@ is-gcd-int-is-gcd-ℕ : {x y d : ℕ} → is-gcd-ℕ x y d → is-gcd-ℤ (int-ℕ x) (int-ℕ y) (int-ℕ d) pr1 (is-gcd-int-is-gcd-ℕ {x} {y} {d} H) = is-nonnegative-int-ℕ d pr1 (pr2 (is-gcd-int-is-gcd-ℕ {x} {y} {d} H) k) = - ( ( ( ( div-div-int-abs-ℤ) ∘ + ( ( ( ( div-div-int-abs-divisor-ℤ) ∘ ( div-int-div-ℕ)) ∘ ( pr1 (H (abs-ℤ k)))) ∘ ( is-common-divisor-is-common-divisor-int-ℕ)) ∘ @@ -157,7 +157,7 @@ pr2 (pr2 (is-gcd-int-is-gcd-ℕ {x} {y} {d} H) k) = ( is-common-divisor-int-is-common-divisor-ℕ)) ∘ ( pr2 (H (abs-ℤ k)))) ∘ ( div-div-int-ℕ)) ∘ - ( div-int-abs-div-ℤ) + ( div-int-abs-divisor-div-ℤ) is-gcd-is-gcd-int-ℕ : {x y d : ℕ} → is-gcd-ℤ (int-ℕ x) (int-ℕ y) (int-ℕ d) → is-gcd-ℕ x y d diff --git a/src/elementary-number-theory/modular-arithmetic.lagda.md b/src/elementary-number-theory/modular-arithmetic.lagda.md index 581cf46b87..b03ca3add5 100644 --- a/src/elementary-number-theory/modular-arithmetic.lagda.md +++ b/src/elementary-number-theory/modular-arithmetic.lagda.md @@ -788,8 +788,8 @@ has-no-fixed-points-succ-Fin {succ-ℕ k} x = is-decidable-div-ℤ : (d x : ℤ) → is-decidable (div-ℤ d x) is-decidable-div-ℤ d x = is-decidable-iff - ( div-div-int-abs-ℤ ∘ div-is-zero-mod-ℤ (abs-ℤ d) x) - ( is-zero-mod-div-ℤ (abs-ℤ d) x ∘ div-int-abs-div-ℤ) + ( div-div-int-abs-divisor-ℤ ∘ div-is-zero-mod-ℤ (abs-ℤ d) x) + ( is-zero-mod-div-ℤ (abs-ℤ d) x ∘ div-int-abs-divisor-div-ℤ) ( has-decidable-equality-ℤ-Mod ( abs-ℤ d) ( mod-ℤ (abs-ℤ d) x) diff --git a/src/elementary-number-theory/unit-integers.lagda.md b/src/elementary-number-theory/unit-integers.lagda.md new file mode 100644 index 0000000000..b9f4270326 --- /dev/null +++ b/src/elementary-number-theory/unit-integers.lagda.md @@ -0,0 +1,178 @@ +# Unit integers + +```agda +module elementary-number-theory.unit-integers where +``` + +
Imports + +```agda +open import commutative-algebra.groups-of-units-commutative-rings +open import commutative-algebra.invertible-elements-commutative-rings + +open import elementary-number-theory.equality-integers +open import elementary-number-theory.integers +open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.nonnegative-integers +open import elementary-number-theory.ring-of-integers + +open import foundation.action-on-identifications-functions +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.identity-types +open import foundation.transport-along-identifications +open import foundation.unit-type +open import foundation.universe-levels +``` + +
+ +## Idea + +An [integer](elementary-number-theory.integers.md) $a$ is said to be a {{#concept "unit" Disambiguation="integer" Agda=unit-ℤ WD="unit" WDID=Q118084}} if it [divides](elementary-number-theory.divisibility-integers.md) the integer $1$, i.e., if there exists an integer $b$ such that + +$$ +ab=1. +$$ + +Note that the predicate of being a unit is expressed without explicitly referring to divisibility of the integers in order to avoid cyclic dependencies in the library. Some properties of divisibility involving unit integers are studied in [`elementary-number-theory.divisibility-integers`](elementary-number-theory.divisibility-integers.md`). + +## Definitions + +### Unit integers + +```agda +is-unit-ℤ : ℤ → UU lzero +is-unit-ℤ = is-invertible-element-Commutative-Ring ℤ-Commutative-Ring + +unit-ℤ : UU lzero +unit-ℤ = type-group-of-units-Commutative-Ring ℤ-Commutative-Ring + +int-unit-ℤ : unit-ℤ → ℤ +int-unit-ℤ = inclusion-group-of-units-Commutative-Ring ℤ-Commutative-Ring + +is-unit-unit-ℤ : (x : unit-ℤ) → is-unit-ℤ (int-unit-ℤ x) +is-unit-unit-ℤ = + is-invertible-element-inclusion-group-of-units-Commutative-Ring + ℤ-Commutative-Ring +``` + +### The predicate of being $1$ or $-1$ + +```agda +is-one-or-neg-one-ℤ : ℤ → UU lzero +is-one-or-neg-one-ℤ x = (is-one-ℤ x) + (is-neg-one-ℤ x) +``` + +### The unit integer $1$ + +```agda +is-unit-one-ℤ : is-unit-ℤ one-ℤ +is-unit-one-ℤ = is-invertible-element-one-Commutative-Ring ℤ-Commutative-Ring + +one-unit-ℤ : unit-ℤ +one-unit-ℤ = unit-group-of-units-Commutative-Ring ℤ-Commutative-Ring + +is-unit-is-one-ℤ : (x : ℤ) → is-one-ℤ x → is-unit-ℤ x +is-unit-is-one-ℤ x p = + is-invertible-element-is-one-Commutative-Ring ℤ-Commutative-Ring x (inv p) +``` + +### The unit integer $-1$ + +```agda +is-unit-neg-one-ℤ : is-unit-ℤ neg-one-ℤ +is-unit-neg-one-ℤ = (neg-one-ℤ , refl , refl) + +neg-one-unit-ℤ : unit-ℤ +pr1 neg-one-unit-ℤ = neg-one-ℤ +pr2 neg-one-unit-ℤ = is-unit-neg-one-ℤ + +is-unit-is-neg-one-ℤ : + (x : ℤ) → is-neg-one-ℤ x → is-unit-ℤ x +is-unit-is-neg-one-ℤ .neg-one-ℤ refl = is-unit-neg-one-ℤ +``` + +## Properties + +### The product `xy` is a unit if and only if both `x` and `y` are units + +```agda +is-unit-mul-ℤ : + (x y : ℤ) → is-unit-ℤ x → is-unit-ℤ y → is-unit-ℤ (x *ℤ y) +is-unit-mul-ℤ = is-invertible-element-mul-Commutative-Ring ℤ-Commutative-Ring + +mul-unit-ℤ : unit-ℤ → unit-ℤ → unit-ℤ +mul-unit-ℤ = mul-group-of-units-Commutative-Ring ℤ-Commutative-Ring + +is-unit-left-factor-mul-ℤ : + (x y : ℤ) → is-unit-ℤ (x *ℤ y) → is-unit-ℤ x +is-unit-left-factor-mul-ℤ = + is-invertible-element-left-factor-Commutative-Ring ℤ-Commutative-Ring + +is-unit-right-factor-ℤ : + (x y : ℤ) → is-unit-ℤ (x *ℤ y) → is-unit-ℤ y +is-unit-right-factor-ℤ = + is-invertible-element-right-factor-Commutative-Ring ℤ-Commutative-Ring +``` + +### An integer is a unit if and only if it is $1$ or $-1$ + +```agda +is-unit-is-one-or-neg-one-ℤ : + (x : ℤ) → is-one-or-neg-one-ℤ x → is-unit-ℤ x +is-unit-is-one-or-neg-one-ℤ x (inl p) = is-unit-is-one-ℤ x p +is-unit-is-one-or-neg-one-ℤ x (inr p) = is-unit-is-neg-one-ℤ x p + +is-one-or-neg-one-is-unit-ℤ : + (x : ℤ) → is-unit-ℤ x → is-one-or-neg-one-ℤ x +is-one-or-neg-one-is-unit-ℤ (inl zero-ℕ) (d , p) = inr refl +is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (inl zero-ℕ , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ neg-one-ℤ (inl (succ-ℕ x)))) +is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (inl (succ-ℕ d) , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ (inl (succ-ℕ d)) (inl (succ-ℕ x)))) +is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (inr (inl star) , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ zero-ℤ (inl (succ-ℕ x)))) +is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (inr (inr zero-ℕ) , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ one-ℤ (inl (succ-ℕ x)))) +is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (inr (inr (succ-ℕ d)) , p , q) = + ex-falso + ( Eq-eq-ℤ (inv q ∙ compute-mul-ℤ (inr (inr (succ-ℕ d))) (inl (succ-ℕ x)))) +is-one-or-neg-one-is-unit-ℤ (inr (inl star)) (d , p , q) = + ex-falso (Eq-eq-ℤ (inv (right-zero-law-mul-ℤ d) ∙ q)) +is-one-or-neg-one-is-unit-ℤ (inr (inr zero-ℕ)) (d , p , q) = inl refl +is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (inl zero-ℕ , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ neg-one-ℤ (inr (inr (succ-ℕ x))))) +is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (inl (succ-ℕ d) , p , q) = + ex-falso + ( Eq-eq-ℤ (inv q ∙ compute-mul-ℤ (inl (succ-ℕ d)) (inr (inr (succ-ℕ x))))) +is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (inr (inl star) , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ zero-ℤ (inr (inr (succ-ℕ x))))) +is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (inr (inr zero-ℕ) , p , q) = + ex-falso (Eq-eq-ℤ (inv q ∙ compute-mul-ℤ one-ℤ (inr (inr (succ-ℕ x))))) +is-one-or-neg-one-is-unit-ℤ + (inr (inr (succ-ℕ x))) (inr (inr (succ-ℕ d)) , p , q) = + ex-falso + ( Eq-eq-ℤ + ( inv q ∙ compute-mul-ℤ (inr (inr (succ-ℕ d))) (inr (inr (succ-ℕ x))))) +``` + +### Unit integers are idempotent + +```agda +idempotent-is-unit-ℤ : {x : ℤ} → is-unit-ℤ x → x *ℤ x = one-ℤ +idempotent-is-unit-ℤ {x} H = + f (is-one-or-neg-one-is-unit-ℤ x H) + where + f : is-one-or-neg-one-ℤ x → x *ℤ x = one-ℤ + f (inl refl) = refl + f (inr refl) = refl + +abstract + is-one-is-unit-int-ℕ : (x : ℕ) → is-unit-ℤ (int-ℕ x) → is-one-ℕ x + is-one-is-unit-int-ℕ x H with is-one-or-neg-one-is-unit-ℤ (int-ℕ x) H + ... | inl p = is-injective-int-ℕ p + ... | inr p = ex-falso (tr is-nonnegative-ℤ p (is-nonnegative-int-ℕ x)) +``` diff --git a/src/elementary-number-theory/unit-similarity-integers.lagda.md b/src/elementary-number-theory/unit-similarity-integers.lagda.md new file mode 100644 index 0000000000..ec8a1ce763 --- /dev/null +++ b/src/elementary-number-theory/unit-similarity-integers.lagda.md @@ -0,0 +1,165 @@ +# Unit similarity of integers + +```agda +module elementary-number-theory.unit-similarity-integers where +``` + +
Imports + +```agda +open import elementary-number-theory.equality-integers +open import elementary-number-theory.integers +open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.unit-integers + +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.negation +open import foundation.universe-levels +``` + +
+ +## Idea + +Two [integers](elementary-number-theory.integers.md) $a$ and $b$ are said to be {{#concept "unit similar" Disambiguation="integers" Agda=sim-unit-ℤ}} if + +```text +¬ (is-nonzero-ℤ a × is-nonzero-ℤ b) → Σ (u : unit-ℤ) ua=b. +``` + +The idea is that $a$ and $b$ are unit similar if there exists a [unit integer](elementary-number-theory.unit-integers.md) $u$ such that $ua=b$. However, this type is not a proposition if both $a$ and $b$ are $0$. The unit similarity relation has therefore a slightly more complicated expression. + +The unit similarity relation is an [equivalence relation](foundation.equivalence-relations.md) on the integers. + +## Definitions + +### The equivalence relation `sim-unit-ℤ` + +```agda +presim-unit-ℤ : ℤ → ℤ → UU lzero +presim-unit-ℤ x y = Σ unit-ℤ (λ u → int-unit-ℤ u *ℤ x = y) + +sim-unit-ℤ : ℤ → ℤ → UU lzero +sim-unit-ℤ x y = ¬ (is-zero-ℤ x × is-zero-ℤ y) → presim-unit-ℤ x y +``` + +## Properties + +### The relations `presim-unit-ℤ` and `sim-unit-ℤ` are logically equivalent + +```agda +sim-unit-presim-unit-ℤ : + {x y : ℤ} → presim-unit-ℤ x y → sim-unit-ℤ x y +sim-unit-presim-unit-ℤ {x} {y} H f = H + +presim-unit-sim-unit-ℤ : + {x y : ℤ} → sim-unit-ℤ x y → presim-unit-ℤ x y +presim-unit-sim-unit-ℤ {inl x} {inl y} H = H (λ t → Eq-eq-ℤ (pr1 t)) +presim-unit-sim-unit-ℤ {inl x} {inr y} H = H (λ t → Eq-eq-ℤ (pr1 t)) +presim-unit-sim-unit-ℤ {inr x} {inl y} H = H (λ t → Eq-eq-ℤ (pr2 t)) +pr1 (presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inl star)} H) = one-unit-ℤ +pr2 (presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inl star)} H) = refl +presim-unit-sim-unit-ℤ {inr (inl star)} {inr (inr y)} H = + H (λ t → Eq-eq-ℤ (pr2 t)) +presim-unit-sim-unit-ℤ {inr (inr x)} {inr (inl star)} H = + H (λ t → Eq-eq-ℤ (pr1 t)) +presim-unit-sim-unit-ℤ {inr (inr x)} {inr (inr y)} H = + H (λ t → Eq-eq-ℤ (pr1 t)) +``` + +### The relations `presim-unit-ℤ` and `sim-unit-ℤ` relate `zero-ℤ` only to itself + +```agda +is-nonzero-presim-unit-ℤ : + {x y : ℤ} → presim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y +is-nonzero-presim-unit-ℤ {x} {y} (pair (pair v (pair u α)) β) f p = + Eq-eq-ℤ (ap (_*ℤ u) (inv q) ∙ (commutative-mul-ℤ v u ∙ α)) + where + q : is-zero-ℤ v + q = is-injective-right-mul-ℤ x f {v} {zero-ℤ} (β ∙ p) + +is-nonzero-sim-unit-ℤ : + {x y : ℤ} → sim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y +is-nonzero-sim-unit-ℤ H f = + is-nonzero-presim-unit-ℤ (H (f ∘ pr1)) f + +is-zero-sim-unit-ℤ : + {x y : ℤ} → sim-unit-ℤ x y → is-zero-ℤ x → is-zero-ℤ y +is-zero-sim-unit-ℤ {x} {y} H p = + double-negation-elim-is-decidable + ( has-decidable-equality-ℤ y zero-ℤ) + ( λ g → g (inv (β g) ∙ (ap ((u g) *ℤ_) p ∙ right-zero-law-mul-ℤ (u g)))) + where + K : is-nonzero-ℤ y → presim-unit-ℤ x y + K g = H (λ (u , v) → g v) + u : is-nonzero-ℤ y → ℤ + u g = pr1 (pr1 (K g)) + v : is-nonzero-ℤ y → ℤ + v g = pr1 (pr2 (pr1 (K g))) + β : (g : is-nonzero-ℤ y) → (u g) *ℤ x = y + β g = pr2 (K g) +``` + +### The relations `presim-unit-ℤ` and `sim-unit-ℤ` are equivalence relations + +```agda +refl-presim-unit-ℤ : is-reflexive presim-unit-ℤ +pr1 (refl-presim-unit-ℤ x) = one-unit-ℤ +pr2 (refl-presim-unit-ℤ x) = left-unit-law-mul-ℤ x + +refl-sim-unit-ℤ : is-reflexive sim-unit-ℤ +refl-sim-unit-ℤ x f = refl-presim-unit-ℤ x + +presim-unit-eq-ℤ : {x y : ℤ} → x = y → presim-unit-ℤ x y +presim-unit-eq-ℤ {x} refl = refl-presim-unit-ℤ x + +sim-unit-eq-ℤ : {x y : ℤ} → x = y → sim-unit-ℤ x y +sim-unit-eq-ℤ {x} refl = refl-sim-unit-ℤ x + +symmetric-presim-unit-ℤ : is-symmetric presim-unit-ℤ +symmetric-presim-unit-ℤ x y (pair (pair u H) p) = + f (is-one-or-neg-one-is-unit-ℤ u H) + where + f : is-one-or-neg-one-ℤ u → presim-unit-ℤ y x + pr1 (f (inl refl)) = one-unit-ℤ + pr2 (f (inl refl)) = inv p + pr1 (f (inr refl)) = neg-one-unit-ℤ + pr2 (f (inr refl)) = inv (inv (neg-neg-ℤ x) ∙ ap (neg-one-ℤ *ℤ_) p) + +symmetric-sim-unit-ℤ : is-symmetric sim-unit-ℤ +symmetric-sim-unit-ℤ x y H f = + symmetric-presim-unit-ℤ x y (H (λ p → f (pair (pr2 p) (pr1 p)))) + +is-nonzero-sim-unit-ℤ' : + {x y : ℤ} → sim-unit-ℤ x y → is-nonzero-ℤ y → is-nonzero-ℤ x +is-nonzero-sim-unit-ℤ' {x} {y} H = + is-nonzero-sim-unit-ℤ (symmetric-sim-unit-ℤ x y H) + +is-zero-sim-unit-ℤ' : + {x y : ℤ} → sim-unit-ℤ x y → is-zero-ℤ y → is-zero-ℤ x +is-zero-sim-unit-ℤ' {x} {y} H = is-zero-sim-unit-ℤ (symmetric-sim-unit-ℤ x y H) + +transitive-presim-unit-ℤ : is-transitive presim-unit-ℤ +transitive-presim-unit-ℤ x y z (pair (pair v K) q) (pair (pair u H) p) = + f (is-one-or-neg-one-is-unit-ℤ u H) (is-one-or-neg-one-is-unit-ℤ v K) + where + f : is-one-or-neg-one-ℤ u → is-one-or-neg-one-ℤ v → presim-unit-ℤ x z + pr1 (f (inl refl) (inl refl)) = one-unit-ℤ + pr2 (f (inl refl) (inl refl)) = p ∙ q + pr1 (f (inl refl) (inr refl)) = neg-one-unit-ℤ + pr2 (f (inl refl) (inr refl)) = ap neg-ℤ p ∙ q + pr1 (f (inr refl) (inl refl)) = neg-one-unit-ℤ + pr2 (f (inr refl) (inl refl)) = p ∙ q + pr1 (f (inr refl) (inr refl)) = one-unit-ℤ + pr2 (f (inr refl) (inr refl)) = inv (neg-neg-ℤ x) ∙ (ap neg-ℤ p ∙ q) + +transitive-sim-unit-ℤ : is-transitive sim-unit-ℤ +transitive-sim-unit-ℤ x y z K H f = + transitive-presim-unit-ℤ x y z + ( K (λ (p , q) → f (is-zero-sim-unit-ℤ' H p , q))) + ( H (λ (p , q) → f (p , is-zero-sim-unit-ℤ K q))) +``` + diff --git a/src/group-theory/invertible-elements-monoids.lagda.md b/src/group-theory/invertible-elements-monoids.lagda.md index c50cf77558..2f6b9aa379 100644 --- a/src/group-theory/invertible-elements-monoids.lagda.md +++ b/src/group-theory/invertible-elements-monoids.lagda.md @@ -19,6 +19,7 @@ open import foundation.injective-maps open import foundation.propositions open import foundation.sets open import foundation.subtypes +open import foundation.transport-along-identifications open import foundation.universe-levels open import group-theory.monoids @@ -289,9 +290,34 @@ module _ left-unit-law-mul-Monoid M (unit-Monoid M) pr2 (pr2 is-invertible-element-unit-Monoid) = left-unit-law-mul-Monoid M (unit-Monoid M) + + is-invertible-element-is-unit-Monoid : + (x : type-Monoid M) → unit-Monoid M = x → + is-invertible-element-Monoid M x + is-invertible-element-is-unit-Monoid .(unit-Monoid M) refl = + is-invertible-element-unit-Monoid +``` + +### The inverse of an invertible element is invertible + +```agda +module _ + {l : Level} (M : Monoid l) + where + + is-invertible-element-inv-is-invertible-element-Monoid : + {x : type-Monoid M} (H : is-invertible-element-Monoid M x) → + is-invertible-element-Monoid M (inv-is-invertible-element-Monoid M H) + pr1 (is-invertible-element-inv-is-invertible-element-Monoid {x} H) = x + pr1 (pr2 (is-invertible-element-inv-is-invertible-element-Monoid H)) = + is-left-inverse-inv-is-invertible-element-Monoid M H + pr2 (pr2 (is-invertible-element-inv-is-invertible-element-Monoid H)) = + is-right-inverse-inv-is-invertible-element-Monoid M H ``` -### Invertible elements are closed under multiplication +### If two of the three elements `x`, `y`, and `xy` are invertible, then so is the third + +#### Invertible elements are closed under multiplication ```agda module _ @@ -349,21 +375,52 @@ module _ ( is-left-invertible-is-invertible-element-Monoid M y K)) ``` -### The inverse of an invertible element is invertible +#### If `y` and `xy` are invertible, then so is `x` ```agda module _ - {l : Level} (M : Monoid l) + {l : Level} (M : Monoid l) (x y : type-Monoid M) where - is-invertible-element-inv-is-invertible-element-Monoid : - {x : type-Monoid M} (H : is-invertible-element-Monoid M x) → - is-invertible-element-Monoid M (inv-is-invertible-element-Monoid M H) - pr1 (is-invertible-element-inv-is-invertible-element-Monoid {x} H) = x - pr1 (pr2 (is-invertible-element-inv-is-invertible-element-Monoid H)) = - is-left-inverse-inv-is-invertible-element-Monoid M H - pr2 (pr2 (is-invertible-element-inv-is-invertible-element-Monoid H)) = - is-right-inverse-inv-is-invertible-element-Monoid M H + is-invertible-element-left-factor-Monoid : + is-invertible-element-Monoid M y → + is-invertible-element-Monoid M (mul-Monoid M x y) → + is-invertible-element-Monoid M x + is-invertible-element-left-factor-Monoid H@(y' , Ly , Ry) K = + tr + ( is-invertible-element-Monoid M) + ( associative-mul-Monoid M x y y' ∙ + ap (mul-Monoid M x) Ly ∙ + right-unit-law-mul-Monoid M x) + ( is-invertible-element-mul-Monoid M + ( mul-Monoid M x y) + ( y') + ( K) + ( is-invertible-element-inv-is-invertible-element-Monoid M H)) +``` + +#### If `x` and `xy` are invertible, then so is `y` + +```agda +module _ + {l : Level} (M : Monoid l) (x y : type-Monoid M) + where + + is-invertible-element-right-factor-Monoid : + is-invertible-element-Monoid M x → + is-invertible-element-Monoid M (mul-Monoid M x y) → + is-invertible-element-Monoid M y + is-invertible-element-right-factor-Monoid H@(x' , Lx , Rx) K = + tr + ( is-invertible-element-Monoid M) + ( inv (associative-mul-Monoid M x' x y) ∙ + ap (mul-Monoid' M y) Rx ∙ + left-unit-law-mul-Monoid M y) + ( is-invertible-element-mul-Monoid M + ( x') + ( mul-Monoid M x y) + ( is-invertible-element-inv-is-invertible-element-Monoid M H) + ( K)) ``` ### An element is invertible if and only if multiplying by it is an equivalence diff --git a/src/ring-theory/groups-of-units-rings.lagda.md b/src/ring-theory/groups-of-units-rings.lagda.md index 16e293804e..38b15cb9d0 100644 --- a/src/ring-theory/groups-of-units-rings.lagda.md +++ b/src/ring-theory/groups-of-units-rings.lagda.md @@ -9,8 +9,10 @@ module ring-theory.groups-of-units-rings where ```agda open import category-theory.functors-large-precategories +open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.propositions +open import foundation.subtypes open import foundation.universe-levels open import group-theory.cores-monoids @@ -32,7 +34,7 @@ open import ring-theory.rings ## Idea -The **group of units** of a [ring](ring-theory.rings.md) `R` is the +The {{#concept "group of units" Disambiguation="ring" WD="unit" WDID=Q118084}} of a [ring](ring-theory.rings.md) `R` is the [group](group-theory.groups.md) consisting of all the [invertible elements](ring-theory.invertible-elements-rings.md) in `R`. Equivalently, the group of units of `R` is the @@ -41,6 +43,8 @@ Equivalently, the group of units of `R` is the ## Definitions +### The group of units of a ring + ```agda module _ {l : Level} (R : Ring l) @@ -136,6 +140,13 @@ module _ inclusion-group-of-units-Ring = inclusion-core-Monoid (multiplicative-monoid-Ring R) + is-invertible-element-inclusion-group-of-units-Ring : + (x : type-group-of-units-Ring) → + is-invertible-element-Ring R + ( inclusion-group-of-units-Ring x) + is-invertible-element-inclusion-group-of-units-Ring = + is-in-subtype-inclusion-subtype subtype-group-of-units-Ring + preserves-mul-inclusion-group-of-units-Ring : {x y : type-group-of-units-Ring} → inclusion-group-of-units-Ring (mul-group-of-units-Ring x y) = @@ -210,7 +221,7 @@ module _ ( hom-multiplicative-monoid-hom-Ring R S f) ``` -#### The functorial laws of `group-of-units-Ring` +#### The functorial laws of the group-of-units functor ```agda module _ @@ -245,7 +256,7 @@ module _ ( hom-multiplicative-monoid-hom-Ring R S f) ``` -#### The functor `group-of-units-Ring` +#### The group-of-units functor ```agda group-of-units-ring-functor-Large-Precategory : @@ -266,3 +277,21 @@ preserves-id-functor-Large-Precategory group-of-units-ring-functor-Large-Precategory {X = R} = preserves-id-hom-group-of-units-hom-Ring R ``` + +### Negatives of units + +```agda +module _ + {l : Level} (R : Ring l) + where + + neg-group-of-units-Ring : + type-group-of-units-Ring R → type-group-of-units-Ring R + pr1 (neg-group-of-units-Ring (x , H)) = neg-Ring R x + pr2 (neg-group-of-units-Ring (x , H)) = is-invertible-element-neg-Ring R x H + + neg-unit-group-of-units-Ring : + type-group-of-units-Ring R + neg-unit-group-of-units-Ring = + neg-group-of-units-Ring (unit-group-of-units-Ring R) +``` diff --git a/src/ring-theory/invertible-elements-rings.lagda.md b/src/ring-theory/invertible-elements-rings.lagda.md index 5066846791..9f9781e0c5 100644 --- a/src/ring-theory/invertible-elements-rings.lagda.md +++ b/src/ring-theory/invertible-elements-rings.lagda.md @@ -221,7 +221,7 @@ module _ ( multiplicative-monoid-Ring R) ``` -### Any invertible element of a monoid has a contractible type of right inverses +### Any invertible element of a ring has a contractible type of right inverses ```agda module _ @@ -236,7 +236,7 @@ module _ ( multiplicative-monoid-Ring R) ``` -### Any invertible element of a monoid has a contractible type of left inverses +### Any invertible element of a ring has a contractible type of left inverses ```agda module _ @@ -251,7 +251,7 @@ module _ ( multiplicative-monoid-Ring R) ``` -### The unit of a monoid is invertible +### The unit of a ring is invertible ```agda module _ @@ -275,9 +275,18 @@ module _ is-invertible-element-one-Ring = is-invertible-element-unit-Monoid ( multiplicative-monoid-Ring R) + + is-invertible-element-is-one-Ring : + (x : type-Ring R) → one-Ring R = x → + is-invertible-element-Ring R x + is-invertible-element-is-one-Ring = + is-invertible-element-is-unit-Monoid + ( multiplicative-monoid-Ring R) ``` -### Invertible elements are closed under multiplication +### If two of the three elements `x`, `y`, and `xy` are invertible, then so is the third + +#### Invertible elements are closed under multiplication ```agda module _ @@ -312,6 +321,42 @@ module _ ( multiplicative-monoid-Ring R) ``` +#### If `y` and `xy` are invertible, then so is `x` + +```agda +module _ + {l : Level} (R : Ring l) (x y : type-Ring R) + where + + is-invertible-element-left-factor-Ring : + is-invertible-element-Ring R y → + is-invertible-element-Ring R (mul-Ring R x y) → + is-invertible-element-Ring R x + is-invertible-element-left-factor-Ring = + is-invertible-element-left-factor-Monoid + ( multiplicative-monoid-Ring R) + ( x) + ( y) +``` + +#### If `x` and `xy` are invertible, then so is `y` + +```agda +module _ + {l : Level} (R : Ring l) (x y : type-Ring R) + where + + is-invertible-element-right-factor-Ring : + is-invertible-element-Ring R x → + is-invertible-element-Ring R (mul-Ring R x y) → + is-invertible-element-Ring R y + is-invertible-element-right-factor-Ring = + is-invertible-element-right-factor-Monoid + ( multiplicative-monoid-Ring R) + ( x) + ( y) +``` + ### Invertible elements are closed under negatives ```agda From 39b6264173ac4ec7c0d4ff412fcd7f9a08289cd6 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 29 Oct 2024 18:28:33 -0400 Subject: [PATCH 13/72] work --- .../divisibility-integers.lagda.md | 124 +----------------- .../unit-integers.lagda.md | 2 +- .../unit-similarity-integers.lagda.md | 12 +- 3 files changed, 15 insertions(+), 123 deletions(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 7d06065f9d..9fb6209303 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -19,6 +19,8 @@ open import elementary-number-theory.nonnegative-integers open import elementary-number-theory.nonpositive-integers open import elementary-number-theory.nonzero-integers open import elementary-number-theory.positive-and-negative-integers +open import elementary-number-theory.unit-integers +open import elementary-number-theory.unit-similarity-integers open import foundation.action-on-identifications-functions open import foundation.binary-relations @@ -102,9 +104,9 @@ concatenate-eq-div-eq-ℤ refl p refl = p ```agda div-is-unit-ℤ : (x y : ℤ) → is-unit-ℤ x → div-ℤ x y -pr1 (div-is-unit-ℤ x y (pair d p)) = y *ℤ d -pr2 (div-is-unit-ℤ x y (pair d p)) = - associative-mul-ℤ y d x ∙ (ap (y *ℤ_) p ∙ right-unit-law-mul-ℤ y) +pr1 (div-is-unit-ℤ x y (d , p , q)) = y *ℤ d +pr2 (div-is-unit-ℤ x y (d , p , q)) = + associative-mul-ℤ y d x ∙ (ap (y *ℤ_) q ∙ right-unit-law-mul-ℤ y) ``` ### Divisibility by a nonzero integer is a property @@ -324,122 +326,6 @@ pr2 (div-div-int-ℕ {succ-ℕ x} {y} (pair d p)) = ( p))) ``` -### An integer is a unit if and only if it is `1` or `-1` - -```agda -is-one-or-neg-one-ℤ : ℤ → UU lzero -is-one-or-neg-one-ℤ x = (is-one-ℤ x) + (is-neg-one-ℤ x) - -is-unit-one-ℤ : is-unit-ℤ one-ℤ -is-unit-one-ℤ = refl-div-ℤ one-ℤ - -one-unit-ℤ : unit-ℤ -pr1 one-unit-ℤ = one-ℤ -pr2 one-unit-ℤ = is-unit-one-ℤ - -is-unit-is-one-ℤ : - (x : ℤ) → is-one-ℤ x → is-unit-ℤ x -is-unit-is-one-ℤ .one-ℤ refl = is-unit-one-ℤ - -is-unit-neg-one-ℤ : is-unit-ℤ neg-one-ℤ -pr1 is-unit-neg-one-ℤ = neg-one-ℤ -pr2 is-unit-neg-one-ℤ = refl - -neg-one-unit-ℤ : unit-ℤ -pr1 neg-one-unit-ℤ = neg-one-ℤ -pr2 neg-one-unit-ℤ = is-unit-neg-one-ℤ - -is-unit-is-neg-one-ℤ : - (x : ℤ) → is-neg-one-ℤ x → is-unit-ℤ x -is-unit-is-neg-one-ℤ .neg-one-ℤ refl = is-unit-neg-one-ℤ - -is-unit-is-one-or-neg-one-ℤ : - (x : ℤ) → is-one-or-neg-one-ℤ x → is-unit-ℤ x -is-unit-is-one-or-neg-one-ℤ x (inl p) = is-unit-is-one-ℤ x p -is-unit-is-one-or-neg-one-ℤ x (inr p) = is-unit-is-neg-one-ℤ x p - -is-one-or-neg-one-is-unit-ℤ : - (x : ℤ) → is-unit-ℤ x → is-one-or-neg-one-ℤ x -is-one-or-neg-one-is-unit-ℤ (inl zero-ℕ) (pair d p) = inr refl -is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (pair (inl zero-ℕ) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ neg-one-ℤ (inl (succ-ℕ x)))) -is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (pair (inl (succ-ℕ d)) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ (inl (succ-ℕ d)) (inl (succ-ℕ x)))) -is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (pair (inr (inl star)) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ zero-ℤ (inl (succ-ℕ x)))) -is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (pair (inr (inr zero-ℕ)) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ one-ℤ (inl (succ-ℕ x)))) -is-one-or-neg-one-is-unit-ℤ (inl (succ-ℕ x)) (pair (inr (inr (succ-ℕ d))) p) = - ex-falso - ( Eq-eq-ℤ (inv p ∙ compute-mul-ℤ (inr (inr (succ-ℕ d))) (inl (succ-ℕ x)))) -is-one-or-neg-one-is-unit-ℤ (inr (inl star)) (pair d p) = - ex-falso (Eq-eq-ℤ (inv (right-zero-law-mul-ℤ d) ∙ p)) -is-one-or-neg-one-is-unit-ℤ (inr (inr zero-ℕ)) (pair d p) = inl refl -is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (pair (inl zero-ℕ) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ neg-one-ℤ (inr (inr (succ-ℕ x))))) -is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (pair (inl (succ-ℕ d)) p) = - ex-falso - ( Eq-eq-ℤ (inv p ∙ compute-mul-ℤ (inl (succ-ℕ d)) (inr (inr (succ-ℕ x))))) -is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (pair (inr (inl star)) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ zero-ℤ (inr (inr (succ-ℕ x))))) -is-one-or-neg-one-is-unit-ℤ (inr (inr (succ-ℕ x))) (pair (inr (inr zero-ℕ)) p) = - ex-falso (Eq-eq-ℤ (inv p ∙ compute-mul-ℤ one-ℤ (inr (inr (succ-ℕ x))))) -is-one-or-neg-one-is-unit-ℤ - (inr (inr (succ-ℕ x))) (pair (inr (inr (succ-ℕ d))) p) = - ex-falso - ( Eq-eq-ℤ - ( inv p ∙ compute-mul-ℤ (inr (inr (succ-ℕ d))) (inr (inr (succ-ℕ x))))) -``` - -### Units are idempotent - -```agda -idempotent-is-unit-ℤ : {x : ℤ} → is-unit-ℤ x → x *ℤ x = one-ℤ -idempotent-is-unit-ℤ {x} H = - f (is-one-or-neg-one-is-unit-ℤ x H) - where - f : is-one-or-neg-one-ℤ x → x *ℤ x = one-ℤ - f (inl refl) = refl - f (inr refl) = refl - -abstract - is-one-is-unit-int-ℕ : (x : ℕ) → is-unit-ℤ (int-ℕ x) → is-one-ℕ x - is-one-is-unit-int-ℕ x H with is-one-or-neg-one-is-unit-ℤ (int-ℕ x) H - ... | inl p = is-injective-int-ℕ p - ... | inr p = ex-falso (tr is-nonnegative-ℤ p (is-nonnegative-int-ℕ x)) -``` - -### The product `xy` is a unit if and only if both `x` and `y` are units - -```agda -is-unit-mul-ℤ : - (x y : ℤ) → is-unit-ℤ x → is-unit-ℤ y → is-unit-ℤ (x *ℤ y) -pr1 (is-unit-mul-ℤ x y (pair d p) (pair e q)) = e *ℤ d -pr2 (is-unit-mul-ℤ x y (pair d p) (pair e q)) = - ( associative-mul-ℤ e d (x *ℤ y)) ∙ - ( ( ap - ( e *ℤ_) - ( ( inv (associative-mul-ℤ d x y)) ∙ - ( ap (_*ℤ y) p))) ∙ - ( q)) - -mul-unit-ℤ : unit-ℤ → unit-ℤ → unit-ℤ -pr1 (mul-unit-ℤ (pair x H) (pair y K)) = x *ℤ y -pr2 (mul-unit-ℤ (pair x H) (pair y K)) = is-unit-mul-ℤ x y H K - -is-unit-left-factor-mul-ℤ : - (x y : ℤ) → is-unit-ℤ (x *ℤ y) → is-unit-ℤ x -pr1 (is-unit-left-factor-mul-ℤ x y (pair d p)) = d *ℤ y -pr2 (is-unit-left-factor-mul-ℤ x y (pair d p)) = - associative-mul-ℤ d y x ∙ (ap (d *ℤ_) (commutative-mul-ℤ y x) ∙ p) - -is-unit-right-factor-ℤ : - (x y : ℤ) → is-unit-ℤ (x *ℤ y) → is-unit-ℤ y -is-unit-right-factor-ℤ x y (pair d p) = - is-unit-left-factor-mul-ℤ y x - ( pair d (ap (d *ℤ_) (commutative-mul-ℤ y x) ∙ p)) -``` - ### `sim-unit-ℤ x y` holds if and only if `x|y` and `y|x` ```agda diff --git a/src/elementary-number-theory/unit-integers.lagda.md b/src/elementary-number-theory/unit-integers.lagda.md index b9f4270326..c3860d33b7 100644 --- a/src/elementary-number-theory/unit-integers.lagda.md +++ b/src/elementary-number-theory/unit-integers.lagda.md @@ -97,7 +97,7 @@ is-unit-is-neg-one-ℤ .neg-one-ℤ refl = is-unit-neg-one-ℤ ## Properties -### The product `xy` is a unit if and only if both `x` and `y` are units +### The product $xy$ is a unit if and only if both $x$ and $y$ are units ```agda is-unit-mul-ℤ : diff --git a/src/elementary-number-theory/unit-similarity-integers.lagda.md b/src/elementary-number-theory/unit-similarity-integers.lagda.md index ec8a1ce763..c43aefff36 100644 --- a/src/elementary-number-theory/unit-similarity-integers.lagda.md +++ b/src/elementary-number-theory/unit-similarity-integers.lagda.md @@ -10,13 +10,19 @@ module elementary-number-theory.unit-similarity-integers where open import elementary-number-theory.equality-integers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.nonzero-integers open import elementary-number-theory.unit-integers +open import foundation.action-on-identifications-functions +open import foundation.binary-relations open import foundation.cartesian-product-types open import foundation.coproduct-types open import foundation.dependent-pair-types +open import foundation.decidable-types +open import foundation.function-types open import foundation.identity-types open import foundation.negation +open import foundation.unit-type open import foundation.universe-levels ``` @@ -75,11 +81,11 @@ presim-unit-sim-unit-ℤ {inr (inr x)} {inr (inr y)} H = ```agda is-nonzero-presim-unit-ℤ : {x y : ℤ} → presim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y -is-nonzero-presim-unit-ℤ {x} {y} (pair (pair v (pair u α)) β) f p = - Eq-eq-ℤ (ap (_*ℤ u) (inv q) ∙ (commutative-mul-ℤ v u ∙ α)) +is-nonzero-presim-unit-ℤ {x} {y} ((v , (u , (α , α'))) , β) f p = + Eq-eq-ℤ (ap (_*ℤ u) (inv q) ∙ commutative-mul-ℤ v u ∙ α') where q : is-zero-ℤ v - q = is-injective-right-mul-ℤ x f {v} {zero-ℤ} (β ∙ p) + q = is-injective-right-mul-ℤ x f (β ∙ p) is-nonzero-sim-unit-ℤ : {x y : ℤ} → sim-unit-ℤ x y → is-nonzero-ℤ x → is-nonzero-ℤ y From dec3433a59b529763c6d02c5ea87691735f90777 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 29 Oct 2024 20:52:55 -0400 Subject: [PATCH 14/72] implementing some algebra into the theory of integers --- ...groups-of-units-commutative-rings.lagda.md | 7 ++-- src/elementary-number-theory.lagda.md | 2 ++ .../congruence-integers.lagda.md | 34 +++++++++++-------- .../divisibility-integers.lagda.md | 25 ++++++++++---- .../greatest-common-divisor-integers.lagda.md | 6 ++-- .../modular-arithmetic.lagda.md | 1 + .../parity-integers.lagda.md | 1 + .../reduced-integer-fractions.lagda.md | 1 + .../unit-integers.lagda.md | 11 ++++-- .../unit-similarity-integers.lagda.md | 14 +++++--- .../groups-of-units-rings.lagda.md | 7 ++-- 11 files changed, 72 insertions(+), 37 deletions(-) diff --git a/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md b/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md index 78457c72b9..218fb440e2 100644 --- a/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md +++ b/src/commutative-algebra/groups-of-units-commutative-rings.lagda.md @@ -36,8 +36,9 @@ open import ring-theory.groups-of-units-rings ## Idea -The {{#concept "group of units" Disambiguation="commutative ring" WD="unit" WDID=Q118084}} of a -[commutative ring](commutative-algebra.commutative-rings.md) `A` is the +The +{{#concept "group of units" Disambiguation="commutative ring" WD="unit" WDID=Q118084}} +of a [commutative ring](commutative-algebra.commutative-rings.md) `A` is the [abelian group](group-theory.abelian-groups.md) consisting of all the [invertible elements](commutative-algebra.invertible-elements-commutative-rings.md) in `A`. Equivalently, the group of units of `A` is the @@ -340,7 +341,7 @@ preserves-id-functor-Large-Precategory module _ {l : Level} (A : Commutative-Ring l) where - + neg-group-of-units-Commutative-Ring : type-group-of-units-Commutative-Ring A → type-group-of-units-Commutative-Ring A diff --git a/src/elementary-number-theory.lagda.md b/src/elementary-number-theory.lagda.md index edc4629d37..f321f56e51 100644 --- a/src/elementary-number-theory.lagda.md +++ b/src/elementary-number-theory.lagda.md @@ -158,6 +158,8 @@ open import elementary-number-theory.triangular-numbers public open import elementary-number-theory.twin-prime-conjecture public open import elementary-number-theory.type-arithmetic-natural-numbers public open import elementary-number-theory.unit-elements-standard-finite-types public +open import elementary-number-theory.unit-integers public +open import elementary-number-theory.unit-similarity-integers public open import elementary-number-theory.unit-similarity-standard-finite-types public open import elementary-number-theory.universal-property-integers public open import elementary-number-theory.universal-property-natural-numbers public diff --git a/src/elementary-number-theory/congruence-integers.lagda.md b/src/elementary-number-theory/congruence-integers.lagda.md index 9844348d01..c993c7b4c6 100644 --- a/src/elementary-number-theory/congruence-integers.lagda.md +++ b/src/elementary-number-theory/congruence-integers.lagda.md @@ -16,6 +16,8 @@ open import elementary-number-theory.divisibility-integers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.unit-integers +open import elementary-number-theory.unit-similarity-integers open import foundation.action-on-identifications-functions open import foundation.binary-relations @@ -57,23 +59,25 @@ is-discrete-cong-ℤ .zero-ℤ refl x y K = eq-diff-ℤ (is-zero-div-zero-ℤ (x -ℤ y) K) is-unit-cong-succ-ℤ : (k x : ℤ) → cong-ℤ k x (succ-ℤ x) → is-unit-ℤ k -pr1 (is-unit-cong-succ-ℤ k x (pair y p)) = neg-ℤ y -pr2 (is-unit-cong-succ-ℤ k x (pair y p)) = - ( left-negative-law-mul-ℤ y k) ∙ - ( is-injective-neg-ℤ - ( ( neg-neg-ℤ (y *ℤ k)) ∙ - ( ( p) ∙ - ( ( ap (x +ℤ_) (neg-succ-ℤ x)) ∙ - ( ( right-predecessor-law-add-ℤ x (neg-ℤ x)) ∙ - ( ap pred-ℤ (right-inverse-law-add-ℤ x))))))) +is-unit-cong-succ-ℤ k x (y , p) = + is-unit-div-one-ℤ k + ( neg-ℤ y , + ( left-negative-law-mul-ℤ y k) ∙ + ( is-injective-neg-ℤ + ( ( neg-neg-ℤ (y *ℤ k)) ∙ + ( ( p) ∙ + ( ( ap (x +ℤ_) (neg-succ-ℤ x)) ∙ + ( ( right-predecessor-law-add-ℤ x (neg-ℤ x)) ∙ + ( ap pred-ℤ (right-inverse-law-add-ℤ x)))))))) is-unit-cong-pred-ℤ : (k x : ℤ) → cong-ℤ k x (pred-ℤ x) → is-unit-ℤ k -pr1 (is-unit-cong-pred-ℤ k x (pair y p)) = y -pr2 (is-unit-cong-pred-ℤ k x (pair y p)) = - ( p) ∙ - ( ( ap (x +ℤ_) (neg-pred-ℤ x)) ∙ - ( ( right-successor-law-add-ℤ x (neg-ℤ x)) ∙ - ( ap succ-ℤ (right-inverse-law-add-ℤ x)))) +is-unit-cong-pred-ℤ k x (y , p) = + is-unit-div-one-ℤ k + ( y , + ( p) ∙ + ( ( ap (x +ℤ_) (neg-pred-ℤ x)) ∙ + ( ( right-successor-law-add-ℤ x (neg-ℤ x)) ∙ + ( ap succ-ℤ (right-inverse-law-add-ℤ x))))) refl-cong-ℤ : (k : ℤ) → is-reflexive (cong-ℤ k) pr1 (refl-cong-ℤ k x) = zero-ℤ diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 9fb6209303..3e6e0117e7 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -157,6 +157,16 @@ is-zero-div-zero-ℤ : is-zero-div-zero-ℤ x (pair d p) = inv p ∙ right-zero-law-mul-ℤ d ``` +### An integer is a unit if and only if it divides $1$ + +```agda +div-one-is-unit-ℤ : (x : ℤ) → is-unit-ℤ x → div-ℤ x one-ℤ +div-one-is-unit-ℤ x (d , p , q) = (d , q) + +is-unit-div-one-ℤ : (x : ℤ) → div-ℤ x one-ℤ → is-unit-ℤ x +is-unit-div-one-ℤ x (d , p) = (d , commutative-mul-ℤ x d ∙ p , p) +``` + ### The quotient of `x` by one is `x` ```agda @@ -337,13 +347,14 @@ antisymmetric-div-ℤ x y (pair d p) (pair e q) H = f : is-decidable (is-zero-ℤ x) → presim-unit-ℤ x y f (inl refl) = presim-unit-eq-ℤ (inv (right-zero-law-mul-ℤ d) ∙ p) pr1 (pr1 (f (inr g))) = d - pr1 (pr2 (pr1 (f (inr g)))) = e - pr2 (pr2 (pr1 (f (inr g)))) = - is-injective-left-mul-ℤ x g - ( ( commutative-mul-ℤ x (e *ℤ d)) ∙ - ( ( associative-mul-ℤ e d x) ∙ - ( ( ap (e *ℤ_) p) ∙ - ( q ∙ inv (right-unit-law-mul-ℤ x))))) + pr2 (pr1 (f (inr g))) = + is-unit-div-one-ℤ d + ( e , + is-injective-left-mul-ℤ x g + ( ( commutative-mul-ℤ x (e *ℤ d)) ∙ + ( ( associative-mul-ℤ e d x) ∙ + ( ( ap (e *ℤ_) p) ∙ + ( q ∙ inv (right-unit-law-mul-ℤ x)))))) pr2 (f (inr g)) = p ``` diff --git a/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md b/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md index 42462305bf..9a1397032f 100644 --- a/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md +++ b/src/elementary-number-theory/greatest-common-divisor-integers.lagda.md @@ -17,6 +17,8 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.nonnegative-integers open import elementary-number-theory.positive-and-negative-integers open import elementary-number-theory.positive-integers +open import elementary-number-theory.unit-integers +open import elementary-number-theory.unit-similarity-integers open import foundation.action-on-identifications-functions open import foundation.cartesian-product-types @@ -316,7 +318,7 @@ is-commutative-gcd-ℤ x y = is-one-is-gcd-one-ℤ : {b x : ℤ} → is-gcd-ℤ one-ℤ b x → is-one-ℤ x is-one-is-gcd-one-ℤ {b} {x} H with ( is-one-or-neg-one-is-unit-ℤ x - ( pr1 (is-common-divisor-is-gcd-ℤ one-ℤ b x H))) + ( is-unit-div-one-ℤ x (pr1 (is-common-divisor-is-gcd-ℤ one-ℤ b x H)))) ... | inl p = p ... | inr p = ex-falso (tr is-nonnegative-ℤ p (pr1 H)) @@ -330,7 +332,7 @@ is-one-gcd-one-ℤ b = is-one-is-gcd-one-ℤ (is-gcd-gcd-ℤ one-ℤ b) is-one-is-gcd-one-ℤ' : {a x : ℤ} → is-gcd-ℤ a one-ℤ x → is-one-ℤ x is-one-is-gcd-one-ℤ' {a} {x} H with ( is-one-or-neg-one-is-unit-ℤ x - ( pr2 (is-common-divisor-is-gcd-ℤ a one-ℤ x H))) + ( is-unit-div-one-ℤ x (pr2 (is-common-divisor-is-gcd-ℤ a one-ℤ x H)))) ... | inl p = p ... | inr p = ex-falso (tr is-nonnegative-ℤ p (pr1 H)) diff --git a/src/elementary-number-theory/modular-arithmetic.lagda.md b/src/elementary-number-theory/modular-arithmetic.lagda.md index b03ca3add5..f2c68101c3 100644 --- a/src/elementary-number-theory/modular-arithmetic.lagda.md +++ b/src/elementary-number-theory/modular-arithmetic.lagda.md @@ -20,6 +20,7 @@ open import elementary-number-theory.multiplication-integers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.nonnegative-integers +open import elementary-number-theory.unit-integers open import foundation.action-on-identifications-binary-functions open import foundation.action-on-identifications-functions diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index ca96303564..ef4d4ad469 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -13,6 +13,7 @@ open import elementary-number-theory.divisibility-integers open import elementary-number-theory.integers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.unit-similarity-integers open import foundation.decidable-types open import foundation.negation diff --git a/src/elementary-number-theory/reduced-integer-fractions.lagda.md b/src/elementary-number-theory/reduced-integer-fractions.lagda.md index aa543cd985..8a63db9ec7 100644 --- a/src/elementary-number-theory/reduced-integer-fractions.lagda.md +++ b/src/elementary-number-theory/reduced-integer-fractions.lagda.md @@ -19,6 +19,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.positive-and-negative-integers open import elementary-number-theory.positive-integers open import elementary-number-theory.relatively-prime-integers +open import elementary-number-theory.unit-similarity-integers open import foundation.action-on-identifications-functions open import foundation.coproduct-types diff --git a/src/elementary-number-theory/unit-integers.lagda.md b/src/elementary-number-theory/unit-integers.lagda.md index c3860d33b7..e6775a28ba 100644 --- a/src/elementary-number-theory/unit-integers.lagda.md +++ b/src/elementary-number-theory/unit-integers.lagda.md @@ -31,13 +31,20 @@ open import foundation.universe-levels ## Idea -An [integer](elementary-number-theory.integers.md) $a$ is said to be a {{#concept "unit" Disambiguation="integer" Agda=unit-ℤ WD="unit" WDID=Q118084}} if it [divides](elementary-number-theory.divisibility-integers.md) the integer $1$, i.e., if there exists an integer $b$ such that +An [integer](elementary-number-theory.integers.md) $a$ is said to be a +{{#concept "unit" Disambiguation="integer" Agda=unit-ℤ WD="unit" WDID=Q118084}} +if it [divides](elementary-number-theory.divisibility-integers.md) the integer +$1$, i.e., if there exists an integer $b$ such that $$ ab=1. $$ -Note that the predicate of being a unit is expressed without explicitly referring to divisibility of the integers in order to avoid cyclic dependencies in the library. Some properties of divisibility involving unit integers are studied in [`elementary-number-theory.divisibility-integers`](elementary-number-theory.divisibility-integers.md`). +Note that the predicate of being a unit is expressed without explicitly +referring to divisibility of the integers in order to avoid cyclic dependencies +in the library. Some properties of divisibility involving unit integers are +studied in +[`elementary-number-theory.divisibility-integers`](elementary-number-theory.divisibility-integers.md`). ## Definitions diff --git a/src/elementary-number-theory/unit-similarity-integers.lagda.md b/src/elementary-number-theory/unit-similarity-integers.lagda.md index c43aefff36..70713851a1 100644 --- a/src/elementary-number-theory/unit-similarity-integers.lagda.md +++ b/src/elementary-number-theory/unit-similarity-integers.lagda.md @@ -17,8 +17,8 @@ open import foundation.action-on-identifications-functions open import foundation.binary-relations open import foundation.cartesian-product-types open import foundation.coproduct-types -open import foundation.dependent-pair-types open import foundation.decidable-types +open import foundation.dependent-pair-types open import foundation.function-types open import foundation.identity-types open import foundation.negation @@ -30,15 +30,20 @@ open import foundation.universe-levels ## Idea -Two [integers](elementary-number-theory.integers.md) $a$ and $b$ are said to be {{#concept "unit similar" Disambiguation="integers" Agda=sim-unit-ℤ}} if +Two [integers](elementary-number-theory.integers.md) $a$ and $b$ are said to be +{{#concept "unit similar" Disambiguation="integers" Agda=sim-unit-ℤ}} if ```text ¬ (is-nonzero-ℤ a × is-nonzero-ℤ b) → Σ (u : unit-ℤ) ua=b. ``` -The idea is that $a$ and $b$ are unit similar if there exists a [unit integer](elementary-number-theory.unit-integers.md) $u$ such that $ua=b$. However, this type is not a proposition if both $a$ and $b$ are $0$. The unit similarity relation has therefore a slightly more complicated expression. +The idea is that $a$ and $b$ are unit similar if there exists a +[unit integer](elementary-number-theory.unit-integers.md) $u$ such that $ua=b$. +However, this type is not a proposition if both $a$ and $b$ are $0$. The unit +similarity relation has therefore a slightly more complicated expression. -The unit similarity relation is an [equivalence relation](foundation.equivalence-relations.md) on the integers. +The unit similarity relation is an +[equivalence relation](foundation.equivalence-relations.md) on the integers. ## Definitions @@ -168,4 +173,3 @@ transitive-sim-unit-ℤ x y z K H f = ( K (λ (p , q) → f (is-zero-sim-unit-ℤ' H p , q))) ( H (λ (p , q) → f (p , is-zero-sim-unit-ℤ K q))) ``` - diff --git a/src/ring-theory/groups-of-units-rings.lagda.md b/src/ring-theory/groups-of-units-rings.lagda.md index 38b15cb9d0..15924cbe35 100644 --- a/src/ring-theory/groups-of-units-rings.lagda.md +++ b/src/ring-theory/groups-of-units-rings.lagda.md @@ -34,8 +34,9 @@ open import ring-theory.rings ## Idea -The {{#concept "group of units" Disambiguation="ring" WD="unit" WDID=Q118084}} of a [ring](ring-theory.rings.md) `R` is the -[group](group-theory.groups.md) consisting of all the +The {{#concept "group of units" Disambiguation="ring" WD="unit" WDID=Q118084}} +of a [ring](ring-theory.rings.md) `R` is the [group](group-theory.groups.md) +consisting of all the [invertible elements](ring-theory.invertible-elements-rings.md) in `R`. Equivalently, the group of units of `R` is the [core](group-theory.cores-monoids.md) of the multiplicative @@ -284,7 +285,7 @@ preserves-id-functor-Large-Precategory module _ {l : Level} (R : Ring l) where - + neg-group-of-units-Ring : type-group-of-units-Ring R → type-group-of-units-Ring R pr1 (neg-group-of-units-Ring (x , H)) = neg-Ring R x From 22a5ace333974a743a2767bfe2a6354a9529c665 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 30 Oct 2024 12:38:05 -0400 Subject: [PATCH 15/72] edits --- .../divisibility-integers.lagda.md | 144 ++++++------------ .../parity-natural-numbers.lagda.md | 4 +- .../unit-similarity-integers.lagda.md | 62 ++++++++ src/literature.lagda.md | 10 +- 4 files changed, 113 insertions(+), 107 deletions(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 3e6e0117e7..65ec923a5f 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -43,8 +43,10 @@ open import foundation.universe-levels ## Idea +On this page we define the {{#concept "divisibility" Disambiguation="integers" Agda=div-ℤ WD="divisibility" WDID=Q5284415}} relation on the [integers](elementary-number-theory.integers.md). + An integer `m` is said to **divide** an integer `n` if there exists an integer -`k` equipped with an identification `km = n`. Using the +`k` equipped with an [identification](foundation-core.identity-types.md) `km = n`. Using the [Curry–Howard interpretation](https://en.wikipedia.org/wiki/Curry–Howard_correspondence) of logic into type theory, we express divisibility as follows: @@ -52,17 +54,13 @@ of logic into type theory, we express divisibility as follows: div-ℤ m n := Σ (k : ℤ), k *ℤ m = n. ``` -If `n` is a nonzero integer, then `div-ℤ m n` is always a proposition in the +If `n` is a [nonzero integer](elementary-number-theory.nonzero-integers.md), then `div-ℤ m n` is always a [proposition](foundation-core.propositions.md) in the sense that the type `div-ℤ m n` contains at most one element. -We also introduce **unit integers**, which are integers that divide the integer -`1`, and an equivalence relation `sim-unit-ℤ` on the integers in which two -integers `x` and `y` are equivalent if there exists a unit integer `u` such that -`ux = y`. These two concepts help establish further properties of the -divisibility relation on the integers. - ## Definitions +### Divisibility of the integers + ```agda div-ℤ : ℤ → ℤ → UU lzero div-ℤ d x = Σ ℤ (λ k → k *ℤ d = x) @@ -99,14 +97,34 @@ concatenate-eq-div-eq-ℤ refl p refl = p ## Properties -### A unit integer divides every integer +### A natural number `m` divides a natural number `n` if and only if `m` divides `n` as integers ```agda -div-is-unit-ℤ : - (x y : ℤ) → is-unit-ℤ x → div-ℤ x y -pr1 (div-is-unit-ℤ x y (d , p , q)) = y *ℤ d -pr2 (div-is-unit-ℤ x y (d , p , q)) = - associative-mul-ℤ y d x ∙ (ap (y *ℤ_) q ∙ right-unit-law-mul-ℤ y) +div-int-div-ℕ : + {m n : ℕ} → div-ℕ m n → div-ℤ (int-ℕ m) (int-ℕ n) +pr1 (div-int-div-ℕ {m} {n} (pair d p)) = int-ℕ d +pr2 (div-int-div-ℕ {m} {n} (pair d p)) = mul-int-ℕ d m ∙ ap int-ℕ p + +div-div-int-ℕ : + {m n : ℕ} → div-ℤ (int-ℕ m) (int-ℕ n) → div-ℕ m n +div-div-int-ℕ {zero-ℕ} {n} (pair d p) = + div-eq-ℕ zero-ℕ n + ( inv (is-injective-int-ℕ (inv p ∙ right-zero-law-mul-ℤ d))) +pr1 (div-div-int-ℕ {succ-ℕ m} {n} (pair d p)) = abs-ℤ d +pr2 (div-div-int-ℕ {succ-ℕ m} {n} (pair d p)) = + is-injective-int-ℕ + ( ( inv (mul-int-ℕ (abs-ℤ d) (succ-ℕ m))) ∙ + ( ( ap + ( _*ℤ (inr (inr m))) + { int-abs-ℤ d} + { d} + ( int-abs-is-nonnegative-ℤ d + ( is-nonnegative-left-factor-mul-ℤ + { d} + { inr (inr m)} + ( is-nonnegative-eq-ℤ (inv p) (is-nonnegative-int-ℕ n)) + ( star)))) ∙ + ( p))) ``` ### Divisibility by a nonzero integer is a property @@ -157,6 +175,16 @@ is-zero-div-zero-ℤ : is-zero-div-zero-ℤ x (pair d p) = inv p ∙ right-zero-law-mul-ℤ d ``` +### A unit integer divides every integer + +```agda +div-is-unit-ℤ : + (x y : ℤ) → is-unit-ℤ x → div-ℤ x y +pr1 (div-is-unit-ℤ x y (d , p , q)) = y *ℤ d +pr2 (div-is-unit-ℤ x y (d , p , q)) = + associative-mul-ℤ y d x ∙ (ap (y *ℤ_) q ∙ right-unit-law-mul-ℤ y) +``` + ### An integer is a unit if and only if it divides $1$ ```agda @@ -306,36 +334,6 @@ div-div-quotient-div-ℤ x y d H K = ( preserves-div-mul-ℤ d x (quotient-div-ℤ d y H) K) ``` -### Comparison of divisibility on `ℕ` and on `ℤ` - -```agda -div-int-div-ℕ : - {x y : ℕ} → div-ℕ x y → div-ℤ (int-ℕ x) (int-ℕ y) -pr1 (div-int-div-ℕ {x} {y} (pair d p)) = int-ℕ d -pr2 (div-int-div-ℕ {x} {y} (pair d p)) = mul-int-ℕ d x ∙ ap int-ℕ p - -div-div-int-ℕ : - {x y : ℕ} → div-ℤ (int-ℕ x) (int-ℕ y) → div-ℕ x y -div-div-int-ℕ {zero-ℕ} {y} (pair d p) = - div-eq-ℕ zero-ℕ y - ( inv (is-injective-int-ℕ (is-zero-div-zero-ℤ (int-ℕ y) (pair d p)))) -pr1 (div-div-int-ℕ {succ-ℕ x} {y} (pair d p)) = abs-ℤ d -pr2 (div-div-int-ℕ {succ-ℕ x} {y} (pair d p)) = - is-injective-int-ℕ - ( ( inv (mul-int-ℕ (abs-ℤ d) (succ-ℕ x))) ∙ - ( ( ap - ( _*ℤ (inr (inr x))) - { int-abs-ℤ d} - { d} - ( int-abs-is-nonnegative-ℤ d - ( is-nonnegative-left-factor-mul-ℤ - { d} - { inr (inr x)} - ( is-nonnegative-eq-ℤ (inv p) (is-nonnegative-int-ℕ y)) - ( star)))) ∙ - ( p))) -``` - ### `sim-unit-ℤ x y` holds if and only if `x|y` and `y|x` ```agda @@ -358,15 +356,9 @@ antisymmetric-div-ℤ x y (pair d p) (pair e q) H = pr2 (f (inr g)) = p ``` -### `sim-unit-ℤ |x| x` holds +### If `x ~ x'` and `y ~ y'` are unit similar, then `x | y` holds if and only if `x' | y'` holds ```agda -sim-unit-abs-ℤ : (x : ℤ) → sim-unit-ℤ (int-abs-ℤ x) x -pr1 (sim-unit-abs-ℤ (inl x) f) = neg-one-unit-ℤ -pr2 (sim-unit-abs-ℤ (inl x) f) = refl -sim-unit-abs-ℤ (inr (inl star)) = refl-sim-unit-ℤ zero-ℤ -sim-unit-abs-ℤ (inr (inr x)) = refl-sim-unit-ℤ (inr (inr x)) - div-presim-unit-ℤ : {x y x' y' : ℤ} → presim-unit-ℤ x x' → presim-unit-ℤ y y' → div-ℤ x y → div-ℤ x' y' @@ -408,53 +400,3 @@ div-div-int-abs-divisor-ℤ : div-div-int-abs-divisor-ℤ {x} {y} = div-sim-unit-ℤ (sim-unit-abs-ℤ x) (refl-sim-unit-ℤ y) ``` - -### If we have that `sim-unit-ℤ x y`, then they must differ only by sign - -```agda -is-plus-or-minus-sim-unit-ℤ : - {x y : ℤ} → sim-unit-ℤ x y → is-plus-or-minus-ℤ x y -is-plus-or-minus-sim-unit-ℤ {x} {y} H with ( is-decidable-is-zero-ℤ x) -is-plus-or-minus-sim-unit-ℤ {x} {y} H | inl z = - inl (z ∙ inv (is-zero-sim-unit-ℤ H z)) -is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz - with - ( is-one-or-neg-one-is-unit-ℤ - ( int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) - ( is-unit-unit-ℤ (pr1 (H (λ u → nz (pr1 u)))))) -is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz | inl pos = - inl - ( equational-reasoning - x - = one-ℤ *ℤ x - by (inv (left-unit-law-mul-ℤ x)) - = (int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) *ℤ x - by inv (ap (_*ℤ x) pos) - = y - by pr2 (H (λ u → nz (pr1 u)))) -is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz | inr p = - inr - ( equational-reasoning - neg-ℤ x - = (int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) *ℤ x - by ap (_*ℤ x) (inv p) - = y - by pr2 (H (λ u → nz (pr1 u)))) -``` - -### If `sim-unit-ℤ x y` holds and both `x` and `y` have the same sign, then `x = y` - -```agda -eq-sim-unit-is-nonnegative-ℤ : - {a b : ℤ} → is-nonnegative-ℤ a → is-nonnegative-ℤ b → sim-unit-ℤ a b → a = b -eq-sim-unit-is-nonnegative-ℤ {a} {b} H H' K = - rec-coproduct - ( id) - ( λ K' → - eq-is-zero-ℤ - ( is-zero-is-nonnegative-neg-is-nonnegative-ℤ H - ( is-nonnegative-eq-ℤ (inv K') H')) - ( is-zero-is-nonnegative-neg-is-nonnegative-ℤ H' - ( is-nonnegative-eq-ℤ (inv (neg-neg-ℤ a) ∙ ap neg-ℤ K') H))) - ( is-plus-or-minus-sim-unit-ℤ K) -``` diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index e2e0de1b25..7adb25d863 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -32,9 +32,9 @@ open import foundation.universe-levels [partitions](foundation.partitions.md) the [natural numbers](elementary-number-theory.natural numbers.md) into two [classes](foundation.equivalence-relations.md): the -{{#concept "even" Disambiguation="natural number" Agda=is-even-ℤ WD="even number" WDID=Q13366104}} +{{#concept "even" Disambiguation="natural number" Agda=is-even-ℕ WD="even number" WDID=Q13366104}} and the -{{#concept "odd" Disambiguation="natural number" Agda=is-odd-ℤ WD="odd number" WDID=Q13366129}} +{{#concept "odd" Disambiguation="natural number" Agda=is-odd-ℕ WD="odd number" WDID=Q13366129}} natural numbers. Even natural numbers are those that are [divisible](elementary-number-theory.divisibility-natural numbers.md) by two, and odd natural numbers are those that aren't. diff --git a/src/elementary-number-theory/unit-similarity-integers.lagda.md b/src/elementary-number-theory/unit-similarity-integers.lagda.md index 70713851a1..c5b26b74e1 100644 --- a/src/elementary-number-theory/unit-similarity-integers.lagda.md +++ b/src/elementary-number-theory/unit-similarity-integers.lagda.md @@ -7,9 +7,11 @@ module elementary-number-theory.unit-similarity-integers where
Imports ```agda +open import elementary-number-theory.absolute-value-integers open import elementary-number-theory.equality-integers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers +open import elementary-number-theory.nonnegative-integers open import elementary-number-theory.nonzero-integers open import elementary-number-theory.unit-integers @@ -173,3 +175,63 @@ transitive-sim-unit-ℤ x y z K H f = ( K (λ (p , q) → f (is-zero-sim-unit-ℤ' H p , q))) ( H (λ (p , q) → f (p , is-zero-sim-unit-ℤ K q))) ``` + +### `sim-unit-ℤ |x| x` holds + +```agda +sim-unit-abs-ℤ : (x : ℤ) → sim-unit-ℤ (int-abs-ℤ x) x +pr1 (sim-unit-abs-ℤ (inl x) f) = neg-one-unit-ℤ +pr2 (sim-unit-abs-ℤ (inl x) f) = refl +sim-unit-abs-ℤ (inr (inl star)) = refl-sim-unit-ℤ zero-ℤ +sim-unit-abs-ℤ (inr (inr x)) = refl-sim-unit-ℤ (inr (inr x)) +``` + +### If we have that `sim-unit-ℤ x y`, then they must differ only by sign + +```agda +is-plus-or-minus-sim-unit-ℤ : + {x y : ℤ} → sim-unit-ℤ x y → is-plus-or-minus-ℤ x y +is-plus-or-minus-sim-unit-ℤ {x} {y} H with ( is-decidable-is-zero-ℤ x) +is-plus-or-minus-sim-unit-ℤ {x} {y} H | inl z = + inl (z ∙ inv (is-zero-sim-unit-ℤ H z)) +is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz + with + ( is-one-or-neg-one-is-unit-ℤ + ( int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) + ( is-unit-unit-ℤ (pr1 (H (λ u → nz (pr1 u)))))) +is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz | inl pos = + inl + ( equational-reasoning + x + = one-ℤ *ℤ x + by (inv (left-unit-law-mul-ℤ x)) + = (int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) *ℤ x + by inv (ap (_*ℤ x) pos) + = y + by pr2 (H (λ u → nz (pr1 u)))) +is-plus-or-minus-sim-unit-ℤ {x} {y} H | inr nz | inr p = + inr + ( equational-reasoning + neg-ℤ x + = (int-unit-ℤ (pr1 (H (λ u → nz (pr1 u))))) *ℤ x + by ap (_*ℤ x) (inv p) + = y + by pr2 (H (λ u → nz (pr1 u)))) +``` + +### If `sim-unit-ℤ x y` holds and both `x` and `y` have the same sign, then `x = y` + +```agda +eq-sim-unit-is-nonnegative-ℤ : + {a b : ℤ} → is-nonnegative-ℤ a → is-nonnegative-ℤ b → sim-unit-ℤ a b → a = b +eq-sim-unit-is-nonnegative-ℤ {a} {b} H H' K = + rec-coproduct + ( id) + ( λ K' → + eq-is-zero-ℤ + ( is-zero-is-nonnegative-neg-is-nonnegative-ℤ H + ( is-nonnegative-eq-ℤ (inv K') H')) + ( is-zero-is-nonnegative-neg-is-nonnegative-ℤ H' + ( is-nonnegative-eq-ℤ (inv (neg-neg-ℤ a) ∙ ap neg-ℤ K') H))) + ( is-plus-or-minus-sim-unit-ℤ K) +``` diff --git a/src/literature.lagda.md b/src/literature.lagda.md index 2164beb6a2..b075a0a16c 100644 --- a/src/literature.lagda.md +++ b/src/literature.lagda.md @@ -1,8 +1,10 @@ -# Formalization of results from the literature +# Formalization of results from the -> This page is currently under construction. To see what's happening behind the -> scenes, see the associated GitHub issue -> [#1055](https://github.com/UniMath/agda-unimath/issues/1055). +On this page we highlight some publications that have been formalized in +agda-unimath. We are still working on giving a more complete overview of the +papers we have formalized. To what's happening behind the scenes, see the +associated GitHub issue +[#1055](https://github.com/UniMath/agda-unimath/issues/1055). ## Modules in the literature namespace From aea9973bb82cf42d5c2b0120cc7d592903aa6b2a Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 30 Oct 2024 12:38:24 -0400 Subject: [PATCH 16/72] edits --- .../divisibility-integers.lagda.md | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 65ec923a5f..2afc5710b3 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -43,10 +43,13 @@ open import foundation.universe-levels ## Idea -On this page we define the {{#concept "divisibility" Disambiguation="integers" Agda=div-ℤ WD="divisibility" WDID=Q5284415}} relation on the [integers](elementary-number-theory.integers.md). +On this page we define the +{{#concept "divisibility" Disambiguation="integers" Agda=div-ℤ WD="divisibility" WDID=Q5284415}} +relation on the [integers](elementary-number-theory.integers.md). An integer `m` is said to **divide** an integer `n` if there exists an integer -`k` equipped with an [identification](foundation-core.identity-types.md) `km = n`. Using the +`k` equipped with an [identification](foundation-core.identity-types.md) +`km = n`. Using the [Curry–Howard interpretation](https://en.wikipedia.org/wiki/Curry–Howard_correspondence) of logic into type theory, we express divisibility as follows: @@ -54,8 +57,9 @@ of logic into type theory, we express divisibility as follows: div-ℤ m n := Σ (k : ℤ), k *ℤ m = n. ``` -If `n` is a [nonzero integer](elementary-number-theory.nonzero-integers.md), then `div-ℤ m n` is always a [proposition](foundation-core.propositions.md) in the -sense that the type `div-ℤ m n` contains at most one element. +If `n` is a [nonzero integer](elementary-number-theory.nonzero-integers.md), +then `div-ℤ m n` is always a [proposition](foundation-core.propositions.md) in +the sense that the type `div-ℤ m n` contains at most one element. ## Definitions From c0dfde5483d8f75f6cb497ba46ec533cb3c32ac4 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 30 Oct 2024 18:56:19 -0400 Subject: [PATCH 17/72] work --- .../parity-integers.lagda.md | 36 ++++++++++--------- .../unit-integers.lagda.md | 32 +++++++++++++++++ 2 files changed, 52 insertions(+), 16 deletions(-) diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index ef4d4ad469..f8009e2b92 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -15,6 +15,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.unit-similarity-integers +open import foundation.coproduct-types open import foundation.decidable-types open import foundation.negation open import foundation.universe-levels @@ -142,7 +143,7 @@ is-odd-one-ℤ : is-odd-ℤ (int-ℕ 1) is-odd-one-ℤ = is-odd-int-is-odd-ℕ 1 is-odd-one-ℕ ``` -### A integer `x` is even if and only if `x + 2` is even +### An integer `x` is even if and only if `x + 2` is even ```agda is-even-is-even-add-two-ℤ : @@ -156,7 +157,7 @@ is-even-add-two-is-even-ℤ a H = div-add-ℤ (int-ℕ 2) a (int-ℕ 2) H (refl-div-ℤ (int-ℕ 2)) ``` -### A integer `x` is odd if and only if `x + 2` is odd +### An integer `x` is odd if and only if `x + 2` is odd ```agda is-odd-is-odd-add-two-ℤ : (a : ℤ) → is-odd-ℤ (a +ℤ int-ℕ 2) → is-odd-ℤ a @@ -166,20 +167,23 @@ is-odd-add-two-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → is-odd-ℤ (a +ℤ in is-odd-add-two-is-odd-ℤ a H K = H (is-even-is-even-add-two-ℤ a K) ``` -### If a integer `x` is odd, then `x + 1` is even +### If an integer `x` is even, then `x + 1` is odd -```text -is-even-succ-is-odd-ℕ : - (n : ℕ) → is-odd-ℕ n → is-even-ℕ (succ-ℕ n) -is-even-succ-is-odd-ℕ zero-ℕ p = ex-falso (p is-even-zero-ℕ) -is-even-succ-is-odd-ℕ (succ-ℕ zero-ℕ) p = (1 , refl) -is-even-succ-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = - is-even-succ-succ-is-even-ℕ - ( succ-ℕ n) - ( is-even-succ-is-odd-ℕ n (is-odd-is-odd-succ-succ-ℕ n p)) +```agda +is-odd-succ-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → is-odd-ℤ (a +ℤ one-ℤ) +is-odd-succ-is-even-ℤ a H K = {!!} +``` + +### If an integer `x` is odd, then `x + 1` is even + +```agda +is-even-succ-is-odd-ℤ : + (a : ℤ) → is-odd-ℤ a → is-even-ℤ (a +ℤ one-ℤ) +is-even-succ-is-odd-ℤ a H = {!a!} ``` -### If a integer `x` is even, then `x + 1` is odd +### If an integer `x` is even, then `x + 1` is odd ```text is-odd-succ-is-even-ℕ : @@ -192,7 +196,7 @@ is-odd-succ-is-even-ℕ (succ-ℕ (succ-ℕ n)) p = ( is-odd-succ-is-even-ℕ n (is-even-is-even-succ-succ-ℕ n p)) ``` -### If a integer `x + 1` is odd, then `x` is even +### If an integer `x + 1` is odd, then `x` is even ```text is-even-is-odd-succ-ℕ : @@ -203,7 +207,7 @@ is-even-is-odd-succ-ℕ n p = ( is-even-succ-is-odd-ℕ (succ-ℕ n) p) ``` -### If a integer `x + 1` is even, then `x` is odd +### If an integer `x + 1` is even, then `x` is odd ```text is-odd-is-even-succ-ℕ : @@ -214,7 +218,7 @@ is-odd-is-even-succ-ℕ n p = ( is-odd-succ-is-even-ℕ (succ-ℕ n) p) ``` -### A integer `x` is odd if and only if there is a integer `y` such that `succ-ℕ (y *ℕ 2) = x` +### An integer `x` is odd if and only if there is an integer `y` such that `succ-ℕ (y *ℕ 2) = x` ```text has-odd-expansion : ℕ → UU lzero diff --git a/src/elementary-number-theory/unit-integers.lagda.md b/src/elementary-number-theory/unit-integers.lagda.md index e6775a28ba..a7ffb3d4df 100644 --- a/src/elementary-number-theory/unit-integers.lagda.md +++ b/src/elementary-number-theory/unit-integers.lagda.md @@ -22,6 +22,7 @@ open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.identity-types +open import foundation.negation open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels @@ -54,6 +55,9 @@ studied in is-unit-ℤ : ℤ → UU lzero is-unit-ℤ = is-invertible-element-Commutative-Ring ℤ-Commutative-Ring +is-nonunit-ℤ : ℤ → UU lzero +is-nonunit-ℤ a = ¬ is-unit-ℤ a + unit-ℤ : UU lzero unit-ℤ = type-group-of-units-Commutative-Ring ℤ-Commutative-Ring @@ -166,6 +170,34 @@ is-one-or-neg-one-is-unit-ℤ ( inv q ∙ compute-mul-ℤ (inr (inr (succ-ℕ d))) (inr (inr (succ-ℕ x))))) ``` +### Some examples of integers that aren't units + +```agda +abstract + is-nonunit-zero-ℤ : is-nonunit-ℤ zero-ℤ + is-nonunit-zero-ℤ H + with + is-one-or-neg-one-is-unit-ℤ zero-ℤ H + ... | inl () + ... | inr () + +abstract + is-nonunit-two-ℤ : is-nonunit-ℤ (int-ℕ 2) + is-nonunit-two-ℤ H + with + is-one-or-neg-one-is-unit-ℤ (int-ℕ 2) H + ... | inl () + ... | inr () + +abstract + is-nonunit-three-ℤ : is-nonunit-ℤ (int-ℕ 3) + is-nonunit-three-ℤ H + with + is-one-or-neg-one-is-unit-ℤ (int-ℕ 3) H + ... | inl () + ... | inr () +``` + ### Unit integers are idempotent ```agda From fbdb94c61dc40d05b7c67ff034f01b52cfba561c Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Thu, 31 Oct 2024 17:22:51 -0400 Subject: [PATCH 18/72] moving decidability of divisibility into the file about divisibility --- .../divisibility-natural-numbers.lagda.md | 290 +++++++++++++----- ...fundamental-theorem-of-arithmetic.lagda.md | 2 +- ...-arithmetic-standard-finite-types.lagda.md | 23 -- .../parity-integers.lagda.md | 17 +- .../proper-divisors-natural-numbers.lagda.md | 5 +- 5 files changed, 236 insertions(+), 101 deletions(-) diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 9bf81abc22..153d9cc343 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -8,7 +8,9 @@ module elementary-number-theory.divisibility-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.decidable-types open import elementary-number-theory.distance-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -16,15 +18,23 @@ open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-propositions +open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.equivalences open import foundation.identity-types open import foundation.logical-equivalences open import foundation.negated-equality open import foundation.negation open import foundation.propositional-maps open import foundation.propositions +open import foundation.sets +open import foundation.subtypes open import foundation.transport-along-identifications +open import foundation.type-arithmetic-dependent-pair-types open import foundation.type-arithmetic-empty-type open import foundation.unit-type open import foundation.universe-levels @@ -34,8 +44,9 @@ open import foundation.universe-levels ## Idea +On this page we define the {{#concept "divisibility" Disambiguation="natural numbers" Agda=div-ℕ WD="divisibility" WDID=Q5284415}} relation on the [natural numbers](elementary-number-theory.natural-numbers.md). A natural number `m` is said to **divide** a natural number `n` if there exists -a natural number `k` equipped with an identification `km = n`. Using the +a natural number `k` equipped with an [identification](foundation-core.identifications.md) `km = n`. Using the [Curry–Howard interpretation](https://en.wikipedia.org/wiki/Curry–Howard_correspondence) of logic into type theory, we express divisibility as follows: @@ -43,33 +54,76 @@ of logic into type theory, we express divisibility as follows: div-ℕ m n := Σ (k : ℕ), k *ℕ m = n. ``` -If `n` is a nonzero natural number, then `div-ℕ m n` is always a proposition in +If `n` is a [nonzero natural number](elementary-number-theory.nonzero-natural-numbers.md), then the type `div-ℕ m n` is always a [proposition](foundation-core.propositions.md) in the sense that the type `div-ℕ m n` contains at most one element. +The divisibility relation is [logically equivalent](foundation.logical-equivalences.md), though not [equivalent](foundation-core.equivalences.md), to the +{{#concept "bounded divisibility" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} relation, which is defined by + +```text + bounded-div-ℕ m n := Σ (k : ℕ), (k ≤ n) × (k *ℕ m = n). +``` + +The discrepancy between divisibility and bounded divisibility is manifested at `(0 , 0)`. Note that `div-ℕ 0 0 ≃ ℕ`, while `bounded-div-ℕ 0 0` is [contractible](foundation-core.contractible-types.md). For all other values we have `div-ℕ m n ≃ bounded-div-ℕ m n`. + +The notion of bounded divisibility has several advantages over ordinary divisibility: Bounded divisibility is always a proposition, and the proof that divisibility is decidable uses the proof that bounded divisibility is decidable. + ## Definitions +### The divisibility relation on the natural numbers + ```agda -div-ℕ : ℕ → ℕ → UU lzero -div-ℕ m n = Σ ℕ (λ k → k *ℕ m = n) +module _ + (m n : ℕ) + where + + div-ℕ : UU lzero + div-ℕ = Σ ℕ (λ k → k *ℕ m = n) -quotient-div-ℕ : (x y : ℕ) → div-ℕ x y → ℕ -quotient-div-ℕ x y H = pr1 H + quotient-div-ℕ : div-ℕ → ℕ + quotient-div-ℕ = pr1 -eq-quotient-div-ℕ : - (x y : ℕ) (H : div-ℕ x y) → (quotient-div-ℕ x y H) *ℕ x = y -eq-quotient-div-ℕ x y H = pr2 H + eq-quotient-div-ℕ : (H : div-ℕ) → quotient-div-ℕ H *ℕ m = n + eq-quotient-div-ℕ = pr2 -eq-quotient-div-ℕ' : - (x y : ℕ) (H : div-ℕ x y) → x *ℕ (quotient-div-ℕ x y H) = y -eq-quotient-div-ℕ' x y H = - commutative-mul-ℕ x (quotient-div-ℕ x y H) ∙ eq-quotient-div-ℕ x y H + eq-quotient-div-ℕ' : (H : div-ℕ) → m *ℕ quotient-div-ℕ H = n + eq-quotient-div-ℕ' H = + commutative-mul-ℕ m (quotient-div-ℕ H) ∙ eq-quotient-div-ℕ H +``` -div-quotient-div-ℕ : - (d x : ℕ) (H : div-ℕ d x) → div-ℕ (quotient-div-ℕ d x H) x -pr1 (div-quotient-div-ℕ d x (u , p)) = d -pr2 (div-quotient-div-ℕ d x (u , p)) = commutative-mul-ℕ d u ∙ p +### The bounded divisibility predicate + +```agda +module _ + (m n : ℕ) + where + + bounded-div-ℕ : UU lzero + bounded-div-ℕ = Σ ℕ (λ k → (k ≤-ℕ n) × (k *ℕ m = n)) + + quotient-bounded-div-ℕ : bounded-div-ℕ → ℕ + quotient-bounded-div-ℕ = pr1 + + upper-bound-quotient-bounded-div-ℕ : + (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H ≤-ℕ n + upper-bound-quotient-bounded-div-ℕ H = pr1 (pr2 H) + + eq-quotient-bounded-div-ℕ : + (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H *ℕ m = n + eq-quotient-bounded-div-ℕ H = pr2 (pr2 H) + + eq-quotient-bounded-div-ℕ' : + (H : bounded-div-ℕ) → m *ℕ quotient-bounded-div-ℕ H = n + eq-quotient-bounded-div-ℕ' H = + commutative-mul-ℕ m (quotient-bounded-div-ℕ H) ∙ eq-quotient-bounded-div-ℕ H + + div-bounded-div-ℕ : bounded-div-ℕ → div-ℕ m n + pr1 (div-bounded-div-ℕ H) = quotient-bounded-div-ℕ H + pr2 (div-bounded-div-ℕ H) = eq-quotient-bounded-div-ℕ H ``` +## Properties + ### Concatenating equality and divisibility ```agda @@ -86,68 +140,31 @@ concatenate-eq-div-eq-ℕ : concatenate-eq-div-eq-ℕ refl p refl = p ``` -## Properties - -### The quotients of a natural number `n` by two natural numbers `p` and `q` are equal if `p` and `q` are equal +### The quotient of `n` by a number `m` it is divisible by also divides `n` ```agda -eq-quotient-div-eq-div-ℕ : - (x y z : ℕ) → is-nonzero-ℕ x → x = y → - (H : div-ℕ x z) → (I : div-ℕ y z) → - quotient-div-ℕ x z H = quotient-div-ℕ y z I -eq-quotient-div-eq-div-ℕ x y z n e H I = - is-injective-left-mul-ℕ - ( x) - ( n) - ( tr - ( λ p → - x *ℕ (quotient-div-ℕ x z H) = - p *ℕ (quotient-div-ℕ y z I)) - ( inv e) - ( commutative-mul-ℕ x (quotient-div-ℕ x z H) ∙ - ( eq-quotient-div-ℕ x z H ∙ - ( inv (eq-quotient-div-ℕ y z I) ∙ - commutative-mul-ℕ (quotient-div-ℕ y z I) y)))) +div-quotient-div-ℕ : + (m n : ℕ) (H : div-ℕ m n) → div-ℕ (quotient-div-ℕ m n H) n +pr1 (div-quotient-div-ℕ m n (u , p)) = m +pr2 (div-quotient-div-ℕ m n (u , p)) = commutative-mul-ℕ m u ∙ p ``` -### Divisibility by a nonzero natural number is a property +### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal ```agda -is-prop-div-ℕ : (k x : ℕ) → is-nonzero-ℕ k → is-prop (div-ℕ k x) -is-prop-div-ℕ k x f = is-prop-map-is-emb (is-emb-right-mul-ℕ k f) x -``` - -### The divisibility relation is a partial order on the natural numbers - -```agda -refl-div-ℕ : is-reflexive div-ℕ -pr1 (refl-div-ℕ x) = 1 -pr2 (refl-div-ℕ x) = left-unit-law-mul-ℕ x - -div-eq-ℕ : (x y : ℕ) → x = y → div-ℕ x y -div-eq-ℕ x .x refl = refl-div-ℕ x - -antisymmetric-div-ℕ : is-antisymmetric div-ℕ -antisymmetric-div-ℕ zero-ℕ zero-ℕ H K = refl -antisymmetric-div-ℕ zero-ℕ (succ-ℕ y) (pair k p) K = - inv (right-zero-law-mul-ℕ k) ∙ p -antisymmetric-div-ℕ (succ-ℕ x) zero-ℕ H (pair l q) = - inv q ∙ right-zero-law-mul-ℕ l -antisymmetric-div-ℕ (succ-ℕ x) (succ-ℕ y) (pair k p) (pair l q) = - ( inv (left-unit-law-mul-ℕ (succ-ℕ x))) ∙ - ( ( ap - ( _*ℕ (succ-ℕ x)) - ( inv - ( is-one-right-is-one-mul-ℕ l k - ( is-one-is-left-unit-mul-ℕ (l *ℕ k) x - ( ( associative-mul-ℕ l k (succ-ℕ x)) ∙ - ( ap (l *ℕ_) p ∙ q)))))) ∙ - ( p)) - -transitive-div-ℕ : is-transitive div-ℕ -pr1 (transitive-div-ℕ x y z (pair l q) (pair k p)) = l *ℕ k -pr2 (transitive-div-ℕ x y z (pair l q) (pair k p)) = - associative-mul-ℕ l k x ∙ (ap (l *ℕ_) p ∙ q) +eq-quotient-div-eq-div-ℕ : + (c d n : ℕ) → is-nonzero-ℕ c → c = d → + (H : div-ℕ c n) → (I : div-ℕ d n) → + quotient-div-ℕ c n H = quotient-div-ℕ d n I +eq-quotient-div-eq-div-ℕ c d n N p H I = + is-injective-left-mul-ℕ c N + ( tr + ( λ x → c *ℕ quotient-div-ℕ c n H = x *ℕ quotient-div-ℕ d n I) + ( inv p) + ( commutative-mul-ℕ c (quotient-div-ℕ c n H) ∙ + eq-quotient-div-ℕ c n H ∙ + inv (eq-quotient-div-ℕ d n I) ∙ + commutative-mul-ℕ (quotient-div-ℕ d n I) d)) ``` ### If `x` is nonzero and `d | x`, then `d ≤ x` @@ -186,6 +203,133 @@ le-div-ℕ : (d x : ℕ) → is-nonzero-ℕ x → div-ℕ d x → d ≠ x → le le-div-ℕ d x H K f = le-leq-neq-ℕ (leq-div-ℕ d x H K) f ``` +### Divisibility is logically equivalent to bounded divisibility + +```agda +bounded-div-div-ℕ : (m n : ℕ) → div-ℕ m n → bounded-div-ℕ m n +bounded-div-div-ℕ m zero-ℕ (q , p) = (0 , refl-leq-ℕ 0 , left-zero-law-mul-ℕ m) +bounded-div-div-ℕ m (succ-ℕ n) (q , p) = + ( q , leq-quotient-div-ℕ m (succ-ℕ n) (is-nonzero-succ-ℕ n) (q , p) , p) + +logical-equivalence-bounded-div-div-ℕ : + (m n : ℕ) → bounded-div-ℕ m n ↔ div-ℕ m n +pr1 (logical-equivalence-bounded-div-div-ℕ m n) = + div-bounded-div-ℕ m n +pr2 (logical-equivalence-bounded-div-div-ℕ m n) = + bounded-div-div-ℕ m n +``` + +### Bounded divisibility is decidable + +```agda +is-decidable-bounded-div-ℕ : + (m n : ℕ) → is-decidable (bounded-div-ℕ m n) +is-decidable-bounded-div-ℕ m n = + is-decidable-bounded-Σ-ℕ' + ( n) + ( λ x → mul-ℕ x m = n) + ( λ x → has-decidable-equality-ℕ (mul-ℕ x m) n) +``` + +### Divisibility is decidable + +```agda +is-decidable-div-ℕ : + (m n : ℕ) → is-decidable (div-ℕ m n) +is-decidable-div-ℕ m n = + is-decidable-iff + ( div-bounded-div-ℕ m n) + ( bounded-div-div-ℕ m n) + ( is-decidable-bounded-div-ℕ m n) +``` + +### Divisibility is a property except at `(0,0)` + +```agda +is-prop-div-ℕ : + (k x : ℕ) → is-nonzero-ℕ k + is-nonzero-ℕ x → is-prop (div-ℕ k x) +is-prop-div-ℕ k x (inl H) = is-prop-map-is-emb (is-emb-right-mul-ℕ k H) x +is-prop-div-ℕ zero-ℕ x (inr H) = + is-prop-is-proof-irrelevant + ( λ (q , p) → ex-falso (H (inv p ∙ right-zero-law-mul-ℕ q))) +is-prop-div-ℕ (succ-ℕ k) x (inr H) = + is-prop-map-is-emb (is-emb-right-mul-ℕ (succ-ℕ k) (is-nonzero-succ-ℕ k)) x + +div-ℕ-Decidable-Prop : (d x : ℕ) → is-nonzero-ℕ d → Decidable-Prop lzero +pr1 (div-ℕ-Decidable-Prop d x H) = div-ℕ d x +pr1 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-prop-div-ℕ d x (inl H) +pr2 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-decidable-div-ℕ d x +``` + +### Bounded divisibility is a property + +```agda +abstract + is-prop-bounded-div-ℕ : + (k n : ℕ) → is-prop (bounded-div-ℕ k n) + is-prop-bounded-div-ℕ zero-ℕ n = + is-prop-all-elements-equal + ( λ (q , b , p) (q' , b' , p') → + eq-type-subtype + ( λ u → product-Prop (leq-ℕ-Prop u n) (Id-Prop ℕ-Set _ _)) + ( is-zero-leq-zero-ℕ q + ( concatenate-leq-eq-ℕ q b (inv p ∙ right-zero-law-mul-ℕ q)) ∙ + inv + ( is-zero-leq-zero-ℕ q' + ( concatenate-leq-eq-ℕ q' b' + ( inv p' ∙ right-zero-law-mul-ℕ q'))))) + is-prop-bounded-div-ℕ (succ-ℕ k) n = + is-prop-equiv + ( equiv-right-swap-Σ ∘e inv-associative-Σ ℕ (λ x → x ≤-ℕ n) _) + ( is-prop-Σ + ( is-prop-div-ℕ (succ-ℕ k) n (inl (is-nonzero-succ-ℕ k))) + ( λ (x , H) → is-prop-leq-ℕ x n)) + +bounded-div-ℕ-Prop : (k n : ℕ) → Prop lzero +pr1 (bounded-div-ℕ-Prop k n) = bounded-div-ℕ k n +pr2 (bounded-div-ℕ-Prop k n) = is-prop-bounded-div-ℕ k n + +bounded-div-ℕ-Decidable-Prop : (k n : ℕ) → Decidable-Prop lzero +pr1 (bounded-div-ℕ-Decidable-Prop k n) = bounded-div-ℕ k n +pr1 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-prop-bounded-div-ℕ k n +pr2 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-decidable-bounded-div-ℕ k n +``` + +### The divisibility relation is a partial order on the natural numbers + +The [poset](order-theory.posets.md) of natural numbers ordered by divisibility is defined in [`elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility`](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). + +```agda +refl-div-ℕ : is-reflexive div-ℕ +pr1 (refl-div-ℕ x) = 1 +pr2 (refl-div-ℕ x) = left-unit-law-mul-ℕ x + +div-eq-ℕ : (x y : ℕ) → x = y → div-ℕ x y +div-eq-ℕ x .x refl = refl-div-ℕ x + +antisymmetric-div-ℕ : is-antisymmetric div-ℕ +antisymmetric-div-ℕ zero-ℕ zero-ℕ H K = refl +antisymmetric-div-ℕ zero-ℕ (succ-ℕ y) (pair k p) K = + inv (right-zero-law-mul-ℕ k) ∙ p +antisymmetric-div-ℕ (succ-ℕ x) zero-ℕ H (pair l q) = + inv q ∙ right-zero-law-mul-ℕ l +antisymmetric-div-ℕ (succ-ℕ x) (succ-ℕ y) (pair k p) (pair l q) = + ( inv (left-unit-law-mul-ℕ (succ-ℕ x))) ∙ + ( ( ap + ( _*ℕ (succ-ℕ x)) + ( inv + ( is-one-right-is-one-mul-ℕ l k + ( is-one-is-left-unit-mul-ℕ (l *ℕ k) x + ( ( associative-mul-ℕ l k (succ-ℕ x)) ∙ + ( ap (l *ℕ_) p ∙ q)))))) ∙ + ( p)) + +transitive-div-ℕ : is-transitive div-ℕ +pr1 (transitive-div-ℕ x y z (pair l q) (pair k p)) = l *ℕ k +pr2 (transitive-div-ℕ x y z (pair l q) (pair k p)) = + associative-mul-ℕ l k x ∙ (ap (l *ℕ_) p ∙ q) +``` + ### `1` divides any number ```agda diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index 33989ff00a..93514b96e5 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -169,7 +169,7 @@ is-prop-is-nontrivial-divisor-ℕ : is-prop-is-nontrivial-divisor-ℕ n x = is-prop-Σ ( is-prop-le-ℕ 1 x) - ( λ p → is-prop-div-ℕ x n (is-nonzero-le-ℕ 1 x p)) + ( λ p → is-prop-div-ℕ x n (inl (is-nonzero-le-ℕ 1 x p))) is-nontrivial-divisor-ℕ-Prop : (n x : ℕ) → Prop lzero pr1 (is-nontrivial-divisor-ℕ-Prop n x) = is-nontrivial-divisor-ℕ n x diff --git a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md index fe5615c83e..e80af86903 100644 --- a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md +++ b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md @@ -1038,26 +1038,3 @@ equiv-neg-Fin : pr1 (equiv-neg-Fin k) = neg-Fin k pr2 (equiv-neg-Fin k) = is-equiv-neg-Fin k ``` - -## Properties - -### Divisibility is a decidable relation on `ℕ` - -```agda -is-decidable-div-ℕ : (d x : ℕ) → is-decidable (div-ℕ d x) -is-decidable-div-ℕ zero-ℕ x = - is-decidable-iff - ( div-eq-ℕ zero-ℕ x) - ( inv ∘ (is-zero-div-zero-ℕ x)) - ( is-decidable-is-zero-ℕ' x) -is-decidable-div-ℕ (succ-ℕ d) x = - is-decidable-iff - ( div-is-zero-mod-succ-ℕ d x) - ( is-zero-mod-succ-ℕ d x) - ( is-decidable-is-zero-Fin (mod-succ-ℕ d x)) - -div-ℕ-Decidable-Prop : (d x : ℕ) → is-nonzero-ℕ d → Decidable-Prop lzero -pr1 (div-ℕ-Decidable-Prop d x H) = div-ℕ d x -pr1 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-prop-div-ℕ d x H -pr2 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-decidable-div-ℕ d x -``` diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index f8009e2b92..eb9869dede 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -172,15 +172,26 @@ is-odd-add-two-is-odd-ℤ a H K = H (is-even-is-even-add-two-ℤ a K) ```agda is-odd-succ-is-even-ℤ : (a : ℤ) → is-even-ℤ a → is-odd-ℤ (a +ℤ one-ℤ) -is-odd-succ-is-even-ℤ a H K = {!!} +is-odd-succ-is-even-ℤ a H K = + is-odd-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a one-ℤ H K) ``` -### If an integer `x` is odd, then `x + 1` is even +### If an integer `x + 1` is even, then `x` is odd ```agda +is-odd-is-even-succ-ℤ : + (a : ℤ) → is-even-ℤ (a +ℤ one-ℤ) → is-odd-ℤ a +is-odd-is-even-succ-ℤ a H K = + is-odd-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a one-ℤ K H) +``` + +### If an integer `x` is odd, then `x + 1` is even + +```text is-even-succ-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → is-even-ℤ (a +ℤ one-ℤ) -is-even-succ-is-odd-ℤ a H = {!a!} +is-even-succ-is-odd-ℤ a H with is-decidable-div-ℤ (int-ℕ 2) (a +ℤ one-ℤ) +... | d = {!!} ``` ### If an integer `x` is even, then `x + 1` is odd diff --git a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md index 0f35339378..4deb6c75a3 100644 --- a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md @@ -16,6 +16,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.cartesian-product-types +open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types @@ -64,7 +65,9 @@ is-prop-is-proper-divisor-ℕ : (n d : ℕ) → is-prop (is-proper-divisor-ℕ n is-prop-is-proper-divisor-ℕ n zero-ℕ (pair f g) = ex-falso (f (inv (is-zero-div-zero-ℕ n g))) is-prop-is-proper-divisor-ℕ n (succ-ℕ d) = - is-prop-product is-prop-neg (is-prop-div-ℕ (succ-ℕ d) n (is-nonzero-succ-ℕ d)) + is-prop-product + ( is-prop-neg) + ( is-prop-div-ℕ (succ-ℕ d) n (inl (is-nonzero-succ-ℕ d))) ``` ### If a natural number has a proper divisor, then `1` is a proper divisor From 9fa26d3366316f5c481cb53f90ccbdf2b6d05d4b Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Thu, 31 Oct 2024 17:33:40 -0400 Subject: [PATCH 19/72] refactoring decidability of divisibility of integers --- .../collatz-conjecture.lagda.md | 2 +- .../dirichlet-convolution.lagda.md | 2 +- .../divisibility-integers.lagda.md | 33 ++++++++++++++++ .../divisibility-natural-numbers.lagda.md | 38 +++++++++++++------ .../modular-arithmetic.lagda.md | 14 ------- 5 files changed, 62 insertions(+), 27 deletions(-) diff --git a/src/elementary-number-theory/collatz-conjecture.lagda.md b/src/elementary-number-theory/collatz-conjecture.lagda.md index 9c44a48d71..2cc9fef51c 100644 --- a/src/elementary-number-theory/collatz-conjecture.lagda.md +++ b/src/elementary-number-theory/collatz-conjecture.lagda.md @@ -7,7 +7,7 @@ module elementary-number-theory.collatz-conjecture where
Imports ```agda -open import elementary-number-theory.modular-arithmetic-standard-finite-types +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers diff --git a/src/elementary-number-theory/dirichlet-convolution.lagda.md b/src/elementary-number-theory/dirichlet-convolution.lagda.md index 0745ebdfe9..6216a71cdb 100644 --- a/src/elementary-number-theory/dirichlet-convolution.lagda.md +++ b/src/elementary-number-theory/dirichlet-convolution.lagda.md @@ -9,7 +9,7 @@ module elementary-number-theory.dirichlet-convolution where ```agda open import elementary-number-theory.arithmetic-functions open import elementary-number-theory.bounded-sums-arithmetic-functions -open import elementary-number-theory.modular-arithmetic-standard-finite-types +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.nonzero-natural-numbers diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 2afc5710b3..027e3b9ac0 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -404,3 +404,36 @@ div-div-int-abs-divisor-ℤ : div-div-int-abs-divisor-ℤ {x} {y} = div-sim-unit-ℤ (sim-unit-abs-ℤ x) (refl-sim-unit-ℤ y) ``` + +### An integer `a` divides an integer `b` if and only if the natural number `|a|` divides the natural number `|b|` + +```agda +div-abs-div-ℤ : + (a b : ℤ) → div-ℤ a b → div-ℕ (abs-ℤ a) (abs-ℤ b) +div-abs-div-ℤ a b H = + div-div-int-ℕ + ( div-sim-unit-ℤ + ( symmetric-sim-unit-ℤ _ _ (sim-unit-abs-ℤ a)) + ( symmetric-sim-unit-ℤ _ _ (sim-unit-abs-ℤ b)) + ( H)) + +div-div-abs-ℤ : + (a b : ℤ) → div-ℕ (abs-ℤ a) (abs-ℤ b) → div-ℤ a b +div-div-abs-ℤ a b H = + div-sim-unit-ℤ + ( sim-unit-abs-ℤ a) + ( sim-unit-abs-ℤ b) + ( div-int-div-ℕ H) +``` + +### Divisibility on the integers is decidable + +```agda +is-decidable-div-ℤ : + (a b : ℤ) → is-decidable (div-ℤ a b) +is-decidable-div-ℤ a b = + is-decidable-iff + ( div-div-abs-ℤ a b) + ( div-abs-div-ℤ a b) + ( is-decidable-div-ℕ _ _) +``` diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 153d9cc343..920601ec54 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -44,9 +44,12 @@ open import foundation.universe-levels ## Idea -On this page we define the {{#concept "divisibility" Disambiguation="natural numbers" Agda=div-ℕ WD="divisibility" WDID=Q5284415}} relation on the [natural numbers](elementary-number-theory.natural-numbers.md). +On this page we define the +{{#concept "divisibility" Disambiguation="natural numbers" Agda=div-ℕ WD="divisibility" WDID=Q5284415}} +relation on the [natural numbers](elementary-number-theory.natural-numbers.md). A natural number `m` is said to **divide** a natural number `n` if there exists -a natural number `k` equipped with an [identification](foundation-core.identifications.md) `km = n`. Using the +a natural number `k` equipped with an +[identification](foundation-core.identifications.md) `km = n`. Using the [Curry–Howard interpretation](https://en.wikipedia.org/wiki/Curry–Howard_correspondence) of logic into type theory, we express divisibility as follows: @@ -54,19 +57,30 @@ of logic into type theory, we express divisibility as follows: div-ℕ m n := Σ (k : ℕ), k *ℕ m = n. ``` -If `n` is a [nonzero natural number](elementary-number-theory.nonzero-natural-numbers.md), then the type `div-ℕ m n` is always a [proposition](foundation-core.propositions.md) in -the sense that the type `div-ℕ m n` contains at most one element. +If `n` is a +[nonzero natural number](elementary-number-theory.nonzero-natural-numbers.md), +then the type `div-ℕ m n` is always a +[proposition](foundation-core.propositions.md) in the sense that the type +`div-ℕ m n` contains at most one element. -The divisibility relation is [logically equivalent](foundation.logical-equivalences.md), though not [equivalent](foundation-core.equivalences.md), to the -{{#concept "bounded divisibility" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} relation, which is defined by +The divisibility relation is +[logically equivalent](foundation.logical-equivalences.md), though not +[equivalent](foundation-core.equivalences.md), to the +{{#concept "bounded divisibility" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} +relation, which is defined by ```text bounded-div-ℕ m n := Σ (k : ℕ), (k ≤ n) × (k *ℕ m = n). ``` -The discrepancy between divisibility and bounded divisibility is manifested at `(0 , 0)`. Note that `div-ℕ 0 0 ≃ ℕ`, while `bounded-div-ℕ 0 0` is [contractible](foundation-core.contractible-types.md). For all other values we have `div-ℕ m n ≃ bounded-div-ℕ m n`. +The discrepancy between divisibility and bounded divisibility is manifested at +`(0 , 0)`. Note that `div-ℕ 0 0 ≃ ℕ`, while `bounded-div-ℕ 0 0` is +[contractible](foundation-core.contractible-types.md). For all other values we +have `div-ℕ m n ≃ bounded-div-ℕ m n`. -The notion of bounded divisibility has several advantages over ordinary divisibility: Bounded divisibility is always a proposition, and the proof that divisibility is decidable uses the proof that bounded divisibility is decidable. +The notion of bounded divisibility has several advantages over ordinary +divisibility: Bounded divisibility is always a proposition, and the proof that +divisibility is decidable uses the proof that bounded divisibility is decidable. ## Definitions @@ -76,7 +90,7 @@ The notion of bounded divisibility has several advantages over ordinary divisibi module _ (m n : ℕ) where - + div-ℕ : UU lzero div-ℕ = Σ ℕ (λ k → k *ℕ m = n) @@ -97,7 +111,7 @@ module _ module _ (m n : ℕ) where - + bounded-div-ℕ : UU lzero bounded-div-ℕ = Σ ℕ (λ k → (k ≤-ℕ n) × (k *ℕ m = n)) @@ -297,7 +311,9 @@ pr2 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-decidable-bounded-div-ℕ k ### The divisibility relation is a partial order on the natural numbers -The [poset](order-theory.posets.md) of natural numbers ordered by divisibility is defined in [`elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility`](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). +The [poset](order-theory.posets.md) of natural numbers ordered by divisibility +is defined in +[`elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility`](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). ```agda refl-div-ℕ : is-reflexive div-ℕ diff --git a/src/elementary-number-theory/modular-arithmetic.lagda.md b/src/elementary-number-theory/modular-arithmetic.lagda.md index f2c68101c3..0e699d2ee4 100644 --- a/src/elementary-number-theory/modular-arithmetic.lagda.md +++ b/src/elementary-number-theory/modular-arithmetic.lagda.md @@ -783,20 +783,6 @@ has-no-fixed-points-succ-Fin {succ-ℕ k} x = has-no-fixed-points-succ-ℤ-Mod (succ-ℕ k) x ``` -### Divisibility is decidable - -```agda -is-decidable-div-ℤ : (d x : ℤ) → is-decidable (div-ℤ d x) -is-decidable-div-ℤ d x = - is-decidable-iff - ( div-div-int-abs-divisor-ℤ ∘ div-is-zero-mod-ℤ (abs-ℤ d) x) - ( is-zero-mod-div-ℤ (abs-ℤ d) x ∘ div-int-abs-divisor-div-ℤ) - ( has-decidable-equality-ℤ-Mod - ( abs-ℤ d) - ( mod-ℤ (abs-ℤ d) x) - ( zero-ℤ-Mod (abs-ℤ d))) -``` - ### `mod-ℤ` is surjective ```agda From 19a1948d87022f31a92f573c1f8f39f98a668271 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 5 Nov 2024 11:39:53 -0500 Subject: [PATCH 20/72] some work --- .../divisibility-integers.lagda.md | 10 +- .../parity-integers.lagda.md | 133 +++++++++++++----- .../parity-natural-numbers.lagda.md | 103 ++++++++++++-- .../powers-of-two.lagda.md | 29 ++-- src/literature/oeis.lagda.md | 110 ++++++++++++--- 5 files changed, 297 insertions(+), 88 deletions(-) diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 027e3b9ac0..231f62e274 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -367,19 +367,19 @@ div-presim-unit-ℤ : {x y x' y' : ℤ} → presim-unit-ℤ x x' → presim-unit-ℤ y y' → div-ℤ x y → div-ℤ x' y' pr1 (div-presim-unit-ℤ {x} {y} {x'} {y'} (pair u q) (pair v r) (pair d p)) = - ((int-unit-ℤ v) *ℤ d) *ℤ (int-unit-ℤ u) + (int-unit-ℤ v *ℤ d) *ℤ (int-unit-ℤ u) pr2 (div-presim-unit-ℤ {x} {y} {x'} {y'} (pair u q) (pair v r) (pair d p)) = ( ap ((((int-unit-ℤ v) *ℤ d) *ℤ (int-unit-ℤ u)) *ℤ_) (inv q)) ∙ ( ( associative-mul-ℤ - ( (int-unit-ℤ v) *ℤ d) + ( int-unit-ℤ v *ℤ d) ( int-unit-ℤ u) - ( (int-unit-ℤ u) *ℤ x)) ∙ + ( int-unit-ℤ u *ℤ x)) ∙ ( ( ap - ( ((int-unit-ℤ v) *ℤ d) *ℤ_) + ( (int-unit-ℤ v *ℤ d) *ℤ_) ( ( inv (associative-mul-ℤ (int-unit-ℤ u) (int-unit-ℤ u) x)) ∙ ( ap (_*ℤ x) (idempotent-is-unit-ℤ (is-unit-unit-ℤ u))))) ∙ ( ( associative-mul-ℤ (int-unit-ℤ v) d x) ∙ - ( ( ap ((int-unit-ℤ v) *ℤ_) p) ∙ + ( ( ap (int-unit-ℤ v *ℤ_) p) ∙ ( r))))) div-sim-unit-ℤ : diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index eb9869dede..8834b09c43 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -9,14 +9,20 @@ module elementary-number-theory.parity-integers where ```agda open import elementary-number-theory.absolute-value-integers open import elementary-number-theory.addition-integers +open import elementary-number-theory.difference-integers open import elementary-number-theory.divisibility-integers open import elementary-number-theory.integers +open import elementary-number-theory.multiplication-integers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.unit-similarity-integers open import foundation.coproduct-types open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.fibers-of-maps +open import foundation.identity-types open import foundation.negation open import foundation.universe-levels ``` @@ -38,16 +44,41 @@ integers are those that aren't. ## Definitions -### Even and odd integers +### Even integers ```agda is-even-ℤ : ℤ → UU lzero is-even-ℤ a = div-ℤ (int-ℕ 2) a +``` + +### The bi-infinite sequence of even integers + +```agda +even-integer-ℤ : ℤ → ℤ +even-integer-ℤ a = int-ℕ 2 *ℤ a +``` + +### Odd integers +```agda is-odd-ℤ : ℤ → UU lzero is-odd-ℤ a = ¬ (is-even-ℤ a) ``` +### The bi-infinite sequence of odd integers + +```agda +odd-integer-ℤ : ℤ → ℤ +odd-integer-ℤ a = int-ℕ 2 *ℤ a +ℤ one-ℤ +``` + +### The type of odd expansions of an integer + +```agda +has-odd-expansion-ℤ : ℤ → UU lzero +has-odd-expansion-ℤ = fiber odd-integer-ℤ +``` + ## Properties ### A natural number is even if and only if it is even as an integer @@ -129,20 +160,37 @@ is-decidable-is-odd-ℤ : (a : ℤ) → is-decidable (is-odd-ℤ a) is-decidable-is-odd-ℤ a = is-decidable-neg (is-decidable-is-even-ℤ a) ``` +### An integer is even if and only if it is not odd + +```agda +is-even-is-not-odd-ℤ : + (a : ℤ) → ¬ (is-odd-ℤ a) → is-even-ℤ a +is-even-is-not-odd-ℤ a = + double-negation-elim-is-decidable (is-decidable-is-even-ℤ a) +``` + ### `0` is an even integer ```agda -is-even-zero-ℤ : is-even-ℤ (int-ℕ 0) +is-even-zero-ℤ : is-even-ℤ zero-ℤ is-even-zero-ℤ = is-even-int-is-even-ℕ 0 is-even-zero-ℕ ``` ### `1` is an odd integer ```agda -is-odd-one-ℤ : is-odd-ℤ (int-ℕ 1) +is-odd-one-ℤ : is-odd-ℤ one-ℤ is-odd-one-ℤ = is-odd-int-is-odd-ℕ 1 is-odd-one-ℕ ``` +### `-1` is an odd integer + +```agda +is-odd-neg-one-ℤ : is-odd-ℤ neg-one-ℤ +is-odd-neg-one-ℤ H = + is-odd-one-ℤ (div-left-summand-ℤ (int-ℕ 2) one-ℤ neg-one-ℤ H is-even-zero-ℤ) +``` + ### An integer `x` is even if and only if `x + 2` is even ```agda @@ -176,6 +224,15 @@ is-odd-succ-is-even-ℤ a H K = is-odd-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a one-ℤ H K) ``` +### If an integer `x` is even, then `x - 1` is odd + +```agda +is-odd-pred-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → is-odd-ℤ (a -ℤ one-ℤ) +is-odd-pred-is-even-ℤ a H K = + is-odd-neg-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a neg-one-ℤ H K) +``` + ### If an integer `x + 1` is even, then `x` is odd ```agda @@ -185,26 +242,22 @@ is-odd-is-even-succ-ℤ a H K = is-odd-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a one-ℤ K H) ``` -### If an integer `x` is odd, then `x + 1` is even +### If an integer `x - 1` is even, then `x` is odd -```text -is-even-succ-is-odd-ℤ : - (a : ℤ) → is-odd-ℤ a → is-even-ℤ (a +ℤ one-ℤ) -is-even-succ-is-odd-ℤ a H with is-decidable-div-ℤ (int-ℕ 2) (a +ℤ one-ℤ) -... | d = {!!} +```agda +is-odd-is-even-pred-ℤ : + (a : ℤ) → is-even-ℤ (a -ℤ one-ℤ) → is-odd-ℤ a +is-odd-is-even-pred-ℤ a H K = + is-odd-neg-one-ℤ (div-right-summand-ℤ (int-ℕ 2) a neg-one-ℤ K H) ``` -### If an integer `x` is even, then `x + 1` is odd +### If an integer `x` is odd, then `x + 1` is even -```text -is-odd-succ-is-even-ℕ : - (n : ℕ) → is-even-ℕ n → is-odd-ℕ (succ-ℕ n) -is-odd-succ-is-even-ℕ zero-ℕ p = is-odd-one-ℕ -is-odd-succ-is-even-ℕ (succ-ℕ zero-ℕ) p = ex-falso (is-odd-one-ℕ p) -is-odd-succ-is-even-ℕ (succ-ℕ (succ-ℕ n)) p = - is-odd-succ-succ-is-odd-ℕ - ( succ-ℕ n) - ( is-odd-succ-is-even-ℕ n (is-even-is-even-succ-succ-ℕ n p)) +```agda +is-even-succ-is-odd-ℤ : + (a : ℤ) → is-odd-ℤ a → is-even-ℤ (a +ℤ one-ℤ) +is-even-succ-is-odd-ℤ a H = + is-even-is-not-odd-ℤ (a +ℤ one-ℤ) ( λ K → {!!}) ``` ### If an integer `x + 1` is odd, then `x` is even @@ -218,34 +271,42 @@ is-even-is-odd-succ-ℕ n p = ( is-even-succ-is-odd-ℕ (succ-ℕ n) p) ``` -### If an integer `x + 1` is even, then `x` is odd +### An integer is odd if and only if it has an odd expansion -```text -is-odd-is-even-succ-ℕ : - (n : ℕ) → is-even-ℕ (succ-ℕ n) → is-odd-ℕ n -is-odd-is-even-succ-ℕ n p = - is-odd-is-odd-succ-succ-ℕ - ( n) - ( is-odd-succ-is-even-ℕ (succ-ℕ n) p) +```agda +is-odd-has-odd-expansion-ℤ : (a : ℤ) → has-odd-expansion-ℤ a → is-odd-ℤ a +is-odd-has-odd-expansion-ℤ ._ (x , refl) = + is-odd-succ-is-even-ℤ (int-ℕ 2 *ℤ x) (x , commutative-mul-ℤ x (int-ℕ 2)) + +has-odd-expansion-neg-one-ℤ : has-odd-expansion-ℤ neg-one-ℤ +pr1 has-odd-expansion-neg-one-ℤ = neg-one-ℤ +pr2 has-odd-expansion-neg-one-ℤ = refl + +has-odd-expansion-is-odd-ℤ : + (a : ℤ) → is-odd-ℤ a → has-odd-expansion-ℤ a +has-odd-expansion-is-odd-ℤ (inl zero-ℕ) H = + has-odd-expansion-neg-one-ℤ +has-odd-expansion-is-odd-ℤ (inl (succ-ℕ zero-ℕ)) H = + ex-falso (H (neg-one-ℤ , refl)) +has-odd-expansion-is-odd-ℤ (inl (succ-ℕ (succ-ℕ a))) H = + {!!} + where + t = has-odd-expansion-is-odd-ℤ (inl a) ? +has-odd-expansion-is-odd-ℤ (inr a) H = {!!} ``` -### An integer `x` is odd if and only if there is an integer `y` such that `succ-ℕ (y *ℕ 2) = x` - ```text -has-odd-expansion : ℕ → UU lzero -has-odd-expansion x = Σ ℕ (λ y → (succ-ℕ (y *ℕ 2)) = x) - -is-odd-has-odd-expansion : (n : ℕ) → has-odd-expansion n → is-odd-ℕ n -is-odd-has-odd-expansion .(succ-ℕ (m *ℕ 2)) (m , refl) = +is-odd-has-odd-expansion-ℕ : (n : ℕ) → has-odd-expansion-ℕ n → is-odd-ℕ n +is-odd-has-odd-expansion-ℕ .(succ-ℕ (m *ℕ 2)) (m , refl) = is-odd-succ-is-even-ℕ (m *ℕ 2) (m , refl) -has-odd-expansion-is-odd : (n : ℕ) → is-odd-ℕ n → has-odd-expansion n +has-odd-expansion-is-odd : (n : ℕ) → is-odd-ℕ n → has-odd-expansion-ℕ n has-odd-expansion-is-odd zero-ℕ p = ex-falso (p is-even-zero-ℕ) has-odd-expansion-is-odd (succ-ℕ zero-ℕ) p = 0 , refl has-odd-expansion-is-odd (succ-ℕ (succ-ℕ n)) p = ( succ-ℕ (pr1 s)) , ap (succ-ℕ ∘ succ-ℕ) (pr2 s) where - s : has-odd-expansion n + s : has-odd-expansion-ℕ n s = has-odd-expansion-is-odd n (is-odd-is-odd-succ-succ-ℕ n p) ``` diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 7adb25d863..52046ea0de 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -10,6 +10,7 @@ module elementary-number-theory.parity-natural-numbers where open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.euclidean-division-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -17,7 +18,9 @@ open import elementary-number-theory.natural-numbers open import foundation.action-on-identifications-functions open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.double-negation open import foundation.empty-types +open import foundation.fibers-of-maps open import foundation.function-types open import foundation.identity-types open import foundation.negation @@ -41,14 +44,41 @@ and odd natural numbers are those that aren't. ## Definition +### Even natural numbers + ```agda is-even-ℕ : ℕ → UU lzero is-even-ℕ n = div-ℕ 2 n +``` + +### The sequence of even numbers + +```agda +even-number-ℕ : ℕ → ℕ +even-number-ℕ n = 2 *ℕ n +``` + +### Odd natural numbers +```agda is-odd-ℕ : ℕ → UU lzero is-odd-ℕ n = ¬ (is-even-ℕ n) ``` +### The sequence of odd numbers + +```agda +odd-number-ℕ : ℕ → ℕ +odd-number-ℕ n = succ-ℕ (2 *ℕ n) +``` + +### The predicate of having an odd expansion + +```agda +has-odd-expansion-ℕ : ℕ → UU lzero +has-odd-expansion-ℕ n = fiber odd-number-ℕ n +``` + ## Properties ### Being even or odd is decidable @@ -61,6 +91,17 @@ is-decidable-is-odd-ℕ : (x : ℕ) → is-decidable (is-odd-ℕ x) is-decidable-is-odd-ℕ x = is-decidable-neg (is-decidable-is-even-ℕ x) ``` +### A number is even if and only if it is not odd + +```agda +is-even-is-not-odd-ℕ : (x : ℕ) → ¬ (is-odd-ℕ x) → is-even-ℕ x +is-even-is-not-odd-ℕ x = + double-negation-elim-is-decidable (is-decidable-is-even-ℕ x) + +is-not-odd-is-even-ℕ : (x : ℕ) → is-even-ℕ x → ¬ (is-odd-ℕ x) +is-not-odd-is-even-ℕ x = intro-double-negation +``` + ### `0` is an even natural number ```agda @@ -150,24 +191,22 @@ is-odd-is-even-succ-ℕ n p = ( is-odd-succ-is-even-ℕ (succ-ℕ n) p) ``` -### A natural number `x` is odd if and only if there is a natural number `y` such that `succ-ℕ (y *ℕ 2) = x` +### A natural number is odd if and only if it has an odd expansion ```agda -has-odd-expansion : ℕ → UU lzero -has-odd-expansion x = Σ ℕ (λ y → (succ-ℕ (y *ℕ 2)) = x) - -is-odd-has-odd-expansion : (n : ℕ) → has-odd-expansion n → is-odd-ℕ n -is-odd-has-odd-expansion .(succ-ℕ (m *ℕ 2)) (m , refl) = - is-odd-succ-is-even-ℕ (m *ℕ 2) (m , refl) - -has-odd-expansion-is-odd : (n : ℕ) → is-odd-ℕ n → has-odd-expansion n -has-odd-expansion-is-odd zero-ℕ p = ex-falso (p is-even-zero-ℕ) -has-odd-expansion-is-odd (succ-ℕ zero-ℕ) p = 0 , refl -has-odd-expansion-is-odd (succ-ℕ (succ-ℕ n)) p = - ( succ-ℕ (pr1 s)) , ap (succ-ℕ ∘ succ-ℕ) (pr2 s) +is-odd-has-odd-expansion-ℕ : (n : ℕ) → has-odd-expansion-ℕ n → is-odd-ℕ n +is-odd-has-odd-expansion-ℕ .(succ-ℕ (2 *ℕ m)) (m , refl) = + is-odd-succ-is-even-ℕ (2 *ℕ m) (m , commutative-mul-ℕ m 2) + +has-odd-expansion-is-odd-ℕ : (n : ℕ) → is-odd-ℕ n → has-odd-expansion-ℕ n +has-odd-expansion-is-odd-ℕ zero-ℕ p = ex-falso (p is-even-zero-ℕ) +has-odd-expansion-is-odd-ℕ (succ-ℕ zero-ℕ) p = (0 , refl) +has-odd-expansion-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = + ( succ-ℕ (pr1 s)) , + ap (succ-ℕ ∘ succ-ℕ) (left-successor-law-add-ℕ _ _ ∙ pr2 s) where - s : has-odd-expansion n - s = has-odd-expansion-is-odd n (is-odd-is-odd-succ-succ-ℕ n p) + s : has-odd-expansion-ℕ n + s = has-odd-expansion-is-odd-ℕ n (is-odd-is-odd-succ-succ-ℕ n p) ``` ### A number is even if and only if it is divisible by an even number @@ -198,3 +237,37 @@ is-even-right-summand-ℕ : (m n : ℕ) → is-even-ℕ m → is-even-ℕ (add-ℕ m n) → is-even-ℕ n is-even-right-summand-ℕ = div-right-summand-ℕ 2 ``` + +### If any two out of `x`, `y`, and `x + y` are odd, then the third number is even + +```agda +is-even-add-is-odd-ℕ : + (m n : ℕ) → is-odd-ℕ m → is-odd-ℕ n → is-even-ℕ (m +ℕ n) +is-even-add-is-odd-ℕ m zero-ℕ H K = ex-falso (K is-even-zero-ℕ) +is-even-add-is-odd-ℕ m (succ-ℕ zero-ℕ) H K = + is-even-succ-is-odd-ℕ m H +is-even-add-is-odd-ℕ m (succ-ℕ (succ-ℕ n)) H K = + is-even-succ-succ-is-even-ℕ + ( add-ℕ m n) + ( is-even-add-is-odd-ℕ m n H (is-odd-is-odd-succ-succ-ℕ n K)) + +is-even-left-summand-is-odd-ℕ : + (m n : ℕ) → is-odd-ℕ n → is-odd-ℕ (m +ℕ n) → is-even-ℕ m +is-even-left-summand-is-odd-ℕ m zero-ℕ H K = ex-falso (H is-even-zero-ℕ) +is-even-left-summand-is-odd-ℕ m (succ-ℕ zero-ℕ) H K = + is-even-is-odd-succ-ℕ m K +is-even-left-summand-is-odd-ℕ m (succ-ℕ (succ-ℕ n)) H K = + is-even-left-summand-is-odd-ℕ m n + ( is-odd-is-odd-succ-succ-ℕ n H) + ( is-odd-is-odd-succ-succ-ℕ (m +ℕ n) K) + +is-even-right-summand-is-odd-ℕ : + (m n : ℕ) → is-odd-ℕ m → is-odd-ℕ (m +ℕ n) → is-even-ℕ n +is-even-right-summand-is-odd-ℕ m zero-ℕ H K = is-even-zero-ℕ +is-even-right-summand-is-odd-ℕ m (succ-ℕ zero-ℕ) H K = + ex-falso (K (is-even-succ-is-odd-ℕ m H)) +is-even-right-summand-is-odd-ℕ m (succ-ℕ (succ-ℕ n)) H K = + is-even-succ-succ-is-even-ℕ n + ( is-even-right-summand-is-odd-ℕ m n H + ( is-odd-is-odd-succ-succ-ℕ (m +ℕ n) K)) +``` diff --git a/src/elementary-number-theory/powers-of-two.lagda.md b/src/elementary-number-theory/powers-of-two.lagda.md index d856d60206..1fa5eaf150 100644 --- a/src/elementary-number-theory/powers-of-two.lagda.md +++ b/src/elementary-number-theory/powers-of-two.lagda.md @@ -65,12 +65,14 @@ abstract ( λ k f → ( λ where ( inl x) → - ( let s = has-odd-expansion-is-odd k (is-odd-is-even-succ-ℕ k x) + ( let s = has-odd-expansion-is-odd-ℕ k (is-odd-is-even-succ-ℕ k x) in - pair - ( 0 , (succ-ℕ (pr1 s))) - ( ( ap ((succ-ℕ ∘ succ-ℕ) ∘ succ-ℕ) (left-unit-law-add-ℕ _)) ∙ - ( ( ap (succ-ℕ ∘ succ-ℕ) (pr2 s))))) + ( ( 0 , succ-ℕ (pr1 s)) , + ( ap + ( succ-ℕ ∘ succ-ℕ) + ( left-unit-law-add-ℕ _ ∙ + ap succ-ℕ (commutative-mul-ℕ (pr1 s) 2) ∙ + pr2 s)))) ( inr x) → ( let e : is-even-ℕ k e = is-even-is-odd-succ-ℕ k x @@ -122,16 +124,25 @@ is-pair-expansion-unique zero-ℕ (succ-ℕ u') v v' p = ex-falso (s t) where s : is-odd-ℕ (succ-ℕ (0 +ℕ (v *ℕ 2))) - s = is-odd-has-odd-expansion _ - ( v , ap succ-ℕ (inv (left-unit-law-add-ℕ _))) + s = + is-odd-has-odd-expansion-ℕ _ + ( v , + ap + ( succ-ℕ) + ( commutative-mul-ℕ 2 v ∙ + inv (left-unit-law-add-ℕ _))) t : is-even-ℕ (succ-ℕ (0 +ℕ (v *ℕ 2))) t = tr is-even-ℕ (inv p) (div-mul-ℕ' _ 2 _ ((exp-ℕ 2 u') , refl)) is-pair-expansion-unique (succ-ℕ u) zero-ℕ v v' p = ex-falso (s t) where s : is-odd-ℕ (succ-ℕ (0 +ℕ (v' *ℕ 2))) - s = is-odd-has-odd-expansion _ - ( v' , ap succ-ℕ (inv (left-unit-law-add-ℕ _))) + s = + is-odd-has-odd-expansion-ℕ _ + ( v' , + ap + ( succ-ℕ) + ( commutative-mul-ℕ 2 v' ∙ inv (left-unit-law-add-ℕ _))) t : is-even-ℕ (succ-ℕ (0 +ℕ (v' *ℕ 2))) t = tr is-even-ℕ p (div-mul-ℕ' _ 2 _ ((exp-ℕ 2 u) , refl)) is-pair-expansion-unique (succ-ℕ u) (succ-ℕ u') v v' p = diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index 31a52f6beb..ba0c1f6f3f 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -20,28 +20,34 @@ open import foundation.function-types ## Sequences -### [A000001](https://oeis.org/A000001) The number of groups of order `n` +### A000001 -- The number of groups of order `n` ```agda open import finite-group-theory.finite-groups using ( number-of-groups-of-order) ``` -### [A000002](https://oeis.org/A000002) The Kolakoski sequence +OEIS: [A000001](https://oeis.org/A000001) + +### A000002 -- The Kolakoski sequence ```agda open import elementary-number-theory.kolakoski-sequence using ( kolakoski) ``` -### [A000004](https://oeis.org/A000004) The zero sequence +OEIS: [A000002](https://oeis.org/A000002) + +### A000004 -- The zero sequence ```agda A000004 : ℕ → ℕ A000004 _ = zero-ℕ ``` -### [A000007](https://oeis.org/A000007) The characteristic function for 0 +OEIS: [A000004](https://oeis.org/A000004) + +### A000007 -- The characteristic function for 0 ```agda A000007 : ℕ → ℕ @@ -49,139 +55,197 @@ A000007 zero-ℕ = 1 A000007 (succ-ℕ _) = 0 ``` -### [A000010](https://oeis.org/A000010) Euler's totient function +OEIS: [A000007](https://oeis.org/A000007) + +### A000010 -- Euler's totient function ```agda open import elementary-number-theory.eulers-totient-function using ( eulers-totient-function-relatively-prime) ``` -### [A000012](https://oeis.org/A000012) All 1's sequence +OEIS: [A000010](https://oeis.org/A000010) + +### A000012 -- All 1's sequence ```agda A000012 : ℕ → ℕ A000012 _ = 1 ``` -### [A000027](https://oeis.org/A000027) The positive integers +OEIS: [A000012](https://oeis.org/A000012) + +### A000027 -- The positive integers ```agda A000027 : ℕ → ℕ A000027 = succ-ℕ ``` -### [A000040](https://oeis.org/A000040) The prime numbers +OEIS: [A000027](https://oeis.org/A000027) + +### A000040 -- The prime numbers ```agda open import elementary-number-theory.infinitude-of-primes using ( prime-ℕ) ``` -### [A000045](https://oeis.org/A000045) The Fibonacci sequence +OEIS: [A000040](https://oeis.org/A000040) + +### A000045 -- The Fibonacci sequence ```agda open import elementary-number-theory.fibonacci-sequence using ( Fibonacci-ℕ) ``` -### [A000058](https://oeis.org/A000058) Sylvester's sequence +OEIS: [A000045](https://oeis.org/A000045) + +### A000058 -- Sylvester's sequence ```agda open import elementary-number-theory.sylvesters-sequence using ( sylvesters-sequence-ℕ) ``` -### [A000079](https://oeis.org/A000079) Powers of `2` +OEIS: [A000058](https://oeis.org/A000058) + +### A000079 -- Powers of `2` ```agda A000079 : ℕ → ℕ A000079 = exp-ℕ 2 ``` -### [A000108](https://oeis.org/A000108) The Catalan numbers +OEIS: [A000079](https://oeis.org/A000079) + +### A000108 -- The Catalan numbers ```agda open import elementary-number-theory.catalan-numbers using ( catalan-numbers) ``` -### [A000142](https://oeis.org/A000142) Factorials +OEIS: [A000108](https://oeis.org/A000108) + +### A000142 -- Factorials ```agda open import elementary-number-theory.factorials using ( factorial-ℕ) ``` -### [A000215](https://oeis.org/A000215) The Fermat numbers +OEIS: [A000142](https://oeis.org/A000142) + +### A000215 -- The Fermat numbers ```agda open import elementary-number-theory.fermat-numbers using ( fermat-number-ℕ) ``` -### [A000244](https://oeis.org/A000244) Powers of `3` +OEIS: [A000215](https://oeis.org/A000215) + +### A000244 -- Powers of `3` ```agda A000244 : ℕ → ℕ A000244 = exp-ℕ 3 ``` -### [A000720](https://oeis.org/A000720) The prime counting function +OEIS: [A000244](https://oeis.org/A000244) + +### A000720 -- The prime counting function ```agda open import elementary-number-theory.infinitude-of-primes using ( prime-counting-ℕ) ``` -### [A000945](https://oeis.org/A000945) The Euclid–Mullin sequence +OEIS: [A000720](https://oeis.org/A000720) + +### A000945 -- The Euclid–Mullin sequence ```agda open import elementary-number-theory.euclid-mullin-sequence using ( euclid-mullin-ℕ) ``` -### [A001175](https://oeis.org/A001175) Pisano periods +OEIS: [A000945](https://oeis.org/A000945) + +### A001175 -- Pisano periods ```agda open import elementary-number-theory.pisano-periods using ( pisano-period) ``` -### [A001177](https://oeis.org/A001177) The cofibonacci sequence +OEIS: [A001175](https://oeis.org/A001175) + +### A001177 -- The cofibonacci sequence ```agda open import elementary-number-theory.cofibonacci using ( cofibonacci) ``` -### [A001477](https://oeis.org/A001477) The natural numbers +OEIS: [A001177](https://oeis.org/A001177) + +### A001477 -- The natural numbers ```agda A001477 : ℕ → ℕ A001477 = id ``` -### [A003090](https://oeis.org/A003090) The number of main classes of Latin squares of order `n` +OEIS: [A001477](https://oeis.org/A001477) + +### A003090 -- The number of main classes of Latin squares of order `n` ```agda open import univalent-combinatorics.main-classes-of-latin-squares using ( number-of-main-classes-of-Latin-squares-of-order) ``` -### [A006369](https://oeis.org/A006369) Collatz' bijection +OEIS: [A003090](https://oeis.org/A003090) + +### A005408 -- The odd numbers + +```agda +open import elementary-number-theory.parity-natural-numbers using + ( odd-number-ℕ) +``` + +OEIS: [A005408](https://oeis.org/A005408) + +### A005843 -- The nonnegative natural numbers + +```agda +open import elementary-number-theory.parity-natural-numbers using + ( even-number-ℕ) +``` + +OEIS: [A005843](https://oeis.org/A005843) + +### A006369 -- Collatz' bijection ```agda open import elementary-number-theory.collatz-bijection using ( map-collatz-bijection) ``` -### [A046859](https://oeis.org/A046859) The main diagonal of the Ackermann–Péter function +OEIS: [A006369](https://oeis.org/A006369) + +### A046859 -- The main diagonal of the Ackermann–Péter function ```agda open import elementary-number-theory.ackermann-function using ( simplified-ackermann-ℕ) ``` +OEIS: [A046859](https://oeis.org/A046859) + ## References {{#bibliography}} From de1f5b5c065f09c4184c27b041eaf41060f3c2ca Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 5 Nov 2024 11:41:35 -0500 Subject: [PATCH 21/72] work --- src/elementary-number-theory/parity-integers.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index 8834b09c43..e36883f038 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -291,7 +291,7 @@ has-odd-expansion-is-odd-ℤ (inl (succ-ℕ zero-ℕ)) H = has-odd-expansion-is-odd-ℤ (inl (succ-ℕ (succ-ℕ a))) H = {!!} where - t = has-odd-expansion-is-odd-ℤ (inl a) ? + t = has-odd-expansion-is-odd-ℤ (inl a) {!!} has-odd-expansion-is-odd-ℤ (inr a) H = {!!} ``` From 032e45d3d2b3ef55e6412ab1aa49f261309b40e7 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 7 Dec 2024 12:09:24 -0500 Subject: [PATCH 22/72] some improvements to number theory --- ...uclidean-division-natural-numbers.lagda.md | 87 +++++++------- .../parity-integers.lagda.md | 89 ++++++++++++--- .../parity-natural-numbers.lagda.md | 107 ++++++++++++++++++ .../sums-of-natural-numbers.lagda.md | 52 +++++++-- .../triangular-numbers.lagda.md | 78 ++++++++++++- 5 files changed, 348 insertions(+), 65 deletions(-) diff --git a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md index 1ef17f8247..c0317e9e84 100644 --- a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md +++ b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md @@ -28,49 +28,19 @@ open import univalent-combinatorics.standard-finite-types ## Idea -Euclidean division is division with remainder, i.e., the Euclidean division of -`n` by `d` is the largest natural number `k ≤ n/d`, and the remainder is the -unique natural number `r < d` such that `kd + r = n`. +{{#concept "Euclidean division" Agda=euclidean-division-ℕ}} is a binary operation on the [natural numbers](elementary-number-theory.natural-numbers.md) given by division with remainder. In other words, the Euclidean division of +`n` by `d` gives the unique pair of natural numbers `q` and `r +## Idea + +The $n$th {{#concept "triangular number" Agda=triangular-number-ℕ}} is the [sum](elementary-number-theory.sums-of-natural-numbers.md) of the [natural numbers](elementary-number-theory.natural-numbers.md) up to $n$. Alternatively, the triangular numbers can be defined inductively by + +```text + T 0 := 0 + T (n + 1) := T n + (n + 1). +``` + +A classic inductive proof, which is often used as the first example when introducing students to mathematical induction, establishes the famous fact that + +```text + T n = n(n + 1)/2. +``` + ## Definition +### The sum definition of the triangular numbers + ```agda triangular-number-ℕ : ℕ → ℕ -triangular-number-ℕ 0 = 0 -triangular-number-ℕ (succ-ℕ n) = (triangular-number-ℕ n) +ℕ (succ-ℕ n) +triangular-number-ℕ n = bounded-sum-ℕ n (λ x H → x) +``` + +### The inductive definition of the triangular numbers + +```agda +inductive-triangular-number-ℕ : ℕ → ℕ +inductive-triangular-number-ℕ 0 = 0 +inductive-triangular-number-ℕ (succ-ℕ n) = + inductive-triangular-number-ℕ n +ℕ succ-ℕ n +``` + +## Properties + +### The two definitions of triangular numbers are the same + +```agda +compute-inductive-triangular-number-ℕ : + (n : ℕ) → inductive-triangular-number-ℕ n = triangular-number-ℕ n +compute-inductive-triangular-number-ℕ zero-ℕ = refl +compute-inductive-triangular-number-ℕ (succ-ℕ n) = + ap (add-ℕ' (succ-ℕ n)) (compute-inductive-triangular-number-ℕ n) +``` + +### The $n$th triangular number is $\frac{1}{2}n(n+1)$ + +**Proof.** The claim is equivalent to the claim that + +$$ + 2 \cdot \left(\sum_{i\leq n} i\right) = n(n+1). +$$ + +We prove this claim to avoid an early division by two, and the proof is by induction on $n$. In the base case both sides of the equality are $0$. For the inductive step, assume that $2 \cdot \sum_{i\leq n} i = n(n+1)$. Then we can compute + +$$ +2 \cdot \sum_{i\leq n+1} i = 2 \cdot \left(\sum_{i\leq n} i\right)+ 2(n+1) = n(n+1) + 2(n+1) = (n+1)(n+2). +$$ + +```agda +compute-triangular-number-ℕ' : + (n : ℕ) → 2 *ℕ triangular-number-ℕ n = n *ℕ succ-ℕ n +compute-triangular-number-ℕ' zero-ℕ = refl +compute-triangular-number-ℕ' (succ-ℕ n) = + left-distributive-mul-add-ℕ 2 (triangular-number-ℕ n) (succ-ℕ n) ∙ + ap-add-ℕ (compute-triangular-number-ℕ' n) (left-two-law-mul-ℕ (succ-ℕ n)) ∙ + inv (associative-add-ℕ (n *ℕ succ-ℕ n) (succ-ℕ n) (succ-ℕ n)) ∙ + inv (right-successor-law-mul-ℕ (succ-ℕ n) (succ-ℕ n)) + +compute-triangular-number-ℕ : + (n : ℕ) → + triangular-number-ℕ n = quotient-euclidean-division-ℕ 2 (n *ℕ succ-ℕ n) +compute-triangular-number-ℕ zero-ℕ = refl +compute-triangular-number-ℕ (succ-ℕ n) = {!!} ``` From c8f84adeea0c61df04deb03c47bb358b7a3d474f Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 16 Dec 2024 17:36:58 -0500 Subject: [PATCH 23/72] refactoring the fundamental theorem of arithmetic --- ...nded-divisibility-natural-numbers.lagda.md | 362 +++++++++ .../divisibility-natural-numbers.lagda.md | 329 ++++---- .../euclid-mullin-sequence.lagda.md | 2 +- ...uclidean-division-natural-numbers.lagda.md | 15 + ...fundamental-theorem-of-arithmetic.lagda.md | 756 +++++++----------- .../inequality-natural-numbers.lagda.md | 174 ++-- .../jacobi-symbol.lagda.md | 4 +- .../lists-of-prime-numbers.lagda.md | 54 ++ ...lication-lists-of-natural-numbers.lagda.md | 14 + .../multiplication-natural-numbers.lagda.md | 10 + ...ve-decompositions-natural-numbers.lagda.md | 162 ++++ ...ntrivial-divisors-natural-numbers.lagda.md | 249 ++++++ ...l-numbers-ordered-by-divisibility.lagda.md | 36 +- .../powers-of-two.lagda.md | 62 +- .../prime-numbers.lagda.md | 30 +- .../triangular-numbers.lagda.md | 2 +- src/lists/predicates-on-lists.lagda.md | 26 +- 17 files changed, 1521 insertions(+), 766 deletions(-) create mode 100644 src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/lists-of-prime-numbers.lagda.md create mode 100644 src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md diff --git a/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md new file mode 100644 index 0000000000..2663d2642c --- /dev/null +++ b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md @@ -0,0 +1,362 @@ +# Bounded divisibility of natural numbers + +```agda +module elementary-number-theory.bounded-divisibility-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.decidable-types +open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.decidable-propositions +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.equivalences +open import foundation.identity-types +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels +``` + +
+ +## Idea + +Consider two [natural numbers](elementary-number-theory.natural-numbers.md) `m` and `n`. The {{#concept "bounded divisibility relation" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} is a [binary relation](foundation.binary-relations.md) on the type of [natural numbers](elementary-number-theory.natural-numbers.md), where we say that a number `n` is bounded divisible by `m` if there is an integer `q ≤ n` such that `q * m = n`. + +The bounded divisibility relation is a slight strengthening of the [divisibility relation](elementary-number-theory.divisibility-natural-numbers.md) by ensuring that the quotient is bounded from above by `n`. This ensures that the bounded divisibility relation is valued in [propositions](foundation-core.propositions.md) for all `m` and `n`, unlike the divisibility relation. Since the bounded divisibility relation is [logically equivalent](foundation.logical-equivalences.md) to the divisibility relation, but it is always valued in the propositions, we use the bounded divisibility relation in the definition of the [poset](order-theory.posets.md) of the [natural numbers ordered by division](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). + +## Definitions + +### The bounded divisibility predicate + +```agda +module _ + (m n : ℕ) + where + + bounded-div-ℕ : UU lzero + bounded-div-ℕ = Σ ℕ (λ k → (k ≤-ℕ n) × (k *ℕ m = n)) + + quotient-bounded-div-ℕ : bounded-div-ℕ → ℕ + quotient-bounded-div-ℕ = pr1 + + upper-bound-quotient-bounded-div-ℕ : + (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H ≤-ℕ n + upper-bound-quotient-bounded-div-ℕ H = pr1 (pr2 H) + + eq-quotient-bounded-div-ℕ : + (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H *ℕ m = n + eq-quotient-bounded-div-ℕ H = pr2 (pr2 H) + + eq-quotient-bounded-div-ℕ' : + (H : bounded-div-ℕ) → m *ℕ quotient-bounded-div-ℕ H = n + eq-quotient-bounded-div-ℕ' H = + commutative-mul-ℕ m (quotient-bounded-div-ℕ H) ∙ eq-quotient-bounded-div-ℕ H +``` + +## Properties + +### If a nonzero number `n` is bounded divisible by `m`, then its quotient is nonzero + +```agda +is-nonzero-quotient-bounded-div-ℕ : + (m n : ℕ) → is-nonzero-ℕ n → (H : bounded-div-ℕ m n) → + is-nonzero-ℕ (quotient-bounded-div-ℕ m n H) +is-nonzero-quotient-bounded-div-ℕ m n N H refl = + N (inv (eq-quotient-bounded-div-ℕ m n H)) +``` + +### If a nonzero number `n` is bounded divisible by `m`, then `m` is bounded from above by `n` + +**Proof.** Suppose that `q ≤ n` is such that `q * m = n`. Since `n` is assumed to be nonzero, it follows that `q` is nonzero. Therefore it follows that `m ≤ q * m = n`. + +```agda +upper-bound-divisor-bounded-div-ℕ : + (m n : ℕ) → is-nonzero-ℕ n → bounded-div-ℕ m n → m ≤-ℕ n +upper-bound-divisor-bounded-div-ℕ m n H K = + concatenate-leq-eq-ℕ m + ( leq-mul-is-nonzero-ℕ' + ( quotient-bounded-div-ℕ m n K) + ( m) + ( is-nonzero-quotient-bounded-div-ℕ m n H K)) + ( eq-quotient-bounded-div-ℕ m n K) +``` + +### If `n` is bounded divisible by a number `m`, then `n` is bounded divisible by its quotient by `m` + +```agda +bounded-div-quotient-bounded-div-ℕ : + (m n : ℕ) (H : bounded-div-ℕ m n) → + bounded-div-ℕ (quotient-bounded-div-ℕ m n H) n +bounded-div-quotient-bounded-div-ℕ m zero-ℕ (q , b , p) = + ( 0 , refl-leq-ℕ 0 , refl) +bounded-div-quotient-bounded-div-ℕ m (succ-ℕ n) H@(q , b , p) = + ( m , + upper-bound-divisor-bounded-div-ℕ m (succ-ℕ n) (is-nonzero-succ-ℕ n) H , + commutative-mul-ℕ m (quotient-bounded-div-ℕ m (succ-ℕ n) H) ∙ + eq-quotient-bounded-div-ℕ m (succ-ℕ n) H) +``` + +### Bounded divisibility is decidable + +```agda +is-decidable-bounded-div-ℕ : + (m n : ℕ) → is-decidable (bounded-div-ℕ m n) +is-decidable-bounded-div-ℕ m n = + is-decidable-bounded-Σ-ℕ' + ( n) + ( λ x → mul-ℕ x m = n) + ( λ x → has-decidable-equality-ℕ (mul-ℕ x m) n) +``` + +### Concatenating equality and bounded divisibility + +```agda +concatenate-eq-bounded-div-ℕ : + {x y z : ℕ} → x = y → bounded-div-ℕ y z → bounded-div-ℕ x z +concatenate-eq-bounded-div-ℕ refl p = p + +concatenate-bounded-div-eq-ℕ : + {x y z : ℕ} → bounded-div-ℕ x y → y = z → bounded-div-ℕ x z +concatenate-bounded-div-eq-ℕ p refl = p + +concatenate-eq-bounded-div-eq-ℕ : + {x y z w : ℕ} → x = y → bounded-div-ℕ y z → z = w → bounded-div-ℕ x w +concatenate-eq-bounded-div-eq-ℕ refl p refl = p +``` + +### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal + +Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. + +```agda +eq-quotient-bounded-div-eq-divisor-ℕ : + (c d n : ℕ) → is-nonzero-ℕ c → c = d → + (H : bounded-div-ℕ c n) → (I : bounded-div-ℕ d n) → + quotient-bounded-div-ℕ c n H = quotient-bounded-div-ℕ d n I +eq-quotient-bounded-div-eq-divisor-ℕ c d n N p H I = + is-injective-left-mul-ℕ c N + ( tr + ( λ x → c *ℕ quotient-bounded-div-ℕ c n H = x *ℕ quotient-bounded-div-ℕ d n I) + ( inv p) + ( commutative-mul-ℕ c (quotient-bounded-div-ℕ c n H) ∙ + eq-quotient-bounded-div-ℕ c n H ∙ + inv (eq-quotient-bounded-div-ℕ d n I) ∙ + commutative-mul-ℕ (quotient-bounded-div-ℕ d n I) d)) +``` + +### If two natural numbers are equal and one is divisible by a number `d`, then the other is divisible by `d` and their quotients are the same + +The first part of the claim was proven above. Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. + +```agda +eq-quotient-bounded-div-eq-is-nonzero-divisor-ℕ : + {d m n : ℕ} → is-nonzero-ℕ d → m = n → (H : bounded-div-ℕ d m) (K : bounded-div-ℕ d n) → + quotient-bounded-div-ℕ d m H = quotient-bounded-div-ℕ d n K +eq-quotient-bounded-div-eq-is-nonzero-divisor-ℕ N refl H K = + is-injective-right-mul-ℕ _ N + ( eq-quotient-bounded-div-ℕ _ _ H ∙ inv (eq-quotient-bounded-div-ℕ _ _ K)) + +eq-quotient-bounded-div-eq-is-nonzero-ℕ : + {d m n : ℕ} → is-nonzero-ℕ m → m = n → (H : bounded-div-ℕ d m) (K : bounded-div-ℕ d n) → + quotient-bounded-div-ℕ d m H = quotient-bounded-div-ℕ d n K +eq-quotient-bounded-div-eq-is-nonzero-ℕ {zero-ℕ} N refl H K = + ex-falso + ( N + ( inv (eq-quotient-bounded-div-ℕ _ _ H) ∙ + right-zero-law-mul-ℕ (quotient-bounded-div-ℕ _ _ H))) +eq-quotient-bounded-div-eq-is-nonzero-ℕ {succ-ℕ d} N refl H K = + eq-quotient-bounded-div-eq-is-nonzero-divisor-ℕ + ( is-nonzero-succ-ℕ d) refl H K +``` + +### Bounded divisibility is a property + +```agda +abstract + is-prop-bounded-div-ℕ : + (k n : ℕ) → is-prop (bounded-div-ℕ k n) + is-prop-bounded-div-ℕ zero-ℕ n = + is-prop-all-elements-equal + ( λ (q , b , p) (q' , b' , p') → + eq-type-subtype + ( λ u → product-Prop (leq-ℕ-Prop u n) (Id-Prop ℕ-Set _ _)) + ( is-zero-leq-zero-ℕ q + ( concatenate-leq-eq-ℕ q b (inv p ∙ right-zero-law-mul-ℕ q)) ∙ + inv + ( is-zero-leq-zero-ℕ q' + ( concatenate-leq-eq-ℕ q' b' + ( inv p' ∙ right-zero-law-mul-ℕ q'))))) + is-prop-bounded-div-ℕ (succ-ℕ k) n = + is-prop-all-elements-equal + ( λ (q , b , p) (q' , b' , p') → + eq-type-subtype + ( λ u → product-Prop (leq-ℕ-Prop u n) (Id-Prop ℕ-Set _ _)) + ( eq-quotient-bounded-div-eq-divisor-ℕ + ( succ-ℕ k) + ( succ-ℕ k) + ( n) + ( is-nonzero-succ-ℕ k) + ( refl) + ( q , b , p) + ( q' , b' , p'))) + +bounded-div-ℕ-Prop : (k n : ℕ) → Prop lzero +pr1 (bounded-div-ℕ-Prop k n) = bounded-div-ℕ k n +pr2 (bounded-div-ℕ-Prop k n) = is-prop-bounded-div-ℕ k n + +bounded-div-ℕ-Decidable-Prop : (k n : ℕ) → Decidable-Prop lzero +pr1 (bounded-div-ℕ-Decidable-Prop k n) = bounded-div-ℕ k n +pr1 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-prop-bounded-div-ℕ k n +pr2 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-decidable-bounded-div-ℕ k n +``` + +### `0` is bounded divisible by any natural number + +```agda +bounded-div-zero-ℕ : + (n : ℕ) → bounded-div-ℕ n 0 +pr1 (bounded-div-zero-ℕ n) = 0 +pr1 (pr2 (bounded-div-zero-ℕ n)) = refl-leq-ℕ 0 +pr2 (pr2 (bounded-div-zero-ℕ n)) = left-zero-law-mul-ℕ n + +is-zero-quotient-bounded-div-zero-ℕ : + (n : ℕ) (H : bounded-div-ℕ n 0) → + is-zero-ℕ (quotient-bounded-div-ℕ n 0 H) +is-zero-quotient-bounded-div-zero-ℕ n H = + is-zero-leq-zero-ℕ + ( quotient-bounded-div-ℕ n 0 H) + ( upper-bound-quotient-bounded-div-ℕ n 0 H) + +is-zero-quotient-bounded-div-is-zero-ℕ : + (m n : ℕ) (H : bounded-div-ℕ m n) → + is-zero-ℕ n → is-zero-ℕ (quotient-bounded-div-ℕ m n H) +is-zero-quotient-bounded-div-is-zero-ℕ m .zero-ℕ H refl = + is-zero-quotient-bounded-div-zero-ℕ m H +``` + +### If a number is bounded divisible by `0`, then it is `0` + +```agda +is-zero-bounded-div-zero-divisor-ℕ : + (n : ℕ) → bounded-div-ℕ 0 n → is-zero-ℕ n +is-zero-bounded-div-zero-divisor-ℕ n H = + inv (eq-quotient-bounded-div-ℕ 0 n H) ∙ + right-zero-law-mul-ℕ (quotient-bounded-div-ℕ 0 n H) + +is-zero-quotient-bounded-div-zero-divisor-ℕ : + (n : ℕ) (H : bounded-div-ℕ 0 n) → + is-zero-ℕ (quotient-bounded-div-ℕ 0 n H) +is-zero-quotient-bounded-div-zero-divisor-ℕ n H = + is-zero-quotient-bounded-div-is-zero-ℕ 0 n H + ( is-zero-bounded-div-zero-divisor-ℕ n H) + +is-zero-bounded-div-is-zero-divisor-ℕ : + (m n : ℕ) → bounded-div-ℕ m n → is-zero-ℕ m → is-zero-ℕ n +is-zero-bounded-div-is-zero-divisor-ℕ .zero-ℕ n H refl = + is-zero-bounded-div-zero-divisor-ℕ n H + +is-zero-quotient-bounded-div-is-zero-divisor-ℕ : + (m n : ℕ) (H : bounded-div-ℕ m n) → + is-zero-ℕ m → is-zero-ℕ (quotient-bounded-div-ℕ m n H) +is-zero-quotient-bounded-div-is-zero-divisor-ℕ m n H p = + is-zero-quotient-bounded-div-is-zero-ℕ m n H + ( is-zero-bounded-div-is-zero-divisor-ℕ m n H p) +``` + +### The divisibility relation is a partial order on the natural numbers + +The [poset](order-theory.posets.md) of natural numbers ordered by divisibility +is defined in +[`elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility`](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). + +```agda +refl-bounded-div-ℕ : is-reflexive bounded-div-ℕ +pr1 (refl-bounded-div-ℕ zero-ℕ) = 0 +pr1 (pr2 (refl-bounded-div-ℕ zero-ℕ)) = refl-leq-ℕ 0 +pr2 (pr2 (refl-bounded-div-ℕ zero-ℕ)) = refl +pr1 (refl-bounded-div-ℕ (succ-ℕ n)) = 1 +pr1 (pr2 (refl-bounded-div-ℕ (succ-ℕ n))) = leq-zero-ℕ n +pr2 (pr2 (refl-bounded-div-ℕ (succ-ℕ n))) = left-unit-law-mul-ℕ (succ-ℕ n) + +bounded-div-eq-ℕ : (x y : ℕ) → x = y → bounded-div-ℕ x y +bounded-div-eq-ℕ x .x refl = refl-bounded-div-ℕ x + +antisymmetric-bounded-div-ℕ : is-antisymmetric bounded-div-ℕ +antisymmetric-bounded-div-ℕ zero-ℕ n H K = + inv (is-zero-bounded-div-zero-divisor-ℕ n H) +antisymmetric-bounded-div-ℕ (succ-ℕ m) zero-ℕ H K = + is-zero-bounded-div-zero-divisor-ℕ (succ-ℕ m) K +antisymmetric-bounded-div-ℕ (succ-ℕ m) (succ-ℕ n) H K = + antisymmetric-leq-ℕ + ( succ-ℕ m) + ( succ-ℕ n) + ( upper-bound-divisor-bounded-div-ℕ + ( succ-ℕ m) + ( succ-ℕ n) + ( is-nonzero-succ-ℕ n) + ( H)) + ( upper-bound-divisor-bounded-div-ℕ + ( succ-ℕ n) + ( succ-ℕ m) + ( is-nonzero-succ-ℕ m) + ( K)) + +quotient-transitive-bounded-div-ℕ : + (m n k : ℕ) → bounded-div-ℕ n k → bounded-div-ℕ m n → ℕ +quotient-transitive-bounded-div-ℕ m n k K H = + quotient-bounded-div-ℕ n k K *ℕ quotient-bounded-div-ℕ m n H + +eq-quotient-transitive-bounded-div-ℕ : + (m n k : ℕ) (K : bounded-div-ℕ n k) (H : bounded-div-ℕ m n) → + quotient-transitive-bounded-div-ℕ m n k K H *ℕ m = k +eq-quotient-transitive-bounded-div-ℕ m n k K H = + ( associative-mul-ℕ + ( quotient-bounded-div-ℕ n k K) + ( quotient-bounded-div-ℕ m n H) + ( m)) ∙ + ( ap + ( mul-ℕ (quotient-bounded-div-ℕ n k K)) + ( eq-quotient-bounded-div-ℕ m n H)) ∙ + ( eq-quotient-bounded-div-ℕ n k K) + +upper-bound-quotient-transitive-bounded-div-ℕ : + (m n k : ℕ) (K : bounded-div-ℕ n k) (H : bounded-div-ℕ m n) → + quotient-transitive-bounded-div-ℕ m n k K H ≤-ℕ k +upper-bound-quotient-transitive-bounded-div-ℕ zero-ℕ n k K H = + leq-is-zero-ℕ + ( quotient-transitive-bounded-div-ℕ 0 n k K H) + ( k) + ( right-zero-law-mul-is-zero-ℕ + ( quotient-bounded-div-ℕ n k K) + ( quotient-bounded-div-ℕ 0 n H) + ( is-zero-quotient-bounded-div-zero-divisor-ℕ n H)) +upper-bound-quotient-transitive-bounded-div-ℕ (succ-ℕ m) n k K H = + leq-left-factor-eq-ℕ + ( quotient-transitive-bounded-div-ℕ (succ-ℕ m) n k K H) + ( succ-ℕ m) + ( k) + ( is-nonzero-succ-ℕ m) + ( eq-quotient-transitive-bounded-div-ℕ (succ-ℕ m) n k K H) + +transitive-bounded-div-ℕ : is-transitive bounded-div-ℕ +pr1 (transitive-bounded-div-ℕ m n k K H) = + quotient-transitive-bounded-div-ℕ m n k K H +pr1 (pr2 (transitive-bounded-div-ℕ m n k K H)) = + upper-bound-quotient-transitive-bounded-div-ℕ m n k K H +pr2 (pr2 (transitive-bounded-div-ℕ m n k K H)) = + eq-quotient-transitive-bounded-div-ℕ m n k K H +``` diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 920601ec54..113973cf82 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -8,6 +8,7 @@ module elementary-number-theory.divisibility-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.decidable-types open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.equality-natural-numbers @@ -66,8 +67,7 @@ then the type `div-ℕ m n` is always a The divisibility relation is [logically equivalent](foundation.logical-equivalences.md), though not [equivalent](foundation-core.equivalences.md), to the -{{#concept "bounded divisibility" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} -relation, which is defined by +[bounded divisibility relation](elementary-number-theory.bounded-divisibility-natural-numbers.md), which is defined by ```text bounded-div-ℕ m n := Σ (k : ℕ), (k ≤ n) × (k *ℕ m = n). @@ -78,132 +78,73 @@ The discrepancy between divisibility and bounded divisibility is manifested at [contractible](foundation-core.contractible-types.md). For all other values we have `div-ℕ m n ≃ bounded-div-ℕ m n`. -The notion of bounded divisibility has several advantages over ordinary -divisibility: Bounded divisibility is always a proposition, and the proof that -divisibility is decidable uses the proof that bounded divisibility is decidable. +When a natural number `n is divisible by a natural number `m`, with an element `H : div-ℕ m n`, then we define the {{#concept "quotient" Disambiguation="divisibility of natural numbers" Agda=quotient-div-ℕ}} to be the unique number `q ≤ n` such that `q * m = n`. ## Definitions ### The divisibility relation on the natural numbers -```agda -module _ - (m n : ℕ) - where - - div-ℕ : UU lzero - div-ℕ = Σ ℕ (λ k → k *ℕ m = n) - - quotient-div-ℕ : div-ℕ → ℕ - quotient-div-ℕ = pr1 - - eq-quotient-div-ℕ : (H : div-ℕ) → quotient-div-ℕ H *ℕ m = n - eq-quotient-div-ℕ = pr2 - - eq-quotient-div-ℕ' : (H : div-ℕ) → m *ℕ quotient-div-ℕ H = n - eq-quotient-div-ℕ' H = - commutative-mul-ℕ m (quotient-div-ℕ H) ∙ eq-quotient-div-ℕ H -``` +We introduce the divisibility relation on the natural numbers, and some infrastructure. -### The bounded divisibility predicate +**Note:** Perhaps a more natural name for `pr1-div-ℕ`, the number `q` such that `q * d = n`, would be `quotient-div-ℕ`. However, since the divisibility relation is not always a proposition, the quotient could have some undesirable properties. Later in this file, we will define `quotient-div-ℕ` to the the quotient of the bounded divisibility relation, which is logically equivalent to the divisibility relation. ```agda module _ (m n : ℕ) where - bounded-div-ℕ : UU lzero - bounded-div-ℕ = Σ ℕ (λ k → (k ≤-ℕ n) × (k *ℕ m = n)) - - quotient-bounded-div-ℕ : bounded-div-ℕ → ℕ - quotient-bounded-div-ℕ = pr1 + div-ℕ : UU lzero + div-ℕ = Σ ℕ (λ k → k *ℕ m = n) - upper-bound-quotient-bounded-div-ℕ : - (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H ≤-ℕ n - upper-bound-quotient-bounded-div-ℕ H = pr1 (pr2 H) + pr1-div-ℕ : div-ℕ → ℕ + pr1-div-ℕ = pr1 - eq-quotient-bounded-div-ℕ : - (H : bounded-div-ℕ) → quotient-bounded-div-ℕ H *ℕ m = n - eq-quotient-bounded-div-ℕ H = pr2 (pr2 H) + eq-pr1-div-ℕ : (H : div-ℕ) → pr1-div-ℕ H *ℕ m = n + eq-pr1-div-ℕ = pr2 - eq-quotient-bounded-div-ℕ' : - (H : bounded-div-ℕ) → m *ℕ quotient-bounded-div-ℕ H = n - eq-quotient-bounded-div-ℕ' H = - commutative-mul-ℕ m (quotient-bounded-div-ℕ H) ∙ eq-quotient-bounded-div-ℕ H + eq-pr1-div-ℕ' : (H : div-ℕ) → m *ℕ pr1-div-ℕ H = n + eq-pr1-div-ℕ' H = + commutative-mul-ℕ m (pr1-div-ℕ H) ∙ eq-pr1-div-ℕ H - div-bounded-div-ℕ : bounded-div-ℕ → div-ℕ m n - pr1 (div-bounded-div-ℕ H) = quotient-bounded-div-ℕ H - pr2 (div-bounded-div-ℕ H) = eq-quotient-bounded-div-ℕ H + div-bounded-div-ℕ : bounded-div-ℕ m n → div-ℕ + pr1 (div-bounded-div-ℕ H) = quotient-bounded-div-ℕ m n H + pr2 (div-bounded-div-ℕ H) = eq-quotient-bounded-div-ℕ m n H ``` ## Properties -### Concatenating equality and divisibility - -```agda -concatenate-eq-div-ℕ : - {x y z : ℕ} → x = y → div-ℕ y z → div-ℕ x z -concatenate-eq-div-ℕ refl p = p - -concatenate-div-eq-ℕ : - {x y z : ℕ} → div-ℕ x y → y = z → div-ℕ x z -concatenate-div-eq-ℕ p refl = p - -concatenate-eq-div-eq-ℕ : - {x y z w : ℕ} → x = y → div-ℕ y z → z = w → div-ℕ x w -concatenate-eq-div-eq-ℕ refl p refl = p -``` - -### The quotient of `n` by a number `m` it is divisible by also divides `n` - -```agda -div-quotient-div-ℕ : - (m n : ℕ) (H : div-ℕ m n) → div-ℕ (quotient-div-ℕ m n H) n -pr1 (div-quotient-div-ℕ m n (u , p)) = m -pr2 (div-quotient-div-ℕ m n (u , p)) = commutative-mul-ℕ m u ∙ p -``` - -### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal +### If `n` is divisible by a number `m` with proof of divisibility `(q , p)`, then `n` is divisible by the number `q`. ```agda -eq-quotient-div-eq-div-ℕ : - (c d n : ℕ) → is-nonzero-ℕ c → c = d → - (H : div-ℕ c n) → (I : div-ℕ d n) → - quotient-div-ℕ c n H = quotient-div-ℕ d n I -eq-quotient-div-eq-div-ℕ c d n N p H I = - is-injective-left-mul-ℕ c N - ( tr - ( λ x → c *ℕ quotient-div-ℕ c n H = x *ℕ quotient-div-ℕ d n I) - ( inv p) - ( commutative-mul-ℕ c (quotient-div-ℕ c n H) ∙ - eq-quotient-div-ℕ c n H ∙ - inv (eq-quotient-div-ℕ d n I) ∙ - commutative-mul-ℕ (quotient-div-ℕ d n I) d)) +div-pr1-div-ℕ : + (m n : ℕ) (H : div-ℕ m n) → div-ℕ (pr1-div-ℕ m n H) n +pr1 (div-pr1-div-ℕ m n (u , p)) = m +pr2 (div-pr1-div-ℕ m n (u , p)) = commutative-mul-ℕ m u ∙ p ``` ### If `x` is nonzero and `d | x`, then `d ≤ x` ```agda leq-div-succ-ℕ : (d x : ℕ) → div-ℕ d (succ-ℕ x) → leq-ℕ d (succ-ℕ x) -leq-div-succ-ℕ d x (pair (succ-ℕ k) p) = +leq-div-succ-ℕ d x (succ-ℕ k , p) = concatenate-leq-eq-ℕ d (leq-mul-ℕ' k d) p leq-div-ℕ : (d x : ℕ) → is-nonzero-ℕ x → div-ℕ d x → leq-ℕ d x leq-div-ℕ d x f H with is-successor-is-nonzero-ℕ f -... | (pair y refl) = leq-div-succ-ℕ d y H +... | (y , refl) = leq-div-succ-ℕ d y H -leq-quotient-div-ℕ : - (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → leq-ℕ (quotient-div-ℕ d x H) x -leq-quotient-div-ℕ d x f H = - leq-div-ℕ (quotient-div-ℕ d x H) x f (div-quotient-div-ℕ d x H) +leq-pr1-div-ℕ : + (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → leq-ℕ (pr1-div-ℕ d x H) x +leq-pr1-div-ℕ d x f H = + leq-div-ℕ (pr1-div-ℕ d x H) x f (div-pr1-div-ℕ d x H) -leq-quotient-div-ℕ' : - (d x : ℕ) → is-nonzero-ℕ d → (H : div-ℕ d x) → leq-ℕ (quotient-div-ℕ d x H) x -leq-quotient-div-ℕ' d zero-ℕ f (zero-ℕ , p) = star -leq-quotient-div-ℕ' d zero-ℕ f (succ-ℕ n , p) = +leq-pr1-div-ℕ' : + (d x : ℕ) → is-nonzero-ℕ d → (H : div-ℕ d x) → leq-ℕ (pr1-div-ℕ d x H) x +leq-pr1-div-ℕ' d zero-ℕ f (zero-ℕ , p) = star +leq-pr1-div-ℕ' d zero-ℕ f (succ-ℕ n , p) = f (is-zero-right-is-zero-add-ℕ _ d p) -leq-quotient-div-ℕ' d (succ-ℕ x) f H = - leq-quotient-div-ℕ d (succ-ℕ x) (is-nonzero-succ-ℕ x) H +leq-pr1-div-ℕ' d (succ-ℕ x) f H = + leq-pr1-div-ℕ d (succ-ℕ x) (is-nonzero-succ-ℕ x) H ``` ### If `x` is nonzero, if `d | x` and `d ≠ x`, then `d < x` @@ -223,7 +164,7 @@ le-div-ℕ d x H K f = le-leq-neq-ℕ (leq-div-ℕ d x H K) f bounded-div-div-ℕ : (m n : ℕ) → div-ℕ m n → bounded-div-ℕ m n bounded-div-div-ℕ m zero-ℕ (q , p) = (0 , refl-leq-ℕ 0 , left-zero-law-mul-ℕ m) bounded-div-div-ℕ m (succ-ℕ n) (q , p) = - ( q , leq-quotient-div-ℕ m (succ-ℕ n) (is-nonzero-succ-ℕ n) (q , p) , p) + ( q , leq-pr1-div-ℕ m (succ-ℕ n) (is-nonzero-succ-ℕ n) (q , p) , p) logical-equivalence-bounded-div-div-ℕ : (m n : ℕ) → bounded-div-ℕ m n ↔ div-ℕ m n @@ -233,16 +174,98 @@ pr2 (logical-equivalence-bounded-div-div-ℕ m n) = bounded-div-div-ℕ m n ``` -### Bounded divisibility is decidable +### The quotient of a natural number by a number it is divisible by + +```agda +module _ + (m n : ℕ) + where + + quotient-div-ℕ : div-ℕ m n → ℕ + quotient-div-ℕ H = + quotient-bounded-div-ℕ m n (bounded-div-div-ℕ m n H) + + eq-quotient-div-ℕ : (H : div-ℕ m n) → quotient-div-ℕ H *ℕ m = n + eq-quotient-div-ℕ H = + eq-quotient-bounded-div-ℕ m n (bounded-div-div-ℕ m n H) + + eq-quotient-div-ℕ' : (H : div-ℕ m n) → m *ℕ quotient-div-ℕ H = n + eq-quotient-div-ℕ' H = + eq-quotient-bounded-div-ℕ' m n (bounded-div-div-ℕ m n H) + + upper-bound-quotient-div-ℕ : + (H : div-ℕ m n) → quotient-div-ℕ H ≤-ℕ n + upper-bound-quotient-div-ℕ H = + upper-bound-quotient-bounded-div-ℕ m n (bounded-div-div-ℕ m n H) + + div-quotient-div-ℕ : + (H : div-ℕ m n) → div-ℕ (quotient-div-ℕ H) n + div-quotient-div-ℕ H = + div-bounded-div-ℕ + ( quotient-div-ℕ H) + ( n) + ( bounded-div-quotient-bounded-div-ℕ m n (bounded-div-div-ℕ m n H)) +``` + +### Concatenating equality and divisibility + +```agda +concatenate-eq-div-ℕ : + {x y z : ℕ} → x = y → div-ℕ y z → div-ℕ x z +concatenate-eq-div-ℕ refl p = p + +concatenate-div-eq-ℕ : + {x y z : ℕ} → div-ℕ x y → y = z → div-ℕ x z +concatenate-div-eq-ℕ p refl = p + +concatenate-eq-div-eq-ℕ : + {x y z w : ℕ} → x = y → div-ℕ y z → z = w → div-ℕ x w +concatenate-eq-div-eq-ℕ refl p refl = p +``` + +### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal + +Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. + +```agda +eq-quotient-div-eq-divisor-ℕ : + (c d n : ℕ) → is-nonzero-ℕ c → c = d → + (H : div-ℕ c n) → (I : div-ℕ d n) → + pr1-div-ℕ c n H = pr1-div-ℕ d n I +eq-quotient-div-eq-divisor-ℕ c d n N p H I = + is-injective-left-mul-ℕ c N + ( tr + ( λ x → c *ℕ pr1-div-ℕ c n H = x *ℕ pr1-div-ℕ d n I) + ( inv p) + ( commutative-mul-ℕ c (pr1-div-ℕ c n H) ∙ + eq-pr1-div-ℕ c n H ∙ + inv (eq-pr1-div-ℕ d n I) ∙ + commutative-mul-ℕ (pr1-div-ℕ d n I) d)) +``` + +### If two natural numbers are equal and one is divisible by a number `d`, then the other is divisible by `d` and their quotients are the same + +The first part of the claim was proven above. Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. ```agda -is-decidable-bounded-div-ℕ : - (m n : ℕ) → is-decidable (bounded-div-ℕ m n) -is-decidable-bounded-div-ℕ m n = - is-decidable-bounded-Σ-ℕ' - ( n) - ( λ x → mul-ℕ x m = n) - ( λ x → has-decidable-equality-ℕ (mul-ℕ x m) n) +eq-quotient-div-eq-is-nonzero-divisor-ℕ : + {d m n : ℕ} → is-nonzero-ℕ d → m = n → (H : div-ℕ d m) (K : div-ℕ d n) → + quotient-div-ℕ d m H = quotient-div-ℕ d n K +eq-quotient-div-eq-is-nonzero-divisor-ℕ N refl H K = + is-injective-right-mul-ℕ _ N + ( eq-quotient-div-ℕ _ _ H ∙ inv (eq-quotient-div-ℕ _ _ K)) + +eq-quotient-div-eq-is-nonzero-ℕ : + {d m n : ℕ} → is-nonzero-ℕ m → m = n → (H : div-ℕ d m) (K : div-ℕ d n) → + quotient-div-ℕ d m H = quotient-div-ℕ d n K +eq-quotient-div-eq-is-nonzero-ℕ {zero-ℕ} N refl H K = + ex-falso + ( N + ( inv (eq-quotient-div-ℕ _ _ H) ∙ + right-zero-law-mul-ℕ (quotient-div-ℕ _ _ H))) +eq-quotient-div-eq-is-nonzero-ℕ {succ-ℕ d} N refl H K = + eq-quotient-div-eq-is-nonzero-divisor-ℕ + ( is-nonzero-succ-ℕ d) refl H K ``` ### Divisibility is decidable @@ -275,40 +298,6 @@ pr1 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-prop-div-ℕ d x (inl H) pr2 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-decidable-div-ℕ d x ``` -### Bounded divisibility is a property - -```agda -abstract - is-prop-bounded-div-ℕ : - (k n : ℕ) → is-prop (bounded-div-ℕ k n) - is-prop-bounded-div-ℕ zero-ℕ n = - is-prop-all-elements-equal - ( λ (q , b , p) (q' , b' , p') → - eq-type-subtype - ( λ u → product-Prop (leq-ℕ-Prop u n) (Id-Prop ℕ-Set _ _)) - ( is-zero-leq-zero-ℕ q - ( concatenate-leq-eq-ℕ q b (inv p ∙ right-zero-law-mul-ℕ q)) ∙ - inv - ( is-zero-leq-zero-ℕ q' - ( concatenate-leq-eq-ℕ q' b' - ( inv p' ∙ right-zero-law-mul-ℕ q'))))) - is-prop-bounded-div-ℕ (succ-ℕ k) n = - is-prop-equiv - ( equiv-right-swap-Σ ∘e inv-associative-Σ ℕ (λ x → x ≤-ℕ n) _) - ( is-prop-Σ - ( is-prop-div-ℕ (succ-ℕ k) n (inl (is-nonzero-succ-ℕ k))) - ( λ (x , H) → is-prop-leq-ℕ x n)) - -bounded-div-ℕ-Prop : (k n : ℕ) → Prop lzero -pr1 (bounded-div-ℕ-Prop k n) = bounded-div-ℕ k n -pr2 (bounded-div-ℕ-Prop k n) = is-prop-bounded-div-ℕ k n - -bounded-div-ℕ-Decidable-Prop : (k n : ℕ) → Decidable-Prop lzero -pr1 (bounded-div-ℕ-Decidable-Prop k n) = bounded-div-ℕ k n -pr1 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-prop-bounded-div-ℕ k n -pr2 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-decidable-bounded-div-ℕ k n -``` - ### The divisibility relation is a partial order on the natural numbers The [poset](order-theory.posets.md) of natural numbers ordered by divisibility @@ -325,11 +314,11 @@ div-eq-ℕ x .x refl = refl-div-ℕ x antisymmetric-div-ℕ : is-antisymmetric div-ℕ antisymmetric-div-ℕ zero-ℕ zero-ℕ H K = refl -antisymmetric-div-ℕ zero-ℕ (succ-ℕ y) (pair k p) K = +antisymmetric-div-ℕ zero-ℕ (succ-ℕ y) (k , p) K = inv (right-zero-law-mul-ℕ k) ∙ p -antisymmetric-div-ℕ (succ-ℕ x) zero-ℕ H (pair l q) = +antisymmetric-div-ℕ (succ-ℕ x) zero-ℕ H (l , q) = inv q ∙ right-zero-law-mul-ℕ l -antisymmetric-div-ℕ (succ-ℕ x) (succ-ℕ y) (pair k p) (pair l q) = +antisymmetric-div-ℕ (succ-ℕ x) (succ-ℕ y) (k , p) (l , q) = ( inv (left-unit-law-mul-ℕ (succ-ℕ x))) ∙ ( ( ap ( _*ℕ (succ-ℕ x)) @@ -341,8 +330,8 @@ antisymmetric-div-ℕ (succ-ℕ x) (succ-ℕ y) (pair k p) (pair l q) = ( p)) transitive-div-ℕ : is-transitive div-ℕ -pr1 (transitive-div-ℕ x y z (pair l q) (pair k p)) = l *ℕ k -pr2 (transitive-div-ℕ x y z (pair l q) (pair k p)) = +pr1 (transitive-div-ℕ x y z (l , q) (k , p)) = l *ℕ k +pr2 (transitive-div-ℕ x y z (l , q) (k , p)) = associative-mul-ℕ l k x ∙ (ap (l *ℕ_) p ∙ q) ``` @@ -412,7 +401,7 @@ leq-one-div-ℕ d x H L = is-zero-div-ℕ : (d x : ℕ) → le-ℕ x d → div-ℕ d x → is-zero-ℕ x is-zero-div-ℕ d zero-ℕ H D = refl -is-zero-div-ℕ d (succ-ℕ x) H (pair (succ-ℕ k) p) = +is-zero-div-ℕ d (succ-ℕ x) H (succ-ℕ k , p) = ex-falso ( contradiction-le-ℕ ( succ-ℕ x) d H @@ -426,7 +415,7 @@ is-zero-div-ℕ d (succ-ℕ x) H (pair (succ-ℕ k) p) = div-mul-ℕ : (k x y : ℕ) → div-ℕ x y → div-ℕ x (k *ℕ y) div-mul-ℕ k x y H = - transitive-div-ℕ x y (k *ℕ y) (pair k refl) H + transitive-div-ℕ x y (k *ℕ y) (k , refl) H div-mul-ℕ' : (k x y : ℕ) → div-ℕ x y → div-ℕ x (y *ℕ k) @@ -439,19 +428,20 @@ div-mul-ℕ' k x y H = ```agda div-add-ℕ : (d x y : ℕ) → div-ℕ d x → div-ℕ d y → div-ℕ d (x +ℕ y) -pr1 (div-add-ℕ d x y (pair n p) (pair m q)) = n +ℕ m -pr2 (div-add-ℕ d x y (pair n p) (pair m q)) = +pr1 (div-add-ℕ d x y (n , p) (m , q)) = n +ℕ m +pr2 (div-add-ℕ d x y (n , p) (m , q)) = ( right-distributive-mul-add-ℕ n m d) ∙ ( ap-add-ℕ p q) div-left-summand-ℕ : (d x y : ℕ) → div-ℕ d y → div-ℕ d (x +ℕ y) → div-ℕ d x -div-left-summand-ℕ zero-ℕ x y (pair m q) (pair n p) = - pair zero-ℕ +div-left-summand-ℕ zero-ℕ x y (m , q) (n , p) = + pair + ( zero-ℕ) ( ( inv (right-zero-law-mul-ℕ n)) ∙ ( p ∙ (ap (x +ℕ_) ((inv q) ∙ (right-zero-law-mul-ℕ m))))) -pr1 (div-left-summand-ℕ (succ-ℕ d) x y (pair m q) (pair n p)) = dist-ℕ m n -pr2 (div-left-summand-ℕ (succ-ℕ d) x y (pair m q) (pair n p)) = +pr1 (div-left-summand-ℕ (succ-ℕ d) x y (m , q) (n , p)) = dist-ℕ m n +pr2 (div-left-summand-ℕ (succ-ℕ d) x y (m , q) (n , p)) = is-injective-right-add-ℕ (m *ℕ (succ-ℕ d)) ( ( inv ( ( right-distributive-mul-add-ℕ m (dist-ℕ m n) (succ-ℕ d)) ∙ @@ -461,7 +451,7 @@ pr2 (div-left-summand-ℕ (succ-ℕ d) x y (pair m q) (pair n p)) = ( ( ap ( _*ℕ (succ-ℕ d)) ( is-additive-right-inverse-dist-ℕ m n - ( reflects-order-mul-ℕ d m n + ( reflects-order-mul-succ-ℕ d m n ( concatenate-eq-leq-eq-ℕ q ( leq-add-ℕ' y x) ( inv p))))) ∙ @@ -485,8 +475,8 @@ is-one-div-ℕ x y H K = is-one-div-one-ℕ x (div-right-summand-ℕ x y 1 H K) ```agda preserves-div-mul-ℕ : (k x y : ℕ) → div-ℕ x y → div-ℕ (k *ℕ x) (k *ℕ y) -pr1 (preserves-div-mul-ℕ k x y (pair q p)) = q -pr2 (preserves-div-mul-ℕ k x y (pair q p)) = +pr1 (preserves-div-mul-ℕ k x y (q , p)) = q +pr2 (preserves-div-mul-ℕ k x y (q , p)) = ( inv (associative-mul-ℕ q k x)) ∙ ( ( ap (_*ℕ x) (commutative-mul-ℕ q k)) ∙ ( ( associative-mul-ℕ k q x) ∙ @@ -498,8 +488,8 @@ pr2 (preserves-div-mul-ℕ k x y (pair q p)) = ```agda reflects-div-mul-ℕ : (k x y : ℕ) → is-nonzero-ℕ k → div-ℕ (k *ℕ x) (k *ℕ y) → div-ℕ x y -pr1 (reflects-div-mul-ℕ k x y H (pair q p)) = q -pr2 (reflects-div-mul-ℕ k x y H (pair q p)) = +pr1 (reflects-div-mul-ℕ k x y H (q , p)) = q +pr2 (reflects-div-mul-ℕ k x y H (q , p)) = is-injective-left-mul-ℕ k H ( ( inv (associative-mul-ℕ k q x)) ∙ ( ( ap (_*ℕ x) (commutative-mul-ℕ k q)) ∙ @@ -535,8 +525,8 @@ div-div-quotient-div-ℕ x y d H K = is-nonzero-quotient-div-ℕ : {d x : ℕ} (H : div-ℕ d x) → is-nonzero-ℕ x → is-nonzero-ℕ (quotient-div-ℕ d x H) -is-nonzero-quotient-div-ℕ {d} {.(k *ℕ d)} (pair k refl) = - is-nonzero-left-factor-mul-ℕ k d +is-nonzero-quotient-div-ℕ {d} {x} H K = + is-nonzero-quotient-bounded-div-ℕ d x K (bounded-div-div-ℕ d x H) ``` ### If `d` divides a number `1 ≤ x`, then `1 ≤ x/d` @@ -568,8 +558,7 @@ is-idempotent-quotient-div-ℕ (succ-ℕ a) nz (u , p) = simplify-mul-quotient-div-ℕ : {a b c : ℕ} → is-nonzero-ℕ c → (H : div-ℕ b a) (K : div-ℕ c b) (L : div-ℕ c a) → - ( (quotient-div-ℕ b a H) *ℕ (quotient-div-ℕ c b K)) = - ( quotient-div-ℕ c a L) + quotient-div-ℕ b a H *ℕ quotient-div-ℕ c b K = quotient-div-ℕ c a L simplify-mul-quotient-div-ℕ {a} {b} {c} nz H K L = is-injective-right-mul-ℕ c nz ( equational-reasoning @@ -649,7 +638,8 @@ is-one-quotient-div-ℕ d .d f H refl = is-idempotent-quotient-div-ℕ d f H eq-is-one-quotient-div-ℕ : (d x : ℕ) → (H : div-ℕ d x) → is-one-ℕ (quotient-div-ℕ d x H) → d = x -eq-is-one-quotient-div-ℕ d x (.1 , q) refl = inv (left-unit-law-mul-ℕ d) ∙ q +eq-is-one-quotient-div-ℕ d (succ-ℕ x) (.1 , q) refl = + inv (left-unit-law-mul-ℕ d) ∙ q ``` ### If `x` is nonzero and `d | x`, then `x/d = x` if and only if `d = 1` @@ -657,14 +647,19 @@ eq-is-one-quotient-div-ℕ d x (.1 , q) refl = inv (left-unit-law-mul-ℕ d) ∙ ```agda compute-quotient-div-is-one-ℕ : (d x : ℕ) → (H : div-ℕ d x) → is-one-ℕ d → quotient-div-ℕ d x H = x -compute-quotient-div-is-one-ℕ .1 x (u , q) refl = +compute-quotient-div-is-one-ℕ .1 zero-ℕ H refl = refl +compute-quotient-div-is-one-ℕ .1 (succ-ℕ x) (u , q) refl = inv (right-unit-law-mul-ℕ u) ∙ q is-one-divisor-ℕ : (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → quotient-div-ℕ d x H = x → is-one-ℕ d -is-one-divisor-ℕ d x f (.x , q) refl = - is-injective-left-mul-ℕ x f (q ∙ inv (right-unit-law-mul-ℕ x)) +is-one-divisor-ℕ d zero-ℕ N H p = ex-falso (N refl) +is-one-divisor-ℕ d (succ-ℕ x) N (.(succ-ℕ x) , q) refl = + is-injective-left-mul-ℕ + ( succ-ℕ x) + ( N) + ( q ∙ inv (right-unit-law-mul-ℕ (succ-ℕ x))) ``` ### If `x` is nonzero and `d | x` is a nontrivial divisor, then `x/d < x` @@ -681,5 +676,13 @@ le-quotient-div-ℕ d x f H g = ( eq-or-le-leq-ℕ ( quotient-div-ℕ d x H) ( x) - ( leq-quotient-div-ℕ d x f H)) + ( upper-bound-quotient-div-ℕ d x H)) ``` + +## See also + +- [Bounded divisibility of natural numbers](elementary-number-theory.bounded-divisibility-natural-numbers.md) +- [Greatest common divisors of natural numbers](elementary-number-theory.greatest-common-divisors-natural-numbers.md) +- [Nontrivial divisors of natural numbers](elementary-number-theory.nontrivial-divisors-natural-numbers.md) +- [Proper divisors of natural numbers](elementary-number-theory.proper-divisors-natural-numbers.md) +- [The poset of natural numbers ordered by divisibility](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md) diff --git a/src/elementary-number-theory/euclid-mullin-sequence.lagda.md b/src/elementary-number-theory/euclid-mullin-sequence.lagda.md index be53df7a65..921c25be1b 100644 --- a/src/elementary-number-theory/euclid-mullin-sequence.lagda.md +++ b/src/elementary-number-theory/euclid-mullin-sequence.lagda.md @@ -7,8 +7,8 @@ module elementary-number-theory.euclid-mullin-sequence where
Imports ```agda -open import elementary-number-theory.fundamental-theorem-of-arithmetic open import elementary-number-theory.natural-numbers +open import elementary-number-theory.nontrivial-divisors-natural-numbers open import elementary-number-theory.products-of-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import elementary-number-theory.strong-induction-natural-numbers diff --git a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md index c0317e9e84..028788652f 100644 --- a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md +++ b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md @@ -10,6 +10,7 @@ module elementary-number-theory.euclidean-division-natural-numbers where open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.congruence-natural-numbers open import elementary-number-theory.distance-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -144,3 +145,17 @@ quotient-euclidean-division-ℕ' zero-ℕ n = zero-ℕ quotient-euclidean-division-ℕ' (succ-ℕ k) n = array-quotient-euclidean-division-ℕ k n k ``` + +## Properties + +### If `d` divides a number `n`, then its quotient by euclidean division is its quotient by division + +Suppose `q * d = n`. Then the congruence class `r` of `n` modulo `d` is `0`, so the distance between `r` and `n` is `n`. + +```text +compute-euclidean-division-div-ℕ : + (d n : ℕ) (H : div-ℕ d n) → + quotient-euclidean-division-ℕ d n = quotient-div-ℕ d n H +compute-euclidean-division-div-ℕ d n H = + {!eq-quotient-div-eq-is-nonzero-divisor-ℕ!} +``` diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index 93514b96e5..f01148e3af 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -13,11 +13,14 @@ open import elementary-number-theory.decidable-total-order-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.lists-of-prime-numbers open import elementary-number-theory.lower-bounds-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-lists-of-natural-numbers open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.multiplicative-decompositions-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.nontrivial-divisors-natural-numbers open import elementary-number-theory.prime-numbers open import elementary-number-theory.relatively-prime-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers @@ -73,7 +76,7 @@ in several ways: [equal](foundation-core.identity-types.md) to `n` is contractible. Note that the [univalence axiom](foundation-core.univalence.md) is neccessary to -prove the second uniqueness property of prime factorizations. +prove the second uniqueness property of prime factorizations. On this page we will prove the fundamental theorem using the first approach. The fundamental theorem of arithmetic is the 80th theorem on [Freek Wiedijk's](http://www.cs.ru.nl/F.Wiedijk/) list of @@ -87,61 +90,88 @@ A list of natural numbers is a prime decomposition of a natural number `n` if: - The list is sorted. - Every element of the list is prime. -- The product of the element of the list is equal to `n`. +- The list is a [multiplicative decomposition](elementary-number-theory.multiplicative-decompositions-natural-numbers.md) of `n`. ```agda -is-prime-list-ℕ : - list ℕ → UU lzero -is-prime-list-ℕ = for-all-list ℕ is-prime-ℕ-Prop - -is-prop-is-prime-list-ℕ : - (l : list ℕ) → is-prop (is-prime-list-ℕ l) -is-prop-is-prime-list-ℕ = is-prop-for-all-list ℕ is-prime-ℕ-Prop - -is-prime-list-primes-ℕ : - (l : list Prime-ℕ) → is-prime-list-ℕ (map-list nat-Prime-ℕ l) -is-prime-list-primes-ℕ nil = raise-star -is-prime-list-primes-ℕ (cons x l) = - (is-prime-Prime-ℕ x) , (is-prime-list-primes-ℕ l) - module _ (x : ℕ) (l : list ℕ) where - is-decomposition-list-ℕ : - UU lzero - is-decomposition-list-ℕ = - mul-list-ℕ l = x - - is-prop-is-decomposition-list-ℕ : - is-prop (is-decomposition-list-ℕ) - is-prop-is-decomposition-list-ℕ = is-set-ℕ (mul-list-ℕ l) x - is-prime-decomposition-list-ℕ : UU lzero is-prime-decomposition-list-ℕ = is-sorted-list ℕ-Decidable-Total-Order l × ( is-prime-list-ℕ l × - is-decomposition-list-ℕ) + is-multiplicative-decomposition-list-ℕ x l) is-sorted-list-is-prime-decomposition-list-ℕ : is-prime-decomposition-list-ℕ → is-sorted-list ℕ-Decidable-Total-Order l - is-sorted-list-is-prime-decomposition-list-ℕ D = pr1 D + is-sorted-list-is-prime-decomposition-list-ℕ D = + pr1 D is-prime-list-is-prime-decomposition-list-ℕ : - is-prime-decomposition-list-ℕ → - is-prime-list-ℕ l + is-prime-decomposition-list-ℕ → is-prime-list-ℕ l is-prime-list-is-prime-decomposition-list-ℕ D = pr1 (pr2 D) - is-decomposition-list-is-prime-decomposition-list-ℕ : + is-multiplicative-decomposition-is-prime-decomposition-list-ℕ : is-prime-decomposition-list-ℕ → - is-decomposition-list-ℕ - is-decomposition-list-is-prime-decomposition-list-ℕ D = + is-multiplicative-decomposition-list-ℕ x l + is-multiplicative-decomposition-is-prime-decomposition-list-ℕ D = pr2 (pr2 D) + le-one-list-is-prime-decomposition-list-ℕ : + is-prime-decomposition-list-ℕ → + for-all-list ℕ (le-ℕ-Prop 1) l + le-one-list-is-prime-decomposition-list-ℕ H = + map-for-all-list + ( is-prime-ℕ-Prop) + ( le-ℕ-Prop 1) + ( le-one-is-prime-ℕ) + ( l) + ( is-prime-list-is-prime-decomposition-list-ℕ H) + + leq-one-list-is-prime-decomposition-list-ℕ : + is-prime-decomposition-list-ℕ → + for-all-list ℕ (leq-ℕ-Prop 1) l + leq-one-list-is-prime-decomposition-list-ℕ H = + map-for-all-list + ( le-ℕ-Prop 1) + ( leq-ℕ-Prop 1) + ( leq-le-ℕ 1) + ( l) + ( le-one-list-is-prime-decomposition-list-ℕ H) + + eq-is-prime-decomposition-list-ℕ : + is-prime-decomposition-list-ℕ → + mul-list-ℕ l = x + eq-is-prime-decomposition-list-ℕ H = + pr2 (is-multiplicative-decomposition-is-prime-decomposition-list-ℕ H) + + leq-one-is-prime-decomposition-list-ℕ : + is-prime-decomposition-list-ℕ → 1 ≤-ℕ x + leq-one-is-prime-decomposition-list-ℕ H = + tr + ( leq-ℕ 1) + ( eq-is-prime-decomposition-list-ℕ H) + ( leq-one-mul-list-ℕ l (leq-one-list-is-prime-decomposition-list-ℕ H)) + + is-list-of-nontrivial-divisors-is-prime-dicomposition-list-ℕ : + is-prime-decomposition-list-ℕ → + is-list-of-nontrivial-divisors-ℕ x l + is-list-of-nontrivial-divisors-is-prime-dicomposition-list-ℕ H = + is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ x l + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ H) + + is-list-of-nontrivial-divisors-is-prime-decomposition-list-ℕ : + is-prime-decomposition-list-ℕ → + is-list-of-nontrivial-divisors-ℕ x l + is-list-of-nontrivial-divisors-is-prime-decomposition-list-ℕ D = + is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ x l + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ D) + is-prop-is-prime-decomposition-list-ℕ : is-prop (is-prime-decomposition-list-ℕ) is-prop-is-prime-decomposition-list-ℕ = @@ -149,323 +179,123 @@ module _ ( is-prop-is-sorted-list ℕ-Decidable-Total-Order l) ( is-prop-product ( is-prop-is-prime-list-ℕ l) - ( is-prop-is-decomposition-list-ℕ)) + ( is-prop-is-multiplicative-decomposition-list-ℕ x l)) is-prime-decomposition-list-ℕ-Prop : Prop lzero pr1 is-prime-decomposition-list-ℕ-Prop = is-prime-decomposition-list-ℕ pr2 is-prime-decomposition-list-ℕ-Prop = is-prop-is-prime-decomposition-list-ℕ -``` - -### Nontrivial divisors - -Nontrivial divisors of a natural number are divisors strictly greater than `1`. -```agda -is-nontrivial-divisor-ℕ : (n x : ℕ) → UU lzero -is-nontrivial-divisor-ℕ n x = (le-ℕ 1 x) × (div-ℕ x n) - -is-prop-is-nontrivial-divisor-ℕ : - (n x : ℕ) → is-prop (is-nontrivial-divisor-ℕ n x) -is-prop-is-nontrivial-divisor-ℕ n x = - is-prop-Σ - ( is-prop-le-ℕ 1 x) - ( λ p → is-prop-div-ℕ x n (inl (is-nonzero-le-ℕ 1 x p))) - -is-nontrivial-divisor-ℕ-Prop : (n x : ℕ) → Prop lzero -pr1 (is-nontrivial-divisor-ℕ-Prop n x) = is-nontrivial-divisor-ℕ n x -pr2 (is-nontrivial-divisor-ℕ-Prop n x) = is-prop-is-nontrivial-divisor-ℕ n x - -is-decidable-is-nontrivial-divisor-ℕ : - (n x : ℕ) → is-decidable (is-nontrivial-divisor-ℕ n x) -is-decidable-is-nontrivial-divisor-ℕ n x = - is-decidable-product (is-decidable-le-ℕ 1 x) (is-decidable-div-ℕ x n) - -is-nontrivial-divisor-diagonal-ℕ : - (n : ℕ) → le-ℕ 1 n → is-nontrivial-divisor-ℕ n n -pr1 (is-nontrivial-divisor-diagonal-ℕ n H) = H -pr2 (is-nontrivial-divisor-diagonal-ℕ n H) = refl-div-ℕ n -``` +module _ + (x y : ℕ) (l : list ℕ) (H : is-prime-decomposition-list-ℕ x (cons y l)) + where -If `l` is a prime decomposition of `n`, then `l` is a list of nontrivial -divisors of `n`. + is-prime-head-is-prime-decomposition-cons-list-ℕ : + is-prime-ℕ y + is-prime-head-is-prime-decomposition-cons-list-ℕ = + pr1 (is-prime-list-is-prime-decomposition-list-ℕ x (cons y l) H) -```agda -is-list-of-nontrivial-divisors-ℕ : - ℕ → list ℕ → UU lzero -is-list-of-nontrivial-divisors-ℕ x nil = unit -is-list-of-nontrivial-divisors-ℕ x (cons y l) = - (is-nontrivial-divisor-ℕ x y) × (is-list-of-nontrivial-divisors-ℕ x l) - -is-nontrivial-divisors-div-list-ℕ : - (x y : ℕ) → div-ℕ x y → (l : list ℕ) → - is-list-of-nontrivial-divisors-ℕ x l → is-list-of-nontrivial-divisors-ℕ y l -is-nontrivial-divisors-div-list-ℕ x y d nil H = star -is-nontrivial-divisors-div-list-ℕ x y d (cons z l) H = - ( pr1 (pr1 H) , transitive-div-ℕ z x y d (pr2 (pr1 H))) , - is-nontrivial-divisors-div-list-ℕ x y d l (pr2 H) - -is-divisor-head-is-decomposition-list-ℕ : - (x : ℕ) (y : ℕ) (l : list ℕ) → - is-decomposition-list-ℕ x (cons y l) → - div-ℕ y x -pr1 (is-divisor-head-is-decomposition-list-ℕ x y l D) = - mul-list-ℕ l -pr2 (is-divisor-head-is-decomposition-list-ℕ x y l D) = - commutative-mul-ℕ (mul-list-ℕ l) y ∙ D - -is-nontrivial-divisor-head-is-decomposition-is-prime-list-ℕ : - (x : ℕ) (y : ℕ) (l : list ℕ) → - is-decomposition-list-ℕ x (cons y l) → - is-prime-list-ℕ (cons y l) → - is-nontrivial-divisor-ℕ x y -pr1 (is-nontrivial-divisor-head-is-decomposition-is-prime-list-ℕ x y l D P) = - le-one-is-prime-ℕ y (pr1 P) -pr2 (is-nontrivial-divisor-head-is-decomposition-is-prime-list-ℕ x y l D P) = - is-divisor-head-is-decomposition-list-ℕ x y l D - -is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ : - (x : ℕ) → (l : list ℕ) → - is-decomposition-list-ℕ x l → - is-prime-list-ℕ l → - is-list-of-nontrivial-divisors-ℕ x l -is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ x nil _ _ = star -is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ - ( x) - ( cons y l) - ( D) - ( P) = - ( is-nontrivial-divisor-head-is-decomposition-is-prime-list-ℕ x y l D P , - is-nontrivial-divisors-div-list-ℕ - ( mul-list-ℕ l) + is-prime-tail-is-prime-decomposition-cons-list-ℕ : + is-prime-list-ℕ l + is-prime-tail-is-prime-decomposition-cons-list-ℕ = + pr2 (is-prime-list-is-prime-decomposition-list-ℕ x (cons y l) H) + + le-one-head-is-prime-decomposition-cons-list-ℕ : + 1 <-ℕ y + le-one-head-is-prime-decomposition-cons-list-ℕ = + le-one-is-prime-ℕ y is-prime-head-is-prime-decomposition-cons-list-ℕ + + div-head-is-prime-decomposition-cons-list-ℕ : + div-ℕ y x + div-head-is-prime-decomposition-cons-list-ℕ = + div-head-is-multiplicative-decomposition-cons-list-ℕ ( x) - ( y , D) + ( y) ( l) - ( is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ - ( mul-list-ℕ l) - ( l) - ( refl) - ( pr2 P))) - -is-divisor-head-prime-decomposition-list-ℕ : - (x : ℕ) (y : ℕ) (l : list ℕ) → - is-prime-decomposition-list-ℕ x (cons y l) → - div-ℕ y x -is-divisor-head-prime-decomposition-list-ℕ x y l D = - is-divisor-head-is-decomposition-list-ℕ - ( x) - ( y) - ( l) - ( is-decomposition-list-is-prime-decomposition-list-ℕ x (cons y l) D) - -is-nontrivial-divisor-head-prime-decomposition-list-ℕ : - (x : ℕ) (y : ℕ) (l : list ℕ) → - is-prime-decomposition-list-ℕ x (cons y l) → - is-nontrivial-divisor-ℕ x y -is-nontrivial-divisor-head-prime-decomposition-list-ℕ x y l D = - is-nontrivial-divisor-head-is-decomposition-is-prime-list-ℕ - ( x) - ( y) - ( l) - ( is-decomposition-list-is-prime-decomposition-list-ℕ x (cons y l) D) - ( is-prime-list-is-prime-decomposition-list-ℕ x (cons y l) D) - -is-list-of-nontrivial-divisors-is-prime-decomposition-list-ℕ : - (x : ℕ) → (l : list ℕ) → - is-prime-decomposition-list-ℕ x l → - is-list-of-nontrivial-divisors-ℕ x l -is-list-of-nontrivial-divisors-is-prime-decomposition-list-ℕ x l D = - is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ - ( x) - ( l) - ( is-decomposition-list-is-prime-decomposition-list-ℕ x l D) - ( is-prime-list-is-prime-decomposition-list-ℕ x l D) -``` - -## Lemmas - -### Every natural number strictly greater than `1` has a least nontrivial divisor - -```agda -least-nontrivial-divisor-ℕ : - (n : ℕ) → le-ℕ 1 n → minimal-element-ℕ (is-nontrivial-divisor-ℕ n) -least-nontrivial-divisor-ℕ n H = - well-ordering-principle-ℕ - ( is-nontrivial-divisor-ℕ n) - ( is-decidable-is-nontrivial-divisor-ℕ n) - ( n , is-nontrivial-divisor-diagonal-ℕ n H) - -nat-least-nontrivial-divisor-ℕ : (n : ℕ) → le-ℕ 1 n → ℕ -nat-least-nontrivial-divisor-ℕ n H = pr1 (least-nontrivial-divisor-ℕ n H) - -nat-least-nontrivial-divisor-ℕ' : ℕ → ℕ -nat-least-nontrivial-divisor-ℕ' zero-ℕ = 0 -nat-least-nontrivial-divisor-ℕ' (succ-ℕ zero-ℕ) = 1 -nat-least-nontrivial-divisor-ℕ' (succ-ℕ (succ-ℕ n)) = - nat-least-nontrivial-divisor-ℕ (succ-ℕ (succ-ℕ n)) star - -le-one-least-nontrivial-divisor-ℕ : - (n : ℕ) (H : le-ℕ 1 n) → le-ℕ 1 (nat-least-nontrivial-divisor-ℕ n H) -le-one-least-nontrivial-divisor-ℕ n H = - pr1 (pr1 (pr2 (least-nontrivial-divisor-ℕ n H))) - -div-least-nontrivial-divisor-ℕ : - (n : ℕ) (H : le-ℕ 1 n) → div-ℕ (nat-least-nontrivial-divisor-ℕ n H) n -div-least-nontrivial-divisor-ℕ n H = - pr2 (pr1 (pr2 (least-nontrivial-divisor-ℕ n H))) - -is-minimal-least-nontrivial-divisor-ℕ : - (n : ℕ) (H : le-ℕ 1 n) (x : ℕ) (K : le-ℕ 1 x) (d : div-ℕ x n) → - leq-ℕ (nat-least-nontrivial-divisor-ℕ n H) x -is-minimal-least-nontrivial-divisor-ℕ n H x K d = - pr2 (pr2 (least-nontrivial-divisor-ℕ n H)) x (K , d) -``` - -### The least nontrivial divisor of a number `> 1` is nonzero - -```agda -abstract - is-nonzero-least-nontrivial-divisor-ℕ : - (n : ℕ) (H : le-ℕ 1 n) → is-nonzero-ℕ (nat-least-nontrivial-divisor-ℕ n H) - is-nonzero-least-nontrivial-divisor-ℕ n H = - is-nonzero-div-ℕ - ( nat-least-nontrivial-divisor-ℕ n H) - ( n) - ( div-least-nontrivial-divisor-ℕ n H) - ( λ where refl → H) -``` - -### The least nontrivial divisor of a number `> 1` is prime + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ x + ( cons y l) + ( H)) -```agda -is-prime-least-nontrivial-divisor-ℕ : - (n : ℕ) (H : le-ℕ 1 n) → is-prime-ℕ (nat-least-nontrivial-divisor-ℕ n H) -pr1 (is-prime-least-nontrivial-divisor-ℕ n H x) (K , L) = - map-right-unit-law-coproduct-is-empty - ( is-one-ℕ x) - ( le-ℕ 1 x) - ( λ p → - contradiction-le-ℕ x l - ( le-div-ℕ x l - ( is-nonzero-least-nontrivial-divisor-ℕ n H) - ( L) - ( K)) - ( is-minimal-least-nontrivial-divisor-ℕ n H x p - ( transitive-div-ℕ x l n (div-least-nontrivial-divisor-ℕ n H) L))) - ( eq-or-le-leq-ℕ' 1 x - ( leq-one-div-ℕ x n - ( transitive-div-ℕ x l n - ( div-least-nontrivial-divisor-ℕ n H) L) - ( leq-le-ℕ 1 n H))) - where - l = nat-least-nontrivial-divisor-ℕ n H -pr1 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = - neq-le-ℕ (le-one-least-nontrivial-divisor-ℕ n H) -pr2 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = - div-one-ℕ _ + is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ : + is-nontrivial-divisor-ℕ x y + is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ = + is-nontrivial-divisor-head-is-multiplicative-decomposition-cons-list-ℕ + ( x) + ( y) + ( l) + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ x + ( cons y l) + ( H)) + + le-one-is-prime-decomposition-cons-list-ℕ : + le-ℕ 1 x + le-one-is-prime-decomposition-cons-list-ℕ = + concatenate-le-leq-ℕ + {x = 1} + {y = y} + {z = x} + ( le-one-head-is-prime-decomposition-cons-list-ℕ) + ( leq-div-ℕ + ( y) + ( x) + ( is-nonzero-leq-one-ℕ x + ( leq-one-is-prime-decomposition-list-ℕ x (cons y l) H)) + ( div-head-is-prime-decomposition-cons-list-ℕ)) ``` -### The least prime divisor of a number `1 < n` +### Definition of the type of prime decomposition of a natural number ```agda -nat-least-prime-divisor-ℕ : (x : ℕ) → le-ℕ 1 x → ℕ -nat-least-prime-divisor-ℕ x H = nat-least-nontrivial-divisor-ℕ x H - -is-prime-least-prime-divisor-ℕ : - (x : ℕ) (H : le-ℕ 1 x) → is-prime-ℕ (nat-least-prime-divisor-ℕ x H) -is-prime-least-prime-divisor-ℕ x H = is-prime-least-nontrivial-divisor-ℕ x H - -least-prime-divisor-ℕ : (x : ℕ) → le-ℕ 1 x → Prime-ℕ -pr1 (least-prime-divisor-ℕ x H) = nat-least-prime-divisor-ℕ x H -pr2 (least-prime-divisor-ℕ x H) = is-prime-least-prime-divisor-ℕ x H - -div-least-prime-divisor-ℕ : - (x : ℕ) (H : le-ℕ 1 x) → div-ℕ (nat-least-prime-divisor-ℕ x H) x -div-least-prime-divisor-ℕ x H = div-least-nontrivial-divisor-ℕ x H - -quotient-div-least-prime-divisor-ℕ : - (x : ℕ) (H : le-ℕ 1 x) → ℕ -quotient-div-least-prime-divisor-ℕ x H = - quotient-div-ℕ - ( nat-least-prime-divisor-ℕ x H) - ( x) - ( div-least-prime-divisor-ℕ x H) - -leq-quotient-div-least-prime-divisor-ℕ : - (x : ℕ) (H : le-ℕ 1 (succ-ℕ x)) → - leq-ℕ (quotient-div-least-prime-divisor-ℕ (succ-ℕ x) H) x -leq-quotient-div-least-prime-divisor-ℕ x H = - leq-quotient-div-is-prime-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ x) H) - ( x) - ( is-prime-least-prime-divisor-ℕ (succ-ℕ x) H) - ( div-least-prime-divisor-ℕ (succ-ℕ x) H) +prime-decomposition-list-ℕ : + (x : ℕ) → 1 ≤-ℕ x → UU lzero +prime-decomposition-list-ℕ x _ = + Σ (list ℕ) (λ l → is-prime-decomposition-list-ℕ x l) ``` -## The fundamental theorem of arithmetic (with lists) +## The fundamental theorem of arithmetic ### Existence -#### The list given by the fundamental theorem of arithmetic +#### The list of primes in the prime decomposition of a number ```agda -list-primes-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → leq-ℕ 1 x → list Prime-ℕ -list-primes-fundamental-theorem-arithmetic-ℕ zero-ℕ () -list-primes-fundamental-theorem-arithmetic-ℕ (succ-ℕ x) H = +successor-step-list-of-primes-fundamental-theorem-arithmetic-ℕ : + (n : ℕ) → 1 ≤-ℕ n → ((m : ℕ) → 1 ≤-ℕ m → m ≤-ℕ n → list Prime-ℕ) → + list Prime-ℕ +successor-step-list-of-primes-fundamental-theorem-arithmetic-ℕ n N f = + cons + ( prime-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) + ( f + ( quotient-div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) + ( leq-one-quotient-div-ℕ + ( nat-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) + ( succ-ℕ n) + ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) + ( preserves-leq-succ-ℕ 1 n N)) + ( upper-bound-quotient-div-least-prime-divisor-ℕ n (le-succ-leq-ℕ 1 n N))) + +list-of-primes-fundamental-theorem-arithmetic-ℕ : + (n : ℕ) → 1 ≤-ℕ n → list Prime-ℕ +list-of-primes-fundamental-theorem-arithmetic-ℕ = based-strong-ind-ℕ 1 - ( λ _ → list Prime-ℕ) + ( λ n → list Prime-ℕ) ( nil) - ( λ n N f → - cons - ( least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( f - ( quotient-div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( succ-ℕ n) - ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( preserves-leq-succ-ℕ 1 n N)) - ( leq-quotient-div-least-prime-divisor-ℕ n (le-succ-leq-ℕ 1 n N)))) - ( succ-ℕ x) - ( H) + ( successor-step-list-of-primes-fundamental-theorem-arithmetic-ℕ) list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → leq-ℕ 1 x → list ℕ + (x : ℕ) → 1 ≤-ℕ x → list ℕ list-fundamental-theorem-arithmetic-ℕ x H = - map-list nat-Prime-ℕ (list-primes-fundamental-theorem-arithmetic-ℕ x H) + map-list nat-Prime-ℕ (list-of-primes-fundamental-theorem-arithmetic-ℕ x H) ``` -#### Computational rules for the list given by the fundamental theorem of arithmetic +#### Computational rules for the list of primes in the prime decomposition of a number ```agda -helper-compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → - based-strong-ind-ℕ 1 - ( λ _ → list Prime-ℕ) - ( nil) - ( λ n N f → - cons - ( least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( f - ( quotient-div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( succ-ℕ n) - ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( preserves-leq-succ-ℕ 1 n N)) - ( leq-quotient-div-least-prime-divisor-ℕ n (le-succ-leq-ℕ 1 n N)))) - ( x) - ( H) = - list-primes-fundamental-theorem-arithmetic-ℕ x H -helper-compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ (succ-ℕ x) H = - refl - -compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ : - (x : ℕ) → (H : 1 ≤-ℕ x) → - list-primes-fundamental-theorem-arithmetic-ℕ (succ-ℕ x) star = +compute-list-of-primes-fundamental-theorem-arithmetic-succ-ℕ : + (x : ℕ) (H : 1 ≤-ℕ x) → + list-of-primes-fundamental-theorem-arithmetic-ℕ (succ-ℕ x) star = cons - ( least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( list-primes-fundamental-theorem-arithmetic-ℕ + ( prime-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) + ( list-of-primes-fundamental-theorem-arithmetic-ℕ ( quotient-div-least-prime-divisor-ℕ ( succ-ℕ x) ( le-succ-leq-ℕ 1 x H)) @@ -474,40 +304,19 @@ compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ : ( succ-ℕ x) ( div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) ( preserves-leq-succ-ℕ 1 x H))) -compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ x H = +compute-list-of-primes-fundamental-theorem-arithmetic-succ-ℕ x H = compute-succ-based-strong-ind-ℕ ( 1) ( λ _ → list Prime-ℕ) ( nil) - ( λ n N f → - cons - ( least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( f - ( quotient-div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( succ-ℕ n) - ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( preserves-leq-succ-ℕ 1 n N)) - ( leq-quotient-div-least-prime-divisor-ℕ n (le-succ-leq-ℕ 1 n N)))) + ( successor-step-list-of-primes-fundamental-theorem-arithmetic-ℕ) ( x) ( H) - ( star) ∙ - ap - ( cons (least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H))) - ( helper-compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ - ( quotient-div-least-prime-divisor-ℕ - ( succ-ℕ x) - ( le-succ-leq-ℕ 1 x H)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( succ-ℕ x) - ( div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( preserves-leq-succ-ℕ 1 x H))) + ( star) compute-list-fundamental-theorem-arithmetic-succ-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → - list-fundamental-theorem-arithmetic-ℕ (succ-ℕ x) star = + (x : ℕ) (H : 1 ≤-ℕ x) → + list-fundamental-theorem-arithmetic-ℕ (succ-ℕ x) _ = cons ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) ( list-fundamental-theorem-arithmetic-ℕ @@ -522,34 +331,50 @@ compute-list-fundamental-theorem-arithmetic-succ-ℕ : compute-list-fundamental-theorem-arithmetic-succ-ℕ x H = ap ( map-list nat-Prime-ℕ) - ( compute-list-primes-fundamental-theorem-arithmetic-succ-ℕ x H) + ( compute-list-of-primes-fundamental-theorem-arithmetic-succ-ℕ x H) ``` -#### Proof that the list given by the fundamental theorem of arithmetic is a prime decomposition +#### The list of primes in the prime decomposition of a number is indeed a prime decomposition ```agda is-prime-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-prime-list-ℕ (list-fundamental-theorem-arithmetic-ℕ x H) is-prime-list-fundamental-theorem-arithmetic-ℕ x H = - is-prime-list-primes-ℕ (list-primes-fundamental-theorem-arithmetic-ℕ x H) + is-prime-list-list-Prime-ℕ + ( list-of-primes-fundamental-theorem-arithmetic-ℕ x H) + +le-one-list-fundamental-theorem-arithmetic-ℕ : + (x : ℕ) (H : 1 ≤-ℕ x) → + for-all-list ℕ (le-ℕ-Prop 1) (list-fundamental-theorem-arithmetic-ℕ x H) +le-one-list-fundamental-theorem-arithmetic-ℕ x H = + map-for-all-list + ( is-prime-ℕ-Prop) + ( le-ℕ-Prop 1) + ( le-one-is-prime-ℕ) + ( list-fundamental-theorem-arithmetic-ℕ x H) + ( is-prime-list-fundamental-theorem-arithmetic-ℕ x H) -is-decomposition-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) → - is-decomposition-list-ℕ x (list-fundamental-theorem-arithmetic-ℕ x H) -is-decomposition-list-fundamental-theorem-arithmetic-ℕ x H = +is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ : + (x : ℕ) (H : 1 ≤-ℕ x) → + is-multiplicative-decomposition-list-ℕ x + ( list-fundamental-theorem-arithmetic-ℕ x H) +is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ = based-strong-ind-ℕ' 1 ( λ n N → - is-decomposition-list-ℕ n (list-fundamental-theorem-arithmetic-ℕ n N)) - ( refl) + is-multiplicative-decomposition-list-ℕ n + ( list-fundamental-theorem-arithmetic-ℕ n N)) + ( for-all-nil-list (le-ℕ-Prop 1) , refl) ( λ n N f → - tr - ( λ p → is-decomposition-list-ℕ (succ-ℕ n) p) - ( inv (compute-list-fundamental-theorem-arithmetic-succ-ℕ n N)) - ( ( ap - ( ( nat-least-prime-divisor-ℕ - ( succ-ℕ n) - ( le-succ-leq-ℕ 1 n N)) *ℕ_) + ( le-one-list-fundamental-theorem-arithmetic-ℕ (succ-ℕ n) star) , + ( ( ap + ( mul-list-ℕ) + ( compute-list-fundamental-theorem-arithmetic-succ-ℕ n N)) ∙ + ( ap + ( ( nat-least-prime-divisor-ℕ + ( succ-ℕ n) + ( le-succ-leq-ℕ 1 n N)) *ℕ_) + ( pr2 ( f ( quotient-div-least-prime-divisor-ℕ ( succ-ℕ n) @@ -559,32 +384,29 @@ is-decomposition-list-fundamental-theorem-arithmetic-ℕ x H = ( succ-ℕ n) ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) ( preserves-leq-succ-ℕ 1 n N)) - ( leq-quotient-div-least-prime-divisor-ℕ + ( upper-bound-quotient-div-least-prime-divisor-ℕ ( n) - ( le-succ-leq-ℕ 1 n N)))) ∙ - eq-quotient-div-ℕ' - ( nat-least-prime-divisor-ℕ - ( succ-ℕ n) - ( le-succ-leq-ℕ 1 n N)) + ( le-succ-leq-ℕ 1 n N))))) ∙ + ( eq-quotient-div-ℕ' + ( nat-least-prime-divisor-ℕ ( succ-ℕ n) - ( div-least-prime-divisor-ℕ - ( succ-ℕ n) - ( le-succ-leq-ℕ 1 n N)))) - ( x) - ( H) + ( le-succ-leq-ℕ 1 n N)) + ( succ-ℕ n) + ( div-least-prime-divisor-ℕ + ( succ-ℕ n) + ( le-succ-leq-ℕ 1 n N))))) is-list-of-nontrivial-divisors-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-list-of-nontrivial-divisors-ℕ x (list-fundamental-theorem-arithmetic-ℕ x H) is-list-of-nontrivial-divisors-fundamental-theorem-arithmetic-ℕ x H = - is-list-of-nontrivial-divisors-is-decomposition-is-prime-list-ℕ + is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ ( x) ( list-fundamental-theorem-arithmetic-ℕ x H) - ( is-decomposition-list-fundamental-theorem-arithmetic-ℕ x H) - ( is-prime-list-fundamental-theorem-arithmetic-ℕ x H) + ( is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ x H) is-least-element-list-least-prime-divisor-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (l : list ℕ) → + (x : ℕ) (H : 1 ≤-ℕ x) (l : list ℕ) → is-list-of-nontrivial-divisors-ℕ ( quotient-div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) ( l) → @@ -611,7 +433,7 @@ is-least-element-list-least-prime-divisor-ℕ x H (cons y l) D = is-least-element-list-least-prime-divisor-ℕ x H l (pr2 D) is-least-element-head-list-fundamental-theorem-arithmetic-succ-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → + (x : ℕ) → (H : 1 ≤-ℕ x) → is-least-element-list ( ℕ-Decidable-Total-Order) ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) @@ -642,7 +464,7 @@ is-least-element-head-list-fundamental-theorem-arithmetic-succ-ℕ x H = ( preserves-leq-succ-ℕ 1 x H))) is-sorted-least-element-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → + (x : ℕ) → (H : 1 ≤-ℕ x) → is-sorted-least-element-list ( ℕ-Decidable-Total-Order) ( list-fundamental-theorem-arithmetic-ℕ x H) @@ -670,12 +492,13 @@ is-sorted-least-element-list-fundamental-theorem-arithmetic-ℕ x H = ( succ-ℕ n) ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) ( preserves-leq-succ-ℕ 1 n N)) - ( leq-quotient-div-least-prime-divisor-ℕ n (le-succ-leq-ℕ 1 n N)))) + ( upper-bound-quotient-div-least-prime-divisor-ℕ n + ( le-succ-leq-ℕ 1 n N)))) ( x) ( H) is-sorted-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → + (x : ℕ) → (H : 1 ≤-ℕ x) → is-sorted-list ( ℕ-Decidable-Total-Order) ( list-fundamental-theorem-arithmetic-ℕ x H) @@ -686,84 +509,40 @@ is-sorted-list-fundamental-theorem-arithmetic-ℕ x H = ( is-sorted-least-element-list-fundamental-theorem-arithmetic-ℕ x H) is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-prime-decomposition-list-ℕ x (list-fundamental-theorem-arithmetic-ℕ x H) pr1 (is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ x H) = is-sorted-list-fundamental-theorem-arithmetic-ℕ x H pr1 (pr2 (is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ x H)) = is-prime-list-fundamental-theorem-arithmetic-ℕ x H pr2 (pr2 (is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ x H)) = - is-decomposition-list-fundamental-theorem-arithmetic-ℕ x H -``` - -### Uniqueness - -#### Definition of the type of prime decomposition of an integer - -```agda -prime-decomposition-list-ℕ : - (x : ℕ) → (leq-ℕ 1 x) → UU lzero -prime-decomposition-list-ℕ x _ = - Σ (list ℕ) (λ l → is-prime-decomposition-list-ℕ x l) -``` - -#### `prime-decomposition-list-ℕ n` is contractible for every `n` + is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ x H -```agda -prime-decomposition-fundamental-theorem-arithmetic-list-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) → prime-decomposition-list-ℕ x H -pr1 (prime-decomposition-fundamental-theorem-arithmetic-list-ℕ x H) = +prime-decomposition-fundamental-theorem-arithmetic-ℕ : + (x : ℕ) (H : 1 ≤-ℕ x) → prime-decomposition-list-ℕ x H +pr1 (prime-decomposition-fundamental-theorem-arithmetic-ℕ x H) = list-fundamental-theorem-arithmetic-ℕ x H -pr2 (prime-decomposition-fundamental-theorem-arithmetic-list-ℕ x H) = +pr2 (prime-decomposition-fundamental-theorem-arithmetic-ℕ x H) = is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ x H +``` -le-one-is-nonempty-prime-decomposition-list-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (y : ℕ) (l : list ℕ) → - is-prime-decomposition-list-ℕ x (cons y l) → - le-ℕ 1 x -le-one-is-nonempty-prime-decomposition-list-ℕ x H y l D = - concatenate-le-leq-ℕ - {x = 1} - {y = y} - {z = x} - ( le-one-is-prime-ℕ - ( y) - ( pr1 (is-prime-list-is-prime-decomposition-list-ℕ x (cons y l) D))) - ( leq-div-ℕ - ( y) - ( x) - ( is-nonzero-leq-one-ℕ x H) - ( mul-list-ℕ l , - ( commutative-mul-ℕ (mul-list-ℕ l) y ∙ - is-decomposition-list-is-prime-decomposition-list-ℕ x (cons y l) D))) +### Uniqueness + +#### The type of prime decompositions of any natural number is contractible -is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (l : list ℕ) → +```text +is-in-prime-decomposition-is-prime-divisor-ℕ : + (x : ℕ) (H : 1 ≤-ℕ x) (l : list ℕ) → is-prime-decomposition-list-ℕ x l → ( y : ℕ) → div-ℕ y x → is-prime-ℕ y → y ∈-list l -is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H nil D y d p = +is-in-prime-decomposition-is-prime-divisor-ℕ x H nil D y d p = ex-falso - ( contradiction-le-ℕ - ( 1) - ( x) - ( concatenate-le-leq-ℕ - {x = 1} - {y = y} - {z = x} - ( le-one-is-prime-ℕ y p) - ( leq-div-ℕ - ( y) - ( x) - ( is-nonzero-leq-one-ℕ x H) - ( d))) - ( leq-eq-ℕ - ( x) - ( 1) - ( inv (is-decomposition-list-is-prime-decomposition-list-ℕ x nil D)))) -is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H (cons z l) D y d p = + ( no-prime-divisors-one-ℕ y p + ( tr (div-ℕ y) (inv (eq-is-prime-decomposition-list-ℕ x nil D)) d)) +is-in-prime-decomposition-is-prime-divisor-ℕ x H (cons z l) D y d p = rec-coproduct ( λ e → tr (λ w → w ∈-list (cons z l)) (inv e) (is-head z l)) ( λ e → @@ -771,25 +550,26 @@ is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H (cons z l) D y d p ( y) ( z) ( l) - ( is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ + ( is-in-prime-decomposition-is-prime-divisor-ℕ ( quotient-div-ℕ ( z) ( x) - ( is-divisor-head-prime-decomposition-list-ℕ x z l D)) + ( div-head-is-prime-decomposition-cons-list-ℕ x z l D)) ( leq-one-quotient-div-ℕ ( z) ( x) - ( is-divisor-head-prime-decomposition-list-ℕ x z l D) + ( div-head-is-prime-decomposition-cons-list-ℕ x z l D) ( H)) ( l) ( ( is-sorted-tail-is-sorted-list ( ℕ-Decidable-Total-Order) ( cons z l) ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons z l) D)) , - ( pr2 - ( is-prime-list-is-prime-decomposition-list-ℕ x (cons z l) D)) , - ( refl)) + ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x z l D) , + ( {!!} , {!!}) {- ( refl) -}) ( y) + {!!} + {- ( div-right-factor-coprime-ℕ ( y) ( z) @@ -797,7 +577,7 @@ is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H (cons z l) D y d p ( tr ( λ x → div-ℕ y x) ( inv - ( is-decomposition-list-is-prime-decomposition-list-ℕ + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ ( x) ( cons z l) ( D))) @@ -806,14 +586,13 @@ is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H (cons z l) D y d p ( y) ( z) ( p) - ( pr1 - ( is-prime-list-is-prime-decomposition-list-ℕ x (cons z l) D)) - ( e))) + ( is-prime-head-is-prime-decomposition-cons-list-ℕ x z l D) + ( e))) -} ( p))) ( has-decidable-equality-ℕ y z) is-lower-bound-head-prime-decomposition-list-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (y : ℕ) (l : list ℕ) → + (x : ℕ) (H : 1 ≤-ℕ x) (y : ℕ) (l : list ℕ) → is-prime-decomposition-list-ℕ x (cons y l) → is-lower-bound-ℕ (is-nontrivial-divisor-ℕ x) y is-lower-bound-head-prime-decomposition-list-ℕ x H y l D m d = @@ -833,7 +612,7 @@ is-lower-bound-head-prime-decomposition-list-ℕ x H y l D m d = ( nat-least-prime-divisor-ℕ m (pr1 d)) ( l) ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons y l) D) - ( is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ + ( is-in-prime-decomposition-is-prime-divisor-ℕ ( x) ( H) ( cons y l) @@ -849,7 +628,7 @@ is-lower-bound-head-prime-decomposition-list-ℕ x H y l D m d = ( refl-leq-ℕ y)) eq-head-prime-decomposition-list-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (y z : ℕ) (p q : list ℕ) → + (x : ℕ) (H : 1 ≤-ℕ x) (y z : ℕ) (p q : list ℕ) → is-prime-decomposition-list-ℕ x (cons y p) → is-prime-decomposition-list-ℕ x (cons z q) → y = z @@ -859,14 +638,14 @@ eq-head-prime-decomposition-list-ℕ x H y z p q I J = ( all-elements-equal-minimal-element-ℕ ( is-nontrivial-divisor-ℕ-Prop x) ( y , - is-nontrivial-divisor-head-prime-decomposition-list-ℕ x y p I , + is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ x y p I , is-lower-bound-head-prime-decomposition-list-ℕ x H y p I) ( z , - is-nontrivial-divisor-head-prime-decomposition-list-ℕ x z q J , + is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ x z q J , is-lower-bound-head-prime-decomposition-list-ℕ x H z q J)) eq-prime-decomposition-list-ℕ : - (x : ℕ) (H : leq-ℕ 1 x) (p q : list ℕ) → + (x : ℕ) (H : 1 ≤-ℕ x) (p q : list ℕ) → is-prime-decomposition-list-ℕ x p → is-prime-decomposition-list-ℕ x q → p = q @@ -877,21 +656,21 @@ eq-prime-decomposition-list-ℕ x H (cons y l) nil I J = ( contradiction-le-ℕ ( 1) ( x) - ( le-one-is-nonempty-prime-decomposition-list-ℕ x H y l I) + ( le-one-is-prime-decomposition-cons-list-ℕ x y l I) ( leq-eq-ℕ ( x) ( 1) - ( inv ( is-decomposition-list-is-prime-decomposition-list-ℕ x nil J)))) + ( inv (eq-is-prime-decomposition-list-ℕ x nil J)))) eq-prime-decomposition-list-ℕ x H nil (cons y l) I J = ex-falso ( contradiction-le-ℕ ( 1) ( x) - ( le-one-is-nonempty-prime-decomposition-list-ℕ x H y l J) + ( le-one-is-prime-decomposition-cons-list-ℕ x y l J) ( leq-eq-ℕ ( x) ( 1) - ( inv (is-decomposition-list-is-prime-decomposition-list-ℕ x nil I)))) + ( inv (eq-is-prime-decomposition-list-ℕ x nil I)))) eq-prime-decomposition-list-ℕ x H (cons y l) (cons z p) I J = eq-Eq-list ( cons y l) @@ -904,11 +683,11 @@ eq-prime-decomposition-list-ℕ x H (cons y l) (cons z p) I J = ( quotient-div-ℕ ( y) ( x) - ( is-divisor-head-prime-decomposition-list-ℕ x y l I)) + ( div-head-is-prime-decomposition-cons-list-ℕ x y l I)) ( leq-one-quotient-div-ℕ ( y) ( x) - ( is-divisor-head-prime-decomposition-list-ℕ x y l I) + ( div-head-is-prime-decomposition-cons-list-ℕ x y l I) ( H)) ( l) ( p) @@ -916,35 +695,33 @@ eq-prime-decomposition-list-ℕ x H (cons y l) (cons z p) I J = ( ℕ-Decidable-Total-Order) ( cons y l) ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons y l) I)) , - pr2 (is-prime-list-is-prime-decomposition-list-ℕ x (cons y l) I) , - refl) + ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x y l I) , + {!!} {- refl -}) ( ( is-sorted-tail-is-sorted-list ( ℕ-Decidable-Total-Order) ( cons z p) ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons z p) J)) , - pr2 (is-prime-list-is-prime-decomposition-list-ℕ x (cons z p) J) , + ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x z p J) , + {!!} + {- tr - ( λ y → is-decomposition-list-ℕ y p) - ( eq-quotient-div-eq-div-ℕ + ( λ y → is-multiplicative-decomposition-list-ℕ y p) + ( eq-quotient-div-eq-divisor-ℕ ( z) ( y) ( x) ( is-nonzero-is-prime-ℕ ( z) - ( pr1 - ( is-prime-list-is-prime-decomposition-list-ℕ - ( x) - ( cons z p) - ( J)))) + ( is-prime-head-is-prime-decomposition-cons-list-ℕ x z p J))) ( inv (eq-head-prime-decomposition-list-ℕ x H y z l p I J)) - ( is-divisor-head-prime-decomposition-list-ℕ x z p J) - ( is-divisor-head-prime-decomposition-list-ℕ x y l I)) - ( refl))))) + ( div-head-is-prime-decomposition-cons-list-ℕ x z p J) + ( div-head-is-prime-decomposition-cons-list-ℕ x y l I)) + ( refl) -})))) fundamental-theorem-arithmetic-list-ℕ : - (x : ℕ) → (H : leq-ℕ 1 x) → is-contr (prime-decomposition-list-ℕ x H) + (x : ℕ) → (H : 1 ≤-ℕ x) → is-contr (prime-decomposition-list-ℕ x H) pr1 (fundamental-theorem-arithmetic-list-ℕ x H) = - prime-decomposition-fundamental-theorem-arithmetic-list-ℕ x H + prime-decomposition-fundamental-theorem-arithmetic-ℕ x H pr2 (fundamental-theorem-arithmetic-list-ℕ x H) d = eq-type-subtype ( is-prime-decomposition-list-ℕ-Prop x) @@ -959,7 +736,7 @@ pr2 (fundamental-theorem-arithmetic-list-ℕ x H) d = ### The sorted list associated with the concatenation of the prime decomposition of `n` and the prime decomposition of `m` is the prime decomposition of `n *ℕ m` -```agda +```text is-prime-list-concat-list-ℕ : (p q : list ℕ) → is-prime-list-ℕ p → is-prime-list-ℕ q → is-prime-list-ℕ (concat-list p q) @@ -1014,8 +791,9 @@ is-prime-list-permute-list-ℕ p t P = is-decomposition-list-concat-list-ℕ : (n m : ℕ) (p q : list ℕ) → - is-decomposition-list-ℕ n p → is-decomposition-list-ℕ m q → - is-decomposition-list-ℕ (n *ℕ m) (concat-list p q) + is-multiplicative-decomposition-list-ℕ n p → + is-multiplicative-decomposition-list-ℕ m q → + is-multiplicative-decomposition-list-ℕ (n *ℕ m) (concat-list p q) is-decomposition-list-concat-list-ℕ n m p q Dp Dq = ( eq-mul-list-concat-list-ℕ p q ∙ ( ap (mul-ℕ (mul-list-ℕ p)) Dq ∙ @@ -1023,13 +801,13 @@ is-decomposition-list-concat-list-ℕ n m p q Dp Dq = is-decomposition-list-permute-list-ℕ : (n : ℕ) (p : list ℕ) (t : Permutation (length-list p)) → - is-decomposition-list-ℕ n p → - is-decomposition-list-ℕ n (permute-list p t) + is-multiplicative-decomposition-list-ℕ n p → + is-multiplicative-decomposition-list-ℕ n (permute-list p t) is-decomposition-list-permute-list-ℕ n p t D = inv (invariant-permutation-mul-list-ℕ p t) ∙ D is-prime-decomposition-list-sort-concatenation-ℕ : - (x y : ℕ) (H : leq-ℕ 1 x) (I : leq-ℕ 1 y) (p q : list ℕ) → + (x y : ℕ) (H : 1 ≤-ℕ x) (I : 1 ≤-ℕ y) (p q : list ℕ) → is-prime-decomposition-list-ℕ x p → is-prime-decomposition-list-ℕ y q → is-prime-decomposition-list-ℕ @@ -1062,7 +840,7 @@ pr2 ( pr2 ( is-prime-decomposition-list-sort-concatenation-ℕ x y H I p q Dp Dq)) = tr - ( λ p → is-decomposition-list-ℕ (x *ℕ y) p) + ( λ p → is-multiplicative-decomposition-list-ℕ (x *ℕ y) p) ( inv ( eq-permute-list-permutation-insertion-sort-list ( ℕ-Decidable-Total-Order) @@ -1078,11 +856,11 @@ pr2 ( y) ( p) ( q) - ( is-decomposition-list-is-prime-decomposition-list-ℕ x p Dp) - ( is-decomposition-list-is-prime-decomposition-list-ℕ y q Dq))) + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ x p Dp) + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ y q Dq))) prime-decomposition-list-sort-concatenation-ℕ : - (x y : ℕ) (H : leq-ℕ 1 x) (I : leq-ℕ 1 y) (p q : list ℕ) → + (x y : ℕ) (H : 1 ≤-ℕ x) (I : 1 ≤-ℕ y) (p q : list ℕ) → is-prime-decomposition-list-ℕ x p → is-prime-decomposition-list-ℕ y q → prime-decomposition-list-ℕ (x *ℕ y) (preserves-leq-mul-ℕ 1 x 1 y H I) diff --git a/src/elementary-number-theory/inequality-natural-numbers.lagda.md b/src/elementary-number-theory/inequality-natural-numbers.lagda.md index 90bb0aa048..0391446218 100644 --- a/src/elementary-number-theory/inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/inequality-natural-numbers.lagda.md @@ -68,7 +68,7 @@ data leq-ℕ' : ℕ → ℕ → UU lzero where ```agda is-prop-leq-ℕ : - (m n : ℕ) → is-prop (leq-ℕ m n) + (m n : ℕ) → is-prop (m ≤-ℕ n) is-prop-leq-ℕ zero-ℕ zero-ℕ = is-prop-unit is-prop-leq-ℕ zero-ℕ (succ-ℕ n) = is-prop-unit is-prop-leq-ℕ (succ-ℕ m) zero-ℕ = is-prop-empty @@ -83,7 +83,7 @@ pr2 (leq-ℕ-Prop m n) = is-prop-leq-ℕ m n ```agda is-decidable-leq-ℕ : - (m n : ℕ) → is-decidable (leq-ℕ m n) + (m n : ℕ) → is-decidable (m ≤-ℕ n) is-decidable-leq-ℕ zero-ℕ zero-ℕ = inl star is-decidable-leq-ℕ zero-ℕ (succ-ℕ n) = inl star is-decidable-leq-ℕ (succ-ℕ m) zero-ℕ = inr id @@ -120,7 +120,7 @@ leq-eq-ℕ m .m refl = refl-leq-ℕ m ### Inequality on the natural numbers is transitive ```agda -transitive-leq-ℕ : is-transitive leq-ℕ +transitive-leq-ℕ : is-transitive _≤-ℕ_ transitive-leq-ℕ zero-ℕ m l p q = star transitive-leq-ℕ (succ-ℕ n) (succ-ℕ m) (succ-ℕ l) p q = transitive-leq-ℕ n m l p q @@ -173,22 +173,19 @@ linear-leq-ℕ (succ-ℕ m) (succ-ℕ n) = linear-leq-ℕ m n cases-order-three-elements-ℕ : (x y z : ℕ) → UU lzero cases-order-three-elements-ℕ x y z = - ( ( leq-ℕ x y × leq-ℕ y z) + - ( leq-ℕ x z × leq-ℕ z y)) + - ( ( ( leq-ℕ y z × leq-ℕ z x) + - ( leq-ℕ y x × leq-ℕ x z)) + - ( ( leq-ℕ z x × leq-ℕ x y) + - ( leq-ℕ z y × leq-ℕ y x))) + ( x ≤-ℕ y × y ≤-ℕ z + x ≤-ℕ z × z ≤-ℕ y) + + ( ( y ≤-ℕ z × z ≤-ℕ x + y ≤-ℕ x × x ≤-ℕ z) + + ( z ≤-ℕ x × x ≤-ℕ y + z ≤-ℕ y × y ≤-ℕ x)) order-three-elements-ℕ : (x y z : ℕ) → cases-order-three-elements-ℕ x y z -order-three-elements-ℕ zero-ℕ zero-ℕ zero-ℕ = inl (inl (pair star star)) -order-three-elements-ℕ zero-ℕ zero-ℕ (succ-ℕ z) = inl (inl (pair star star)) -order-three-elements-ℕ zero-ℕ (succ-ℕ y) zero-ℕ = inl (inr (pair star star)) +order-three-elements-ℕ zero-ℕ zero-ℕ zero-ℕ = inl (inl (star , star)) +order-three-elements-ℕ zero-ℕ zero-ℕ (succ-ℕ z) = inl (inl (star , star)) +order-three-elements-ℕ zero-ℕ (succ-ℕ y) zero-ℕ = inl (inr (star , star)) order-three-elements-ℕ zero-ℕ (succ-ℕ y) (succ-ℕ z) = inl (map-coproduct (pair star) (pair star) (linear-leq-ℕ y z)) order-three-elements-ℕ (succ-ℕ x) zero-ℕ zero-ℕ = - inr (inl (inl (pair star star))) + inr (inl (inl (star , star))) order-three-elements-ℕ (succ-ℕ x) zero-ℕ (succ-ℕ z) = inr (inl (map-coproduct (pair star) (pair star) (linear-leq-ℕ z x))) order-three-elements-ℕ (succ-ℕ x) (succ-ℕ y) zero-ℕ = @@ -201,19 +198,23 @@ order-three-elements-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) = ```agda leq-zero-ℕ : - (n : ℕ) → zero-ℕ ≤-ℕ n + (n : ℕ) → 0 ≤-ℕ n leq-zero-ℕ n = star + +leq-is-zero-ℕ : + (m n : ℕ) → is-zero-ℕ m → m ≤-ℕ n +leq-is-zero-ℕ .zero-ℕ n refl = leq-zero-ℕ n ``` ### Any natural number less than zero is zero ```agda is-zero-leq-zero-ℕ : - (x : ℕ) → x ≤-ℕ zero-ℕ → is-zero-ℕ x + (x : ℕ) → x ≤-ℕ 0 → is-zero-ℕ x is-zero-leq-zero-ℕ zero-ℕ star = refl is-zero-leq-zero-ℕ' : - (x : ℕ) → x ≤-ℕ zero-ℕ → is-zero-ℕ' x + (x : ℕ) → x ≤-ℕ 0 → is-zero-ℕ' x is-zero-leq-zero-ℕ' zero-ℕ star = refl ``` @@ -233,7 +234,7 @@ is-nonzero-leq-one-ℕ .zero-ℕ () refl ### Any natural number is less than or equal to its own successor ```agda -succ-leq-ℕ : (n : ℕ) → n ≤-ℕ (succ-ℕ n) +succ-leq-ℕ : (n : ℕ) → n ≤-ℕ succ-ℕ n succ-leq-ℕ zero-ℕ = star succ-leq-ℕ (succ-ℕ n) = succ-leq-ℕ n ``` @@ -242,7 +243,7 @@ succ-leq-ℕ (succ-ℕ n) = succ-leq-ℕ n ```agda decide-leq-succ-ℕ : - (m n : ℕ) → m ≤-ℕ (succ-ℕ n) → (m ≤-ℕ n) + (m = succ-ℕ n) + (m n : ℕ) → m ≤-ℕ succ-ℕ n → (m ≤-ℕ n) + (m = succ-ℕ n) decide-leq-succ-ℕ zero-ℕ zero-ℕ l = inl star decide-leq-succ-ℕ zero-ℕ (succ-ℕ n) l = inl star decide-leq-succ-ℕ (succ-ℕ m) zero-ℕ l = @@ -255,7 +256,7 @@ decide-leq-succ-ℕ (succ-ℕ m) (succ-ℕ n) l = ```agda preserves-leq-succ-ℕ : - (m n : ℕ) → m ≤-ℕ n → m ≤-ℕ (succ-ℕ n) + (m n : ℕ) → m ≤-ℕ n → m ≤-ℕ succ-ℕ n preserves-leq-succ-ℕ m n p = transitive-leq-ℕ m n (succ-ℕ n) (succ-leq-ℕ n) p ``` @@ -263,7 +264,7 @@ preserves-leq-succ-ℕ m n p = transitive-leq-ℕ m n (succ-ℕ n) (succ-leq-ℕ ```agda neg-succ-leq-ℕ : - (n : ℕ) → ¬ (leq-ℕ (succ-ℕ n) n) + (n : ℕ) → ¬ (succ-ℕ n ≤-ℕ n) neg-succ-leq-ℕ zero-ℕ = id neg-succ-leq-ℕ (succ-ℕ n) = neg-succ-leq-ℕ n ``` @@ -272,7 +273,7 @@ neg-succ-leq-ℕ (succ-ℕ n) = neg-succ-leq-ℕ n ```agda cases-leq-succ-ℕ : - {m n : ℕ} → leq-ℕ m (succ-ℕ n) → (leq-ℕ m n) + (m = succ-ℕ n) + {m n : ℕ} → m ≤-ℕ succ-ℕ n → (m ≤-ℕ n) + (m = succ-ℕ n) cases-leq-succ-ℕ {zero-ℕ} {n} star = inl star cases-leq-succ-ℕ {succ-ℕ m} {zero-ℕ} p = inr (ap succ-ℕ (antisymmetric-leq-ℕ m zero-ℕ p star)) @@ -289,10 +290,10 @@ cases-leq-succ-reflexive-leq-ℕ {succ-ℕ n} = ### `m ≤ n` if and only if `n + 1 ≰ m` ```agda -contradiction-leq-ℕ : (m n : ℕ) → m ≤-ℕ n → ¬ ((succ-ℕ n) ≤-ℕ m) +contradiction-leq-ℕ : (m n : ℕ) → m ≤-ℕ n → ¬ (succ-ℕ n ≤-ℕ m) contradiction-leq-ℕ (succ-ℕ m) (succ-ℕ n) H K = contradiction-leq-ℕ m n H K -contradiction-leq-ℕ' : (m n : ℕ) → (succ-ℕ n) ≤-ℕ m → ¬ (m ≤-ℕ n) +contradiction-leq-ℕ' : (m n : ℕ) → succ-ℕ n ≤-ℕ m → ¬ (m ≤-ℕ n) contradiction-leq-ℕ' m n K H = contradiction-leq-ℕ m n H K ``` @@ -300,11 +301,11 @@ contradiction-leq-ℕ' m n K H = contradiction-leq-ℕ m n H K ```agda preserves-leq-left-add-ℕ : - (k m n : ℕ) → m ≤-ℕ n → (m +ℕ k) ≤-ℕ (n +ℕ k) + (k m n : ℕ) → m ≤-ℕ n → m +ℕ k ≤-ℕ n +ℕ k preserves-leq-left-add-ℕ zero-ℕ m n = id preserves-leq-left-add-ℕ (succ-ℕ k) m n H = preserves-leq-left-add-ℕ k m n H -preserves-leq-right-add-ℕ : (k m n : ℕ) → m ≤-ℕ n → (k +ℕ m) ≤-ℕ (k +ℕ n) +preserves-leq-right-add-ℕ : (k m n : ℕ) → m ≤-ℕ n → k +ℕ m ≤-ℕ k +ℕ n preserves-leq-right-add-ℕ k m n H = concatenate-eq-leq-eq-ℕ ( commutative-add-ℕ k m) @@ -312,7 +313,7 @@ preserves-leq-right-add-ℕ k m n H = ( commutative-add-ℕ n k) preserves-leq-add-ℕ : - {m m' n n' : ℕ} → m ≤-ℕ m' → n ≤-ℕ n' → (m +ℕ n) ≤-ℕ (m' +ℕ n') + {m m' n n' : ℕ} → m ≤-ℕ m' → n ≤-ℕ n' → m +ℕ n ≤-ℕ m' +ℕ n' preserves-leq-add-ℕ {m} {m'} {n} {n'} H K = transitive-leq-ℕ ( m +ℕ n) @@ -326,12 +327,12 @@ preserves-leq-add-ℕ {m} {m'} {n} {n'} H K = ```agda reflects-leq-left-add-ℕ : - (k m n : ℕ) → (m +ℕ k) ≤-ℕ (n +ℕ k) → m ≤-ℕ n + (k m n : ℕ) → m +ℕ k ≤-ℕ n +ℕ k → m ≤-ℕ n reflects-leq-left-add-ℕ zero-ℕ m n = id reflects-leq-left-add-ℕ (succ-ℕ k) m n = reflects-leq-left-add-ℕ k m n reflects-leq-right-add-ℕ : - (k m n : ℕ) → (k +ℕ m) ≤-ℕ (k +ℕ n) → m ≤-ℕ n + (k m n : ℕ) → k +ℕ m ≤-ℕ k +ℕ n → m ≤-ℕ n reflects-leq-right-add-ℕ k m n H = reflects-leq-left-add-ℕ k m n ( concatenate-eq-leq-eq-ℕ @@ -343,7 +344,7 @@ reflects-leq-right-add-ℕ k m n H = ### `m ≤ m + n` for any two natural numbers `m` and `n` ```agda -leq-add-ℕ : (m n : ℕ) → m ≤-ℕ (m +ℕ n) +leq-add-ℕ : (m n : ℕ) → m ≤-ℕ m +ℕ n leq-add-ℕ m zero-ℕ = refl-leq-ℕ m leq-add-ℕ m (succ-ℕ n) = transitive-leq-ℕ @@ -353,17 +354,31 @@ leq-add-ℕ m (succ-ℕ n) = ( succ-leq-ℕ (m +ℕ n)) ( leq-add-ℕ m n) -leq-add-ℕ' : (m n : ℕ) → m ≤-ℕ (n +ℕ m) +leq-add-ℕ' : (m n : ℕ) → m ≤-ℕ n +ℕ m leq-add-ℕ' m n = concatenate-leq-eq-ℕ m (leq-add-ℕ m n) (commutative-add-ℕ m n) ``` +### If `m + n ≤ k`, then both `m ≤ k` and `n ≤ k` hold + +```agda +leq-right-summand-ℕ : + (m n k : ℕ) → m +ℕ n ≤-ℕ k → n ≤-ℕ k +leq-right-summand-ℕ m n k H = + transitive-leq-ℕ n (m +ℕ n) k H (leq-add-ℕ' n m) + +leq-left-summand-ℕ : + (m n k : ℕ) → m +ℕ n ≤-ℕ k → m ≤-ℕ k +leq-left-summand-ℕ m n k H = + transitive-leq-ℕ m (m +ℕ n) k H (leq-add-ℕ m n) +``` + ### We have `n ≤ m` if and only if there is a number `l` such that `l+n=m` ```agda subtraction-leq-ℕ : (n m : ℕ) → n ≤-ℕ m → Σ ℕ (λ l → l +ℕ n = m) -subtraction-leq-ℕ zero-ℕ m p = pair m refl -subtraction-leq-ℕ (succ-ℕ n) (succ-ℕ m) p = pair (pr1 P) (ap succ-ℕ (pr2 P)) +subtraction-leq-ℕ zero-ℕ m p = (m , refl) +subtraction-leq-ℕ (succ-ℕ n) (succ-ℕ m) p = (pr1 P , ap succ-ℕ (pr2 P)) where P : Σ ℕ (λ l' → l' +ℕ n = m) P = subtraction-leq-ℕ n m p @@ -378,7 +393,7 @@ leq-subtraction-ℕ (succ-ℕ n) (succ-ℕ m) l p = ```agda preserves-leq-left-mul-ℕ : - (k m n : ℕ) → m ≤-ℕ n → (m *ℕ k) ≤-ℕ (n *ℕ k) + (k m n : ℕ) → m ≤-ℕ n → m *ℕ k ≤-ℕ n *ℕ k preserves-leq-left-mul-ℕ k zero-ℕ n p = star preserves-leq-left-mul-ℕ k (succ-ℕ m) (succ-ℕ n) p = preserves-leq-left-add-ℕ k @@ -387,7 +402,7 @@ preserves-leq-left-mul-ℕ k (succ-ℕ m) (succ-ℕ n) p = ( preserves-leq-left-mul-ℕ k m n p) preserves-leq-right-mul-ℕ : - (k m n : ℕ) → m ≤-ℕ n → (k *ℕ m) ≤-ℕ (k *ℕ n) + (k m n : ℕ) → m ≤-ℕ n → k *ℕ m ≤-ℕ k *ℕ n preserves-leq-right-mul-ℕ k m n H = concatenate-eq-leq-eq-ℕ ( commutative-mul-ℕ k m) @@ -395,7 +410,7 @@ preserves-leq-right-mul-ℕ k m n H = ( commutative-mul-ℕ n k) preserves-leq-mul-ℕ : - (m m' n n' : ℕ) → m ≤-ℕ m' → n ≤-ℕ n' → (m *ℕ n) ≤-ℕ (m' *ℕ n') + (m m' n n' : ℕ) → m ≤-ℕ m' → n ≤-ℕ n' → m *ℕ n ≤-ℕ m' *ℕ n' preserves-leq-mul-ℕ m m' n n' H K = transitive-leq-ℕ ( m *ℕ n) @@ -408,54 +423,111 @@ preserves-leq-mul-ℕ m m' n n' H K = ### Multiplication by a nonzero natural number reflects inequality of natural numbers ```agda -reflects-order-mul-ℕ : - (k m n : ℕ) → (m *ℕ (succ-ℕ k)) ≤-ℕ (n *ℕ (succ-ℕ k)) → m ≤-ℕ n -reflects-order-mul-ℕ k zero-ℕ n p = star -reflects-order-mul-ℕ k (succ-ℕ m) (succ-ℕ n) p = - reflects-order-mul-ℕ k m n +reflects-order-mul-succ-ℕ : + (k m n : ℕ) → m *ℕ succ-ℕ k ≤-ℕ n *ℕ succ-ℕ k → m ≤-ℕ n +reflects-order-mul-succ-ℕ k zero-ℕ n p = star +reflects-order-mul-succ-ℕ k (succ-ℕ m) (succ-ℕ n) p = + reflects-order-mul-succ-ℕ k m n ( reflects-leq-left-add-ℕ ( succ-ℕ k) ( m *ℕ (succ-ℕ k)) ( n *ℕ (succ-ℕ k)) ( p)) -reflects-order-mul-ℕ' : - (k m n : ℕ) → ((succ-ℕ k) *ℕ m) ≤-ℕ ((succ-ℕ k) *ℕ n) → m ≤-ℕ n -reflects-order-mul-ℕ' k m n H = - reflects-order-mul-ℕ k m n +reflects-order-mul-succ-ℕ' : + (k m n : ℕ) → succ-ℕ k *ℕ m ≤-ℕ succ-ℕ k *ℕ n → m ≤-ℕ n +reflects-order-mul-succ-ℕ' k m n H = + reflects-order-mul-succ-ℕ k m n ( concatenate-eq-leq-eq-ℕ ( commutative-mul-ℕ m (succ-ℕ k)) ( H) ( commutative-mul-ℕ (succ-ℕ k) n)) + +reflects-order-mul-ℕ : + (k m n : ℕ) → is-nonzero-ℕ k → m *ℕ k ≤-ℕ n *ℕ k → m ≤-ℕ n +reflects-order-mul-ℕ k m n H with + is-successor-is-nonzero-ℕ H +... | (k' , refl) = reflects-order-mul-succ-ℕ k' m n + +reflects-order-mul-ℕ' : + (k m n : ℕ) → is-nonzero-ℕ k → k *ℕ m ≤-ℕ k *ℕ n → m ≤-ℕ n +reflects-order-mul-ℕ' k m n H K = + reflects-order-mul-ℕ k m n H + ( concatenate-eq-leq-eq-ℕ + ( commutative-mul-ℕ m k) + ( K) + ( commutative-mul-ℕ k n)) ``` ### Any number `x` is less than a nonzero multiple of itself ```agda leq-mul-ℕ : - (k x : ℕ) → x ≤-ℕ (x *ℕ (succ-ℕ k)) + (k x : ℕ) → x ≤-ℕ x *ℕ succ-ℕ k leq-mul-ℕ k x = concatenate-eq-leq-ℕ - ( x *ℕ (succ-ℕ k)) + ( x *ℕ succ-ℕ k) ( inv (right-unit-law-mul-ℕ x)) ( preserves-leq-right-mul-ℕ x 1 (succ-ℕ k) (leq-zero-ℕ k)) leq-mul-ℕ' : - (k x : ℕ) → x ≤-ℕ ((succ-ℕ k) *ℕ x) + (k x : ℕ) → x ≤-ℕ succ-ℕ k *ℕ x leq-mul-ℕ' k x = concatenate-leq-eq-ℕ x ( leq-mul-ℕ k x) ( commutative-mul-ℕ x (succ-ℕ k)) leq-mul-is-nonzero-ℕ : - (k x : ℕ) → is-nonzero-ℕ k → x ≤-ℕ (x *ℕ k) + (k x : ℕ) → is-nonzero-ℕ k → x ≤-ℕ x *ℕ k leq-mul-is-nonzero-ℕ k x H with is-successor-is-nonzero-ℕ H -... | pair l refl = leq-mul-ℕ l x +... | (l , refl) = leq-mul-ℕ l x leq-mul-is-nonzero-ℕ' : - (k x : ℕ) → is-nonzero-ℕ k → x ≤-ℕ (k *ℕ x) + (k x : ℕ) → is-nonzero-ℕ k → x ≤-ℕ k *ℕ x leq-mul-is-nonzero-ℕ' k x H with is-successor-is-nonzero-ℕ H -... | pair l refl = leq-mul-ℕ' l x +... | (l , refl) = leq-mul-ℕ' l x +``` + +### If `m * n ≤ k` and `m` is nonzero, then `n ≤ k` + +```agda +leq-right-factor-mul-succ-ℕ : + (m n k : ℕ) → succ-ℕ m *ℕ n ≤-ℕ k → n ≤-ℕ k +leq-right-factor-mul-succ-ℕ m n = + leq-right-summand-ℕ (m *ℕ n) n + +leq-right-factor-ℕ : + (m n k : ℕ) → is-nonzero-ℕ m → m *ℕ n ≤-ℕ k → n ≤-ℕ k +leq-right-factor-ℕ zero-ℕ n k H K = ex-falso (H refl) +leq-right-factor-ℕ (succ-ℕ m) n k H K = leq-right-factor-mul-succ-ℕ m n k K + +leq-right-factor-eq-ℕ : + (m n k : ℕ) → is-nonzero-ℕ m → m *ℕ n = k → n ≤-ℕ k +leq-right-factor-eq-ℕ m n .(m *ℕ n) H refl = + leq-mul-is-nonzero-ℕ' m n H +``` + +### If `m * n ≤ k` and `n` is nonzero, then `m ≤ k` + +```agda +leq-left-factor-mul-succ-ℕ : + (m n k : ℕ) → m *ℕ succ-ℕ n ≤-ℕ k → m ≤-ℕ k +leq-left-factor-mul-succ-ℕ m n k H = + leq-right-summand-ℕ + ( m *ℕ n) + ( m) + ( k) + ( concatenate-eq-leq-ℕ k (inv (right-successor-law-mul-ℕ m n)) H) + +leq-left-factor-ℕ : + (m n k : ℕ) → is-nonzero-ℕ n → m *ℕ n ≤-ℕ k → m ≤-ℕ k +leq-left-factor-ℕ m zero-ℕ k H K = ex-falso (H refl) +leq-left-factor-ℕ m (succ-ℕ n) k H K = leq-left-factor-mul-succ-ℕ m n k K + +leq-left-factor-eq-ℕ : + (m n k : ℕ) → is-nonzero-ℕ n → m *ℕ n = k → m ≤-ℕ k +leq-left-factor-eq-ℕ m n .(m *ℕ n) H refl = + leq-mul-is-nonzero-ℕ n m H ``` ## See also diff --git a/src/elementary-number-theory/jacobi-symbol.lagda.md b/src/elementary-number-theory/jacobi-symbol.lagda.md index b7c291a319..f4b447aa87 100644 --- a/src/elementary-number-theory/jacobi-symbol.lagda.md +++ b/src/elementary-number-theory/jacobi-symbol.lagda.md @@ -53,5 +53,7 @@ jacobi-symbol a (succ-ℕ (succ-ℕ n)) = ( mul-ℤ) ( map-list ( swap-Π legendre-symbol a) - ( list-primes-fundamental-theorem-arithmetic-ℕ (succ-ℕ (succ-ℕ n)) star)) + ( list-of-primes-fundamental-theorem-arithmetic-ℕ + ( succ-ℕ (succ-ℕ n)) + ( star))) ``` diff --git a/src/elementary-number-theory/lists-of-prime-numbers.lagda.md b/src/elementary-number-theory/lists-of-prime-numbers.lagda.md new file mode 100644 index 0000000000..9aabd05901 --- /dev/null +++ b/src/elementary-number-theory/lists-of-prime-numbers.lagda.md @@ -0,0 +1,54 @@ +# Lists of prime numbers + +```agda +module elementary-number-theory.lists-of-prime-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.prime-numbers + +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.unit-type +open import foundation.universe-levels + +open import lists.functoriality-lists +open import lists.lists +open import lists.predicates-on-lists +``` + +
+ +## Idea + +A {{#concept "list of prime numbers"}} is a [list](lists.lists.md) of [natural numbers](elementary-number-theory.natural-numbers.md) such that each number in the list is [prime](elementary-number-theory.prime-numbers.md). + +## Definitions + +### The predicate on lists of natural numbers of being a list of prime numbers + +```agda +is-prime-list-ℕ : + list ℕ → UU lzero +is-prime-list-ℕ = for-all-list ℕ is-prime-ℕ-Prop + +is-prop-is-prime-list-ℕ : + (l : list ℕ) → is-prop (is-prime-list-ℕ l) +is-prop-is-prime-list-ℕ = is-prop-for-all-list ℕ is-prime-ℕ-Prop +``` + +## Properties + +### Any list of prime numbers is a prime list + +```agda +is-prime-list-list-Prime-ℕ : + (l : list Prime-ℕ) → is-prime-list-ℕ (map-list nat-Prime-ℕ l) +is-prime-list-list-Prime-ℕ nil = raise-star +is-prime-list-list-Prime-ℕ (cons x l) = + ( is-prime-Prime-ℕ x , is-prime-list-list-Prime-ℕ l) +``` + diff --git a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md index 85b452a28c..c6589caff0 100644 --- a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md @@ -8,17 +8,20 @@ module elementary-number-theory.multiplication-lists-of-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import finite-group-theory.permutations-standard-finite-types open import foundation.action-on-identifications-functions +open import foundation.dependent-pair-types open import foundation.identity-types open import lists.concatenation-lists open import lists.lists open import lists.permutation-lists +open import lists.predicates-on-lists ```
@@ -38,6 +41,17 @@ mul-list-ℕ = fold-list 1 mul-ℕ ## Properties +### The product of any list of natural numbers greater than one is greater than one + +```agda +leq-one-mul-list-ℕ : + (l : list ℕ) (H : for-all-list ℕ (leq-ℕ-Prop 1) l) → + 1 ≤-ℕ mul-list-ℕ l +leq-one-mul-list-ℕ nil H = refl-leq-ℕ 1 +leq-one-mul-list-ℕ (cons n l) (H , K) = + preserves-leq-mul-ℕ 1 n 1 (mul-list-ℕ l) H (leq-one-mul-list-ℕ l K) +``` + ### `mul-list-ℕ` is invariant by permutation ```agda diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index d28741d96b..cb7bd10022 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -77,11 +77,21 @@ abstract (x : ℕ) → zero-ℕ *ℕ x = zero-ℕ left-zero-law-mul-ℕ x = refl + left-zero-law-mul-is-zero-ℕ : + (x y : ℕ) → is-zero-ℕ x → is-zero-ℕ (x *ℕ y) + left-zero-law-mul-is-zero-ℕ .zero-ℕ y refl = + left-zero-law-mul-ℕ y + right-zero-law-mul-ℕ : (x : ℕ) → x *ℕ zero-ℕ = zero-ℕ right-zero-law-mul-ℕ zero-ℕ = refl right-zero-law-mul-ℕ (succ-ℕ x) = ( right-unit-law-add-ℕ (x *ℕ zero-ℕ)) ∙ (right-zero-law-mul-ℕ x) + + right-zero-law-mul-is-zero-ℕ : + (x y : ℕ) → is-zero-ℕ y → is-zero-ℕ (x *ℕ y) + right-zero-law-mul-is-zero-ℕ x .zero-ℕ refl = + right-zero-law-mul-ℕ x ``` ### The unit laws for multiplication diff --git a/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md new file mode 100644 index 0000000000..8f10ecc72c --- /dev/null +++ b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md @@ -0,0 +1,162 @@ +# Multiplicative decompositions of natural numbers + +```agda +module + elementary-number-theory.multiplicative-decompositions-natural-numbers + where +``` + +
Imports + +```agda +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.multiplication-lists-of-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.nontrivial-divisors-natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers + +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.propositions +open import foundation.unit-type +open import foundation.universe-levels + +open import lists.lists +open import lists.predicates-on-lists +``` + +
+ +## Idea + +A {{#concept "multiplicative decomposition" Disambiguation="natural numbers" Agda=is-multiplicative-decomposition-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a [list](lists.lists.md) `l` of natural numbers [strictly greater than](elementary-number-theory.strict-inequality-natural-numbers.md) `1`, such that the [product](elementary-number-theory.multiplication-lists-of-natural-numbers.md) of the numbers in this list is [equal to](foundation-core.identity-types.md) `n`. + +## Definitions + +### The predicate on lists of natural numbers of being a product decomposition of a given number + +```agda +module _ + (n : ℕ) + (l : list ℕ) + where + + is-multiplicative-decomposition-list-ℕ : + UU lzero + is-multiplicative-decomposition-list-ℕ = + for-all-list ℕ (le-ℕ-Prop 1) l × + (mul-list-ℕ l = n) + + is-prop-is-multiplicative-decomposition-list-ℕ : + is-prop (is-multiplicative-decomposition-list-ℕ) + is-prop-is-multiplicative-decomposition-list-ℕ = + is-prop-product + ( is-prop-for-all-list ℕ (le-ℕ-Prop 1) l) + ( is-set-ℕ (mul-list-ℕ l) n) + + le-one-list-is-multiplicative-decomposition-list-ℕ : + is-multiplicative-decomposition-list-ℕ → + for-all-list ℕ (le-ℕ-Prop 1) l + le-one-list-is-multiplicative-decomposition-list-ℕ = pr1 + + eq-is-multiplicative-decomposition-list-ℕ : + is-multiplicative-decomposition-list-ℕ → + mul-list-ℕ l = n + eq-is-multiplicative-decomposition-list-ℕ = pr2 +``` + +## Properties + +### Any list of numbers strictly greater than 1 is a multiplicative decomposition of its product + +```agda +is-multiplicative-decomposition-mul-list-ℕ : + (l : list ℕ) → + for-all-list ℕ (le-ℕ-Prop 1) l → + is-multiplicative-decomposition-list-ℕ (mul-list-ℕ l) l +pr1 (is-multiplicative-decomposition-mul-list-ℕ l H) = H +pr2 (is-multiplicative-decomposition-mul-list-ℕ l H) = refl +``` + +### The head and tail of a multiplicative decomposition of a number are strictly greater than 1 + +```agda +module _ + (n x : ℕ) (l : list ℕ) + where + + le-one-head-is-multiplicative-decomposition-cons-list-ℕ : + is-multiplicative-decomposition-list-ℕ n (cons x l) → + 1 <-ℕ x + le-one-head-is-multiplicative-decomposition-cons-list-ℕ H = + pr1 (le-one-list-is-multiplicative-decomposition-list-ℕ n (cons x l) H) + + le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ : + is-multiplicative-decomposition-list-ℕ n (cons x l) → + for-all-list ℕ (le-ℕ-Prop 1) l + le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ H = + pr2 (le-one-list-is-multiplicative-decomposition-list-ℕ n (cons x l) H) +``` + +### If `l` is a multiplicative decomposition of a number `n`, then `l` is a list of nontrivial divisors of `l` + +```agda +div-head-is-multiplicative-decomposition-cons-list-ℕ : + (n x : ℕ) (l : list ℕ) → + is-multiplicative-decomposition-list-ℕ n (cons x l) → + div-ℕ x n +pr1 (div-head-is-multiplicative-decomposition-cons-list-ℕ n x l H) = + mul-list-ℕ l +pr2 (div-head-is-multiplicative-decomposition-cons-list-ℕ n x l H) = + commutative-mul-ℕ (mul-list-ℕ l) x ∙ + eq-is-multiplicative-decomposition-list-ℕ n (cons x l) H + +div-tail-is-multiplicative-decomposition-cons-list-ℕ : + (n x : ℕ) (l : list ℕ) → + is-multiplicative-decomposition-list-ℕ n (cons x l) → + div-ℕ (mul-list-ℕ l) n +pr1 (div-tail-is-multiplicative-decomposition-cons-list-ℕ n x l H) = + x +pr2 (div-tail-is-multiplicative-decomposition-cons-list-ℕ n x l H) = + eq-is-multiplicative-decomposition-list-ℕ n (cons x l) H + +is-nontrivial-divisor-head-is-multiplicative-decomposition-cons-list-ℕ : + (n x : ℕ) (l : list ℕ) → + is-multiplicative-decomposition-list-ℕ n (cons x l) → + is-nontrivial-divisor-ℕ n x +pr1 + ( is-nontrivial-divisor-head-is-multiplicative-decomposition-cons-list-ℕ + n x l H) = + le-one-head-is-multiplicative-decomposition-cons-list-ℕ n x l H +pr2 + ( is-nontrivial-divisor-head-is-multiplicative-decomposition-cons-list-ℕ + n x l H) = + div-head-is-multiplicative-decomposition-cons-list-ℕ n x l H + +is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ : + (n : ℕ) (l : list ℕ) → + is-multiplicative-decomposition-list-ℕ n l → + is-list-of-nontrivial-divisors-ℕ n l +is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ n nil H = + raise-star +is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ + ( n) + ( cons x l) + ( H) = + ( is-nontrivial-divisor-head-is-multiplicative-decomposition-cons-list-ℕ + n x l H , + is-list-of-nontrivial-divisors-div-ℕ + ( mul-list-ℕ l) + ( n) + ( div-tail-is-multiplicative-decomposition-cons-list-ℕ n x l H) + ( l) + ( is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ + ( mul-list-ℕ l) + ( l) + ( is-multiplicative-decomposition-mul-list-ℕ l + ( le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ n x l + ( H))))) +``` diff --git a/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md new file mode 100644 index 0000000000..b49688b451 --- /dev/null +++ b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md @@ -0,0 +1,249 @@ +# Nontrivial divisors of natural numbers + +```agda +module elementary-number-theory.nontrivial-divisors-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.prime-numbers +open import elementary-number-theory.strict-inequality-natural-numbers +open import elementary-number-theory.well-ordering-principle-natural-numbers + +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.propositions +open import foundation.type-arithmetic-empty-type +open import foundation.unit-type +open import foundation.universe-levels + +open import lists.lists +open import lists.predicates-on-lists +``` + +
+ +## Idea + +A {{#concept "nontrivial divisor"}} of a [natural number](elementary-number-theory.natural-numbers.md) is a [divisor](elementary-number-theory.divisibility-natural-numbers.md) [strictly greater](elementary-number-theory.strict-inequality-natural-numbers.md) than `1`. + +Every number strictly greater than `1` has a least nontrivial divisor, which is a [prime number](elementary-number-theory.prime-numbers.md). This fact is essential in the [Fundamental Theorem of Arithmetic](elementary-number-theory.fundamental-theorem-of-arithmetic.md). + +## Definitions + +### Nontrivial divisors + +Nontrivial divisors of a natural number are divisors strictly greater than `1`. + +```agda +module _ + (n x : ℕ) + where + + is-nontrivial-divisor-ℕ : UU lzero + is-nontrivial-divisor-ℕ = (1 <-ℕ x) × (div-ℕ x n) + + is-prop-is-nontrivial-divisor-ℕ : is-prop is-nontrivial-divisor-ℕ + is-prop-is-nontrivial-divisor-ℕ = + is-prop-Σ + ( is-prop-le-ℕ 1 x) + ( λ p → is-prop-div-ℕ x n (inl (is-nonzero-le-ℕ 1 x p))) + + is-nontrivial-divisor-ℕ-Prop : Prop lzero + pr1 is-nontrivial-divisor-ℕ-Prop = is-nontrivial-divisor-ℕ + pr2 is-nontrivial-divisor-ℕ-Prop = is-prop-is-nontrivial-divisor-ℕ + + le-one-is-nontrivial-divisor-ℕ : + is-nontrivial-divisor-ℕ → 1 <-ℕ x + le-one-is-nontrivial-divisor-ℕ = pr1 + + div-is-nontrivial-divisor-ℕ : + is-nontrivial-divisor-ℕ → div-ℕ x n + div-is-nontrivial-divisor-ℕ = pr2 +``` + +### Lists of nontrivial divisors of a number + +```agda +is-list-of-nontrivial-divisors-ℕ : + ℕ → list ℕ → UU lzero +is-list-of-nontrivial-divisors-ℕ n = + for-all-list ℕ (is-nontrivial-divisor-ℕ-Prop n) +``` + +## Properties + +### Being a nontrivial divisor is decidable + +```agda +is-decidable-is-nontrivial-divisor-ℕ : + (n x : ℕ) → is-decidable (is-nontrivial-divisor-ℕ n x) +is-decidable-is-nontrivial-divisor-ℕ n x = + is-decidable-product (is-decidable-le-ℕ 1 x) (is-decidable-div-ℕ x n) +``` + +### Any number strictly greater than `1` is a nontrivial divisor of itself + +```agda +is-nontrivial-divisor-diagonal-ℕ : + (n : ℕ) → le-ℕ 1 n → is-nontrivial-divisor-ℕ n n +pr1 (is-nontrivial-divisor-diagonal-ℕ n H) = H +pr2 (is-nontrivial-divisor-diagonal-ℕ n H) = refl-div-ℕ n +``` + +### Any nontrivial divisor of a number `x` that divides a number `y` is a nontrivial divisor of `y` + +```agda +is-nontrivial-divisor-div-ℕ : + (m n k : ℕ) → div-ℕ n k → + is-nontrivial-divisor-ℕ n m → is-nontrivial-divisor-ℕ k m +pr1 (is-nontrivial-divisor-div-ℕ m n k H K) = + le-one-is-nontrivial-divisor-ℕ n m K +pr2 (is-nontrivial-divisor-div-ℕ m n k H K) = + transitive-div-ℕ m n k H (div-is-nontrivial-divisor-ℕ n m K) +``` + +### Any list of nontrivial divisors of a number `x` that divides a number `y` is a list of nontrivial divisors of `y` + +```agda +is-list-of-nontrivial-divisors-div-ℕ : + (x y : ℕ) → div-ℕ x y → (l : list ℕ) → + is-list-of-nontrivial-divisors-ℕ x l → is-list-of-nontrivial-divisors-ℕ y l +is-list-of-nontrivial-divisors-div-ℕ x y d nil H = raise-star +pr1 (is-list-of-nontrivial-divisors-div-ℕ x y d (cons z l) (H , K)) = + is-nontrivial-divisor-div-ℕ z x y d H +pr2 (is-list-of-nontrivial-divisors-div-ℕ x y d (cons z l) (H , K)) = + is-list-of-nontrivial-divisors-div-ℕ x y d l K +``` + +### Every natural number strictly greater than `1` has a least nontrivial divisor + +```agda +least-nontrivial-divisor-ℕ : + (n : ℕ) → le-ℕ 1 n → minimal-element-ℕ (is-nontrivial-divisor-ℕ n) +least-nontrivial-divisor-ℕ n H = + well-ordering-principle-ℕ + ( is-nontrivial-divisor-ℕ n) + ( is-decidable-is-nontrivial-divisor-ℕ n) + ( n , is-nontrivial-divisor-diagonal-ℕ n H) + +nat-least-nontrivial-divisor-ℕ : (n : ℕ) → le-ℕ 1 n → ℕ +nat-least-nontrivial-divisor-ℕ n H = pr1 (least-nontrivial-divisor-ℕ n H) + +nat-least-nontrivial-divisor-ℕ' : ℕ → ℕ +nat-least-nontrivial-divisor-ℕ' zero-ℕ = 0 +nat-least-nontrivial-divisor-ℕ' (succ-ℕ zero-ℕ) = 1 +nat-least-nontrivial-divisor-ℕ' (succ-ℕ (succ-ℕ n)) = + nat-least-nontrivial-divisor-ℕ (succ-ℕ (succ-ℕ n)) star + +le-one-least-nontrivial-divisor-ℕ : + (n : ℕ) (H : le-ℕ 1 n) → le-ℕ 1 (nat-least-nontrivial-divisor-ℕ n H) +le-one-least-nontrivial-divisor-ℕ n H = + pr1 (pr1 (pr2 (least-nontrivial-divisor-ℕ n H))) + +div-least-nontrivial-divisor-ℕ : + (n : ℕ) (H : le-ℕ 1 n) → div-ℕ (nat-least-nontrivial-divisor-ℕ n H) n +div-least-nontrivial-divisor-ℕ n H = + pr2 (pr1 (pr2 (least-nontrivial-divisor-ℕ n H))) + +is-minimal-least-nontrivial-divisor-ℕ : + (n : ℕ) (H : le-ℕ 1 n) (x : ℕ) (K : 1 <-ℕ x) (d : div-ℕ x n) → + leq-ℕ (nat-least-nontrivial-divisor-ℕ n H) x +is-minimal-least-nontrivial-divisor-ℕ n H x K d = + pr2 (pr2 (least-nontrivial-divisor-ℕ n H)) x (K , d) +``` + +### The least nontrivial divisor of a number `> 1` is nonzero + +```agda +abstract + is-nonzero-least-nontrivial-divisor-ℕ : + (n : ℕ) (H : le-ℕ 1 n) → is-nonzero-ℕ (nat-least-nontrivial-divisor-ℕ n H) + is-nonzero-least-nontrivial-divisor-ℕ n H = + is-nonzero-div-ℕ + ( nat-least-nontrivial-divisor-ℕ n H) + ( n) + ( div-least-nontrivial-divisor-ℕ n H) + ( λ where refl → H) +``` + +### The least nontrivial divisor of a number `> 1` is prime + +```agda +is-prime-least-nontrivial-divisor-ℕ : + (n : ℕ) (H : le-ℕ 1 n) → is-prime-ℕ (nat-least-nontrivial-divisor-ℕ n H) +pr1 (is-prime-least-nontrivial-divisor-ℕ n H x) (K , L) = + map-right-unit-law-coproduct-is-empty + ( is-one-ℕ x) + ( 1 <-ℕ x) + ( λ p → + contradiction-le-ℕ x + ( nat-least-nontrivial-divisor-ℕ n H) + ( le-div-ℕ x + ( nat-least-nontrivial-divisor-ℕ n H) + ( is-nonzero-least-nontrivial-divisor-ℕ n H) + ( L) + ( K)) + ( is-minimal-least-nontrivial-divisor-ℕ n H x p + ( transitive-div-ℕ x + ( nat-least-nontrivial-divisor-ℕ n H) + ( n) + ( div-least-nontrivial-divisor-ℕ n H) + ( L)))) + ( eq-or-le-leq-ℕ' 1 x + ( leq-one-div-ℕ x n + ( transitive-div-ℕ x + ( nat-least-nontrivial-divisor-ℕ n H) + ( n) + ( div-least-nontrivial-divisor-ℕ n H) + ( L)) + ( leq-le-ℕ 1 n H))) +pr1 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = + neq-le-ℕ (le-one-least-nontrivial-divisor-ℕ n H) +pr2 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = + div-one-ℕ _ +``` + +### The least prime divisor of a number `1 < n` + +```agda +nat-least-prime-divisor-ℕ : (x : ℕ) → 1 <-ℕ x → ℕ +nat-least-prime-divisor-ℕ x H = nat-least-nontrivial-divisor-ℕ x H + +is-prime-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ x) → is-prime-ℕ (nat-least-prime-divisor-ℕ x H) +is-prime-least-prime-divisor-ℕ x H = is-prime-least-nontrivial-divisor-ℕ x H + +prime-least-prime-divisor-ℕ : (x : ℕ) → 1 <-ℕ x → Prime-ℕ +pr1 (prime-least-prime-divisor-ℕ x H) = nat-least-prime-divisor-ℕ x H +pr2 (prime-least-prime-divisor-ℕ x H) = is-prime-least-prime-divisor-ℕ x H + +div-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ x) → div-ℕ (nat-least-prime-divisor-ℕ x H) x +div-least-prime-divisor-ℕ x H = div-least-nontrivial-divisor-ℕ x H + +quotient-div-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ x) → ℕ +quotient-div-least-prime-divisor-ℕ x H = + quotient-div-ℕ + ( nat-least-prime-divisor-ℕ x H) + ( x) + ( div-least-prime-divisor-ℕ x H) + +upper-bound-quotient-div-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ succ-ℕ x) → + quotient-div-least-prime-divisor-ℕ (succ-ℕ x) H ≤-ℕ x +upper-bound-quotient-div-least-prime-divisor-ℕ x H = + upper-bound-quotient-div-is-prime-ℕ + ( nat-least-prime-divisor-ℕ (succ-ℕ x) H) + ( x) + ( is-prime-least-prime-divisor-ℕ (succ-ℕ x) H) + ( div-least-prime-divisor-ℕ (succ-ℕ x) H) +``` diff --git a/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md b/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md index 4977a5e18b..e0fd9653a5 100644 --- a/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md +++ b/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md @@ -9,6 +9,7 @@ module
Imports ```agda +open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers @@ -32,50 +33,29 @@ open import order-theory.preorders The **poset of natural numbers ordered by divisibility** consists of the [natural numbers](elementary-number-theory.natural-numbers.md) and its ordering -is defined by `m ≤ n := m | n`, i.e., by -[divisibility](elementary-number-theory.divisibility-natural-numbers.md). - -The divisibility relation `m | n` on the natural numbers, however, is only -valued in the [propositions](foundation.propositions.md) when both `m` and `n` -are [nonzero](elementary-number-theory.nonzero-natural-numbers.md). We therefore -redefine the divisibility relation in the following way: A number `m` is said to -**divide** a number `n` if there -[merely exists](foundation.existential-quantification.md) a number `k` such that -`km = n`. Since mere existence is defined via the -[propoositional truncation](foundation.propositional-truncations.md), this can -be stated alternatively as the proposition - -```text - trunc-Prop (div-ℕ m n). -``` +is defined by [bounded divisibility](elementary-number-theory.bounded-divisibility-natural-numbers.md), i.e., the type `m ≤ n` is defined to be the type of natural numbers `q ≤ n` such that `q * m = n`. -In other words, we simply force the divisibility relation to take values in -propositions by identifying all witnesses of divisibility. +Recall that bounded divisibility is [logically equivalent](foundation.logical-equivalences.md) to the more standard [divisibility relation](elementar-number-theory.divisibility-natural-numbers.md). However, the divisibility relation `m | n` is only valued in the [propositions](foundation.propositions.md) when both `m` and `n` +are [nonzero](elementary-number-theory.nonzero-natural-numbers.md). On the other hand, bounded divisibility is always valued in propositions. By using bounded divisibility we avoid the need for [propoositional truncation](foundation.propositional-truncations.md). ## Definition ```agda leq-prop-ℕ-Div : ℕ → ℕ → Prop lzero -leq-prop-ℕ-Div m n = trunc-Prop (div-ℕ m n) +leq-prop-ℕ-Div = bounded-div-ℕ-Prop leq-ℕ-Div : ℕ → ℕ → UU lzero leq-ℕ-Div m n = type-Prop (leq-prop-ℕ-Div m n) refl-leq-ℕ-Div : (n : ℕ) → leq-ℕ-Div n n -refl-leq-ℕ-Div n = unit-trunc-Prop (refl-div-ℕ n) +refl-leq-ℕ-Div = refl-bounded-div-ℕ antisymmetric-leq-ℕ-Div : (m n : ℕ) → leq-ℕ-Div m n → leq-ℕ-Div n m → m = n -antisymmetric-leq-ℕ-Div m n H K = - apply-twice-universal-property-trunc-Prop H K - ( Id-Prop ℕ-Set _ _) - ( antisymmetric-div-ℕ m n) +antisymmetric-leq-ℕ-Div = antisymmetric-bounded-div-ℕ transitive-leq-ℕ-Div : (m n o : ℕ) → leq-ℕ-Div n o → leq-ℕ-Div m n → leq-ℕ-Div m o -transitive-leq-ℕ-Div m n o H K = - apply-twice-universal-property-trunc-Prop H K - ( leq-prop-ℕ-Div m o) - ( λ H' K' → unit-trunc-Prop (transitive-div-ℕ m n o H' K')) +transitive-leq-ℕ-Div = transitive-bounded-div-ℕ ℕ-Div-Preorder : Preorder lzero lzero pr1 ℕ-Div-Preorder = ℕ diff --git a/src/elementary-number-theory/powers-of-two.lagda.md b/src/elementary-number-theory/powers-of-two.lagda.md index 1fa5eaf150..4fed8c7c76 100644 --- a/src/elementary-number-theory/powers-of-two.lagda.md +++ b/src/elementary-number-theory/powers-of-two.lagda.md @@ -74,33 +74,47 @@ abstract ap succ-ℕ (commutative-mul-ℕ (pr1 s) 2) ∙ pr2 s)))) ( inr x) → - ( let e : is-even-ℕ k - e = is-even-is-odd-succ-ℕ k x - - t : (pr1 e) ≤-ℕ k - t = leq-quotient-div-ℕ' 2 k is-nonzero-two-ℕ e - - s : (pair-expansion (pr1 e)) - s = f (pr1 e) t (is-decidable-is-even-ℕ (pr1 e)) + ( let + e : is-even-ℕ k + e = is-even-is-odd-succ-ℕ k x + + t : quotient-div-ℕ 2 k e ≤-ℕ k + t = upper-bound-quotient-div-ℕ 2 k e + + s : (pair-expansion (quotient-div-ℕ 2 k e)) + s = + f (quotient-div-ℕ 2 k e) + ( t) + ( is-decidable-is-even-ℕ (quotient-div-ℕ 2 k e)) in - pair - ( succ-ℕ (pr1 (pr1 s)) , pr2 (pr1 s)) - ( ( ap - ( _*ℕ (succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) - ( commutative-mul-ℕ (exp-ℕ 2 (pr1 (pr1 s))) 2)) ∙ - ( ( associative-mul-ℕ 2 - ( exp-ℕ 2 (pr1 (pr1 s))) - ( succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) ∙ - ( ( ap (2 *ℕ_) (pr2 s)) ∙ - ( ( ap succ-ℕ - ( left-successor-law-add-ℕ (0 +ℕ (pr1 e)) (pr1 e))) ∙ + pair + ( succ-ℕ (pr1 (pr1 s)) , pr2 (pr1 s)) + ( ( ap + ( _*ℕ (succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) + ( commutative-mul-ℕ (exp-ℕ 2 (pr1 (pr1 s))) 2)) ∙ + ( ( associative-mul-ℕ + ( 2) + ( exp-ℕ 2 (pr1 (pr1 s))) + ( succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) ∙ + ( ( ap (2 *ℕ_) (pr2 s)) ∙ + ( ( ap + ( succ-ℕ) + ( ( left-successor-law-add-ℕ + ( 0 +ℕ quotient-div-ℕ 2 k e) + ( quotient-div-ℕ 2 k e)))) ∙ + ( ( ap + ( succ-ℕ ∘ succ-ℕ) + ( ap + ( _+ℕ quotient-div-ℕ 2 k e) + ( left-unit-law-add-ℕ (quotient-div-ℕ 2 k e)))) ∙ ( ( ap ( succ-ℕ ∘ succ-ℕ) - ( ap (_+ℕ (pr1 e)) (left-unit-law-add-ℕ (pr1 e)))) ∙ - ( ( ap - ( succ-ℕ ∘ succ-ℕ) - ( inv (right-two-law-mul-ℕ (pr1 e)))) ∙ - ( ( ap (succ-ℕ ∘ succ-ℕ) (pr2 e)))))))))))) + ( inv + ( right-two-law-mul-ℕ + ( quotient-div-ℕ 2 k e)))) ∙ + ( ap + ( succ-ℕ ∘ succ-ℕ) + ( eq-quotient-div-ℕ 2 k e))))))))))) ( n) has-pair-expansion : (n : ℕ) → pair-expansion n diff --git a/src/elementary-number-theory/prime-numbers.lagda.md b/src/elementary-number-theory/prime-numbers.lagda.md index 930d240933..b969562206 100644 --- a/src/elementary-number-theory/prime-numbers.lagda.md +++ b/src/elementary-number-theory/prime-numbers.lagda.md @@ -27,6 +27,7 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.identity-types open import foundation.logical-equivalences open import foundation.negated-equality +open import foundation.negation open import foundation.propositions open import foundation.torsorial-type-families open import foundation.transport-along-identifications @@ -224,24 +225,39 @@ is-prime-two-ℕ = ### If a prime number `p` divides a nonzero number `x`, then `x/p < x` ```agda -le-quotient-div-is-prime-ℕ : +strict-upper-bound-quotient-div-is-prime-ℕ : (p x : ℕ) → is-nonzero-ℕ x → is-prime-ℕ p → - (H : div-ℕ p x) → le-ℕ (quotient-div-ℕ p x H) x -le-quotient-div-is-prime-ℕ p x N P H = + (H : div-ℕ p x) → quotient-div-ℕ p x H <-ℕ x +strict-upper-bound-quotient-div-is-prime-ℕ p x N P H = le-quotient-div-ℕ p x N H (is-not-one-is-prime-ℕ p P) ``` ### If a prime number `p` divides a number `x + 1`, then `(x + 1)/p ≤ x` +Note that this upper bound is slightly sharper than the usual upper bound `x + 1` we get for arbitrary quotients of divisible natural numbers. + ```agda -leq-quotient-div-is-prime-ℕ : +upper-bound-quotient-div-is-prime-ℕ : (p x : ℕ) → is-prime-ℕ p → - (H : div-ℕ p (succ-ℕ x)) → leq-ℕ (quotient-div-ℕ p (succ-ℕ x) H) x -leq-quotient-div-is-prime-ℕ p x P H = + (H : div-ℕ p (succ-ℕ x)) → quotient-div-ℕ p (succ-ℕ x) H ≤-ℕ x +upper-bound-quotient-div-is-prime-ℕ p x P H = leq-le-succ-ℕ ( quotient-div-ℕ p (succ-ℕ x) H) ( x) - ( le-quotient-div-is-prime-ℕ p (succ-ℕ x) (is-nonzero-succ-ℕ x) P H) + ( strict-upper-bound-quotient-div-is-prime-ℕ p + ( succ-ℕ x) + ( is-nonzero-succ-ℕ x) + ( P) + ( H)) +``` + +### The number `1` has no prime divisors + +```agda +no-prime-divisors-one-ℕ : + (p : ℕ) → is-prime-ℕ p → ¬ (div-ℕ p 1) +no-prime-divisors-one-ℕ p H K = + is-not-one-is-prime-ℕ p H (is-one-div-one-ℕ p K) ``` ## See also diff --git a/src/elementary-number-theory/triangular-numbers.lagda.md b/src/elementary-number-theory/triangular-numbers.lagda.md index b92b578c2f..41c27c4114 100644 --- a/src/elementary-number-theory/triangular-numbers.lagda.md +++ b/src/elementary-number-theory/triangular-numbers.lagda.md @@ -75,7 +75,7 @@ $$ We prove this claim to avoid an early division by two, and the proof is by induction on $n$. In the base case both sides of the equality are $0$. For the inductive step, assume that $2 \cdot \sum_{i\leq n} i = n(n+1)$. Then we can compute $$ -2 \cdot \sum_{i\leq n+1} i = 2 \cdot \left(\sum_{i\leq n} i\right)+ 2(n+1) = n(n+1) + 2(n+1) = (n+1)(n+2). +2 \cdot \sum_{i\leq n+1} i = 2 \cdot \left(\sum_{i\leq n} i\right)+ 2(n+1) = n(n+1) + ((n+1) + (n+1)) = (n+1)(n+2). $$ ```agda diff --git a/src/lists/predicates-on-lists.lagda.md b/src/lists/predicates-on-lists.lagda.md index 4d52edb669..56e2eeb219 100644 --- a/src/lists/predicates-on-lists.lagda.md +++ b/src/lists/predicates-on-lists.lagda.md @@ -7,6 +7,7 @@ module lists.predicates-on-lists where
Imports ```agda +open import foundation.dependent-pair-types open import foundation.propositions open import foundation.unit-type open import foundation.universe-levels @@ -18,7 +19,7 @@ open import lists.lists ## Definitions -### For all +### Universal quantification over the elements of a list ```agda module _ @@ -38,3 +39,26 @@ module _ (l : list X) → is-prop (for-all-list l) is-prop-for-all-list l = is-prop-type-Prop (for-all-list-Prop l) ``` + +## Properties + +### Universal quantification over the elements of the empty list + +```agda +for-all-nil-list : + {l1 l2 : Level} {X : UU l1} (P : X → Prop l2) → + for-all-list X P nil +for-all-nil-list P = raise-star +``` + +### If a predicate implies another predicate, then its universal quantification over the elements of a list implies the universal-quantification of the other predicate + +```agda +map-for-all-list : + {l1 l2 l3 : Level} {X : UU l1} (P : X → Prop l2) (Q : X → Prop l3) → + ((x : X) → type-Prop (P x) → type-Prop (Q x)) → + (l : list X) → for-all-list X P l → for-all-list X Q l +map-for-all-list P Q f nil H = raise-star +pr1 (map-for-all-list P Q f (cons x l) (H , K)) = f _ H +pr2 (map-for-all-list P Q f (cons x l) (H , K)) = map-for-all-list P Q f l K +``` From 5578aaacec8afa9b85b4c0972507ffb14176cf2f Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 17 Dec 2024 12:22:09 -0500 Subject: [PATCH 24/72] further refactoring of fundamental theorem --- .../divisibility-natural-numbers.lagda.md | 5 + ...fundamental-theorem-of-arithmetic.lagda.md | 316 +++++++++--------- .../lists-of-prime-numbers.lagda.md | 11 +- ...lication-lists-of-natural-numbers.lagda.md | 5 +- ...ve-decompositions-natural-numbers.lagda.md | 45 ++- ...ntrivial-divisors-natural-numbers.lagda.md | 40 ++- src/lists.lagda.md | 2 +- src/lists/predicates-on-lists.lagda.md | 64 ---- src/lists/sorted-lists.lagda.md | 30 +- .../universal-quantification-lists.lagda.md | 141 ++++++++ 10 files changed, 405 insertions(+), 254 deletions(-) delete mode 100644 src/lists/predicates-on-lists.lagda.md create mode 100644 src/lists/universal-quantification-lists.lagda.md diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 113973cf82..913a825491 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -540,6 +540,11 @@ leq-one-quotient-div-ℕ d x H K = ( x) ( div-quotient-div-ℕ d x H) ( K) + +leq-one-quotient-div-is-nonzero-ℕ : + (d x : ℕ) (H : div-ℕ d x) → is-nonzero-ℕ x → leq-ℕ 1 (quotient-div-ℕ d x H) +leq-one-quotient-div-is-nonzero-ℕ d x H N = + leq-one-quotient-div-ℕ d x H (leq-one-is-nonzero-ℕ x N) ``` ### `a/a = 1` diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index f01148e3af..762a082354 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -47,7 +47,7 @@ open import lists.concatenation-lists open import lists.functoriality-lists open import lists.lists open import lists.permutation-lists -open import lists.predicates-on-lists +open import lists.universal-quantification-lists open import lists.sort-by-insertion-lists open import lists.sorted-lists ``` @@ -76,7 +76,8 @@ in several ways: [equal](foundation-core.identity-types.md) to `n` is contractible. Note that the [univalence axiom](foundation-core.univalence.md) is neccessary to -prove the second uniqueness property of prime factorizations. On this page we will prove the fundamental theorem using the first approach. +prove the second uniqueness property of prime factorizations. On this page we +will prove the fundamental theorem using the first approach. The fundamental theorem of arithmetic is the 80th theorem on [Freek Wiedijk's](http://www.cs.ru.nl/F.Wiedijk/) list of @@ -90,7 +91,9 @@ A list of natural numbers is a prime decomposition of a natural number `n` if: - The list is sorted. - Every element of the list is prime. -- The list is a [multiplicative decomposition](elementary-number-theory.multiplicative-decompositions-natural-numbers.md) of `n`. +- The list is a + [multiplicative decomposition](elementary-number-theory.multiplicative-decompositions-natural-numbers.md) + of `n`. ```agda module _ @@ -105,10 +108,10 @@ module _ ( is-prime-list-ℕ l × is-multiplicative-decomposition-list-ℕ x l) - is-sorted-list-is-prime-decomposition-list-ℕ : + is-sorted-is-prime-decomposition-list-ℕ : is-prime-decomposition-list-ℕ → is-sorted-list ℕ-Decidable-Total-Order l - is-sorted-list-is-prime-decomposition-list-ℕ D = + is-sorted-is-prime-decomposition-list-ℕ D = pr1 D is-prime-list-is-prime-decomposition-list-ℕ : @@ -124,24 +127,18 @@ module _ le-one-list-is-prime-decomposition-list-ℕ : is-prime-decomposition-list-ℕ → - for-all-list ℕ (le-ℕ-Prop 1) l + for-all-list l (le-ℕ 1) le-one-list-is-prime-decomposition-list-ℕ H = - map-for-all-list - ( is-prime-ℕ-Prop) - ( le-ℕ-Prop 1) + map-for-all-list l ( le-one-is-prime-ℕ) - ( l) ( is-prime-list-is-prime-decomposition-list-ℕ H) leq-one-list-is-prime-decomposition-list-ℕ : is-prime-decomposition-list-ℕ → - for-all-list ℕ (leq-ℕ-Prop 1) l + for-all-list l (leq-ℕ 1) leq-one-list-is-prime-decomposition-list-ℕ H = - map-for-all-list - ( le-ℕ-Prop 1) - ( leq-ℕ-Prop 1) + map-for-all-list l ( leq-le-ℕ 1) - ( l) ( le-one-list-is-prime-decomposition-list-ℕ H) eq-is-prime-decomposition-list-ℕ : @@ -203,7 +200,7 @@ module _ 1 <-ℕ y le-one-head-is-prime-decomposition-cons-list-ℕ = le-one-is-prime-ℕ y is-prime-head-is-prime-decomposition-cons-list-ℕ - + div-head-is-prime-decomposition-cons-list-ℕ : div-ℕ y x div-head-is-prime-decomposition-cons-list-ℕ = @@ -215,6 +212,23 @@ module _ ( cons y l) ( H)) + quotient-head-is-prime-decomposition-cons-list-ℕ : + ℕ + quotient-head-is-prime-decomposition-cons-list-ℕ = + quotient-div-ℕ y x div-head-is-prime-decomposition-cons-list-ℕ + + eq-quotient-head-is-prime-decomposition-cons-list-ℕ : + quotient-head-is-prime-decomposition-cons-list-ℕ *ℕ y = x + eq-quotient-head-is-prime-decomposition-cons-list-ℕ = + eq-quotient-div-ℕ y x div-head-is-prime-decomposition-cons-list-ℕ + + leq-one-quotient-head-is-prime-decomposition-cons-list-ℕ : + 1 ≤-ℕ quotient-head-is-prime-decomposition-cons-list-ℕ + leq-one-quotient-head-is-prime-decomposition-cons-list-ℕ = + leq-one-quotient-div-ℕ y x + ( div-head-is-prime-decomposition-cons-list-ℕ) + ( leq-one-is-prime-decomposition-list-ℕ x (cons y l) H) + is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ : is-nontrivial-divisor-ℕ x y is-nontrivial-divisor-head-is-prime-decomposition-cons-list-ℕ = @@ -240,14 +254,58 @@ module _ ( is-nonzero-leq-one-ℕ x ( leq-one-is-prime-decomposition-list-ℕ x (cons y l) H)) ( div-head-is-prime-decomposition-cons-list-ℕ)) + + is-sorted-tail-is-prime-decomposition-cons-list-ℕ : + is-sorted-list ℕ-Decidable-Total-Order l + is-sorted-tail-is-prime-decomposition-cons-list-ℕ = + is-sorted-tail-is-sorted-list + ( ℕ-Decidable-Total-Order) + ( cons y l) + ( is-sorted-is-prime-decomposition-list-ℕ x (cons y l) H) + + le-one-tail-is-prime-decomposition-cons-list-ℕ : + for-all-list l (le-ℕ 1) + le-one-tail-is-prime-decomposition-cons-list-ℕ = + pr2 (le-one-list-is-prime-decomposition-list-ℕ x (cons y l) H) + + compute-quotient-head-is-prime-decomposition-cons-list-ℕ : + quotient-head-is-prime-decomposition-cons-list-ℕ = + mul-list-ℕ l + compute-quotient-head-is-prime-decomposition-cons-list-ℕ = + is-injective-right-mul-ℕ y + ( is-nonzero-is-prime-ℕ y + ( is-prime-head-is-prime-decomposition-cons-list-ℕ)) + ( ( eq-quotient-head-is-prime-decomposition-cons-list-ℕ) ∙ + ( inv (eq-is-prime-decomposition-list-ℕ x (cons y l) H)) ∙ + ( commutative-mul-ℕ y (mul-list-ℕ l))) + + is-multiplicative-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ : + is-multiplicative-decomposition-list-ℕ + ( quotient-head-is-prime-decomposition-cons-list-ℕ) + ( l) + pr1 + is-multiplicative-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ = + le-one-tail-is-prime-decomposition-cons-list-ℕ + pr2 + is-multiplicative-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ = + inv compute-quotient-head-is-prime-decomposition-cons-list-ℕ + + is-prime-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ : + is-prime-decomposition-list-ℕ + ( quotient-head-is-prime-decomposition-cons-list-ℕ) + ( l) + is-prime-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ = + ( is-sorted-tail-is-prime-decomposition-cons-list-ℕ , + is-prime-tail-is-prime-decomposition-cons-list-ℕ , + is-multiplicative-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ) ``` -### Definition of the type of prime decomposition of a natural number +### The type of prime decompositions of a natural number ```agda prime-decomposition-list-ℕ : - (x : ℕ) → 1 ≤-ℕ x → UU lzero -prime-decomposition-list-ℕ x _ = + (x : ℕ) → UU lzero +prime-decomposition-list-ℕ x = Σ (list ℕ) (λ l → is-prime-decomposition-list-ℕ x l) ``` @@ -287,7 +345,7 @@ list-fundamental-theorem-arithmetic-ℕ x H = map-list nat-Prime-ℕ (list-of-primes-fundamental-theorem-arithmetic-ℕ x H) ``` -#### Computational rules for the list of primes in the prime decomposition of a number +#### Computation rules for the list of primes in the prime decomposition of a number ```agda compute-list-of-primes-fundamental-theorem-arithmetic-succ-ℕ : @@ -334,7 +392,7 @@ compute-list-fundamental-theorem-arithmetic-succ-ℕ x H = ( compute-list-of-primes-fundamental-theorem-arithmetic-succ-ℕ x H) ``` -#### The list of primes in the prime decomposition of a number is indeed a prime decomposition +#### The list of primes in the prime decomposition of a number is a prime decomposition ```agda is-prime-list-fundamental-theorem-arithmetic-ℕ : @@ -346,13 +404,11 @@ is-prime-list-fundamental-theorem-arithmetic-ℕ x H = le-one-list-fundamental-theorem-arithmetic-ℕ : (x : ℕ) (H : 1 ≤-ℕ x) → - for-all-list ℕ (le-ℕ-Prop 1) (list-fundamental-theorem-arithmetic-ℕ x H) + for-all-list (list-fundamental-theorem-arithmetic-ℕ x H) (le-ℕ 1) le-one-list-fundamental-theorem-arithmetic-ℕ x H = map-for-all-list - ( is-prime-ℕ-Prop) - ( le-ℕ-Prop 1) - ( le-one-is-prime-ℕ) ( list-fundamental-theorem-arithmetic-ℕ x H) + ( le-one-is-prime-ℕ) ( is-prime-list-fundamental-theorem-arithmetic-ℕ x H) is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ : @@ -364,7 +420,7 @@ is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ = ( λ n N → is-multiplicative-decomposition-list-ℕ n ( list-fundamental-theorem-arithmetic-ℕ n N)) - ( for-all-nil-list (le-ℕ-Prop 1) , refl) + ( for-all-nil-list (le-ℕ 1) , refl) ( λ n N f → ( le-one-list-fundamental-theorem-arithmetic-ℕ (succ-ℕ n) star) , ( ( ap @@ -416,7 +472,7 @@ is-least-element-list-least-prime-divisor-ℕ : ( l) is-least-element-list-least-prime-divisor-ℕ x H nil D = raise-star is-least-element-list-least-prime-divisor-ℕ x H (cons y l) D = - is-minimal-least-nontrivial-divisor-ℕ + ( is-minimal-least-nontrivial-divisor-ℕ ( succ-ℕ x) ( le-succ-leq-ℕ 1 x H) ( y) @@ -425,15 +481,12 @@ is-least-element-list-least-prime-divisor-ℕ x H (cons y l) D = ( y) ( quotient-div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) ( succ-ℕ x) - ( div-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( succ-ℕ x) - ( div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H))) + ( div-quotient-div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) ( pr2 (pr1 D))) , - is-least-element-list-least-prime-divisor-ℕ x H l (pr2 D) + is-least-element-list-least-prime-divisor-ℕ x H l (pr2 D)) is-least-element-head-list-fundamental-theorem-arithmetic-succ-ℕ : - (x : ℕ) → (H : 1 ≤-ℕ x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-least-element-list ( ℕ-Decidable-Total-Order) ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) @@ -450,21 +503,17 @@ is-least-element-head-list-fundamental-theorem-arithmetic-succ-ℕ x H = ( H) ( list-fundamental-theorem-arithmetic-ℕ ( quotient-div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) + ( leq-one-quotient-div-least-prime-divisor-ℕ ( succ-ℕ x) - ( div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( preserves-leq-succ-ℕ 1 x H))) + ( le-succ-leq-ℕ 1 x H))) ( is-list-of-nontrivial-divisors-fundamental-theorem-arithmetic-ℕ ( quotient-div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) + ( leq-one-quotient-div-least-prime-divisor-ℕ ( succ-ℕ x) - ( div-least-prime-divisor-ℕ (succ-ℕ x) (le-succ-leq-ℕ 1 x H)) - ( preserves-leq-succ-ℕ 1 x H))) + ( le-succ-leq-ℕ 1 x H))) is-sorted-least-element-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → (H : 1 ≤-ℕ x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-sorted-least-element-list ( ℕ-Decidable-Total-Order) ( list-fundamental-theorem-arithmetic-ℕ x H) @@ -483,22 +532,19 @@ is-sorted-least-element-list-fundamental-theorem-arithmetic-ℕ x H = ( l)) ( inv (compute-list-fundamental-theorem-arithmetic-succ-ℕ n N)) ( is-least-element-head-list-fundamental-theorem-arithmetic-succ-ℕ n N , - f - ( quotient-div-least-prime-divisor-ℕ + f ( quotient-div-least-prime-divisor-ℕ ( succ-ℕ n) ( le-succ-leq-ℕ 1 n N)) - ( leq-one-quotient-div-ℕ - ( nat-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) + ( leq-one-quotient-div-least-prime-divisor-ℕ ( succ-ℕ n) - ( div-least-prime-divisor-ℕ (succ-ℕ n) (le-succ-leq-ℕ 1 n N)) - ( preserves-leq-succ-ℕ 1 n N)) + ( le-succ-leq-ℕ 1 n N)) ( upper-bound-quotient-div-least-prime-divisor-ℕ n ( le-succ-leq-ℕ 1 n N)))) ( x) ( H) is-sorted-list-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) → (H : 1 ≤-ℕ x) → + (x : ℕ) (H : 1 ≤-ℕ x) → is-sorted-list ( ℕ-Decidable-Total-Order) ( list-fundamental-theorem-arithmetic-ℕ x H) @@ -519,7 +565,7 @@ pr2 (pr2 (is-prime-decomposition-list-fundamental-theorem-arithmetic-ℕ x H)) = is-multiplicative-decomposition-list-fundamental-theorem-arithmetic-ℕ x H prime-decomposition-fundamental-theorem-arithmetic-ℕ : - (x : ℕ) (H : 1 ≤-ℕ x) → prime-decomposition-list-ℕ x H + (x : ℕ) (H : 1 ≤-ℕ x) → prime-decomposition-list-ℕ x pr1 (prime-decomposition-fundamental-theorem-arithmetic-ℕ x H) = list-fundamental-theorem-arithmetic-ℕ x H pr2 (prime-decomposition-fundamental-theorem-arithmetic-ℕ x H) = @@ -530,64 +576,42 @@ pr2 (prime-decomposition-fundamental-theorem-arithmetic-ℕ x H) = #### The type of prime decompositions of any natural number is contractible -```text +```agda is-in-prime-decomposition-is-prime-divisor-ℕ : - (x : ℕ) (H : 1 ≤-ℕ x) (l : list ℕ) → - is-prime-decomposition-list-ℕ x l → - ( y : ℕ) → - div-ℕ y x → - is-prime-ℕ y → - y ∈-list l + (x : ℕ) (H : 1 ≤-ℕ x) (l : list ℕ) → is-prime-decomposition-list-ℕ x l → + (y : ℕ) → div-ℕ y x → is-prime-ℕ y → y ∈-list l is-in-prime-decomposition-is-prime-divisor-ℕ x H nil D y d p = ex-falso ( no-prime-divisors-one-ℕ y p ( tr (div-ℕ y) (inv (eq-is-prime-decomposition-list-ℕ x nil D)) d)) is-in-prime-decomposition-is-prime-divisor-ℕ x H (cons z l) D y d p = rec-coproduct - ( λ e → tr (λ w → w ∈-list (cons z l)) (inv e) (is-head z l)) + ( λ p → tr (_∈-list cons z l) (inv p) (is-head z l)) ( λ e → - is-in-tail - ( y) - ( z) - ( l) + is-in-tail y z l ( is-in-prime-decomposition-is-prime-divisor-ℕ - ( quotient-div-ℕ - ( z) - ( x) - ( div-head-is-prime-decomposition-cons-list-ℕ x z l D)) - ( leq-one-quotient-div-ℕ - ( z) - ( x) - ( div-head-is-prime-decomposition-cons-list-ℕ x z l D) - ( H)) + ( quotient-head-is-prime-decomposition-cons-list-ℕ x z l D) + ( leq-one-quotient-head-is-prime-decomposition-cons-list-ℕ + x z l D) ( l) - ( ( is-sorted-tail-is-sorted-list - ( ℕ-Decidable-Total-Order) - ( cons z l) - ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons z l) D)) , - ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x z l D) , - ( {!!} , {!!}) {- ( refl) -}) + ( is-prime-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ + x z l D) ( y) - {!!} - {- - ( div-right-factor-coprime-ℕ - ( y) - ( z) - ( mul-list-ℕ l) + ( div-right-factor-coprime-ℕ y z + ( quotient-head-is-prime-decomposition-cons-list-ℕ x z l D) ( tr - ( λ x → div-ℕ y x) - ( inv - ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ - ( x) - ( cons z l) - ( D))) + ( div-ℕ y) + ( ( inv + ( eq-quotient-head-is-prime-decomposition-cons-list-ℕ + x z l D)) ∙ + ( commutative-mul-ℕ _ z)) ( d)) ( is-relatively-prime-is-prime-ℕ ( y) ( z) ( p) ( is-prime-head-is-prime-decomposition-cons-list-ℕ x z l D) - ( e))) -} + ( e))) ( p))) ( has-decidable-equality-ℕ y z) @@ -611,7 +635,7 @@ is-lower-bound-head-prime-decomposition-list-ℕ x H y l D m d = ( y) ( nat-least-prime-divisor-ℕ m (pr1 d)) ( l) - ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons y l) D) + ( is-sorted-is-prime-decomposition-list-ℕ x (cons y l) D) ( is-in-prime-decomposition-is-prime-divisor-ℕ ( x) ( H) @@ -680,46 +704,27 @@ eq-prime-decomposition-list-ℕ x H (cons y l) (cons z p) I J = ( l) ( p) ( eq-prime-decomposition-list-ℕ - ( quotient-div-ℕ - ( y) - ( x) - ( div-head-is-prime-decomposition-cons-list-ℕ x y l I)) - ( leq-one-quotient-div-ℕ - ( y) - ( x) - ( div-head-is-prime-decomposition-cons-list-ℕ x y l I) - ( H)) + ( quotient-head-is-prime-decomposition-cons-list-ℕ x y l I) + ( leq-one-quotient-head-is-prime-decomposition-cons-list-ℕ x y l I) ( l) ( p) - ( ( is-sorted-tail-is-sorted-list - ( ℕ-Decidable-Total-Order) - ( cons y l) - ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons y l) I)) , - ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x y l I) , - {!!} {- refl -}) - ( ( is-sorted-tail-is-sorted-list - ( ℕ-Decidable-Total-Order) - ( cons z p) - ( is-sorted-list-is-prime-decomposition-list-ℕ x (cons z p) J)) , + ( is-prime-decomposition-tail-quotient-head-is-prime-decomposition-cons-list-ℕ + x y l I) + ( ( is-sorted-tail-is-prime-decomposition-cons-list-ℕ x z p J) , ( is-prime-tail-is-prime-decomposition-cons-list-ℕ x z p J) , - {!!} - {- - tr - ( λ y → is-multiplicative-decomposition-list-ℕ y p) - ( eq-quotient-div-eq-divisor-ℕ - ( z) - ( y) - ( x) - ( is-nonzero-is-prime-ℕ - ( z) - ( is-prime-head-is-prime-decomposition-cons-list-ℕ x z p J))) - ( inv (eq-head-prime-decomposition-list-ℕ x H y z l p I J)) - ( div-head-is-prime-decomposition-cons-list-ℕ x z p J) - ( div-head-is-prime-decomposition-cons-list-ℕ x y l I)) - ( refl) -})))) + ( le-one-tail-is-prime-decomposition-cons-list-ℕ x z p J) , + ( eq-quotient-div-eq-divisor-ℕ z y x + ( is-nonzero-is-prime-ℕ z + ( is-prime-head-is-prime-decomposition-cons-list-ℕ x z p J)) + ( inv (eq-head-prime-decomposition-list-ℕ x H y z l p I J)) + ( div-head-is-prime-decomposition-cons-list-ℕ x z p J) + ( div-head-is-prime-decomposition-cons-list-ℕ x y l I)) ∙ + ( inv + ( compute-quotient-head-is-prime-decomposition-cons-list-ℕ + x y l I)))))) fundamental-theorem-arithmetic-list-ℕ : - (x : ℕ) → (H : 1 ≤-ℕ x) → is-contr (prime-decomposition-list-ℕ x H) + (x : ℕ) → (H : 1 ≤-ℕ x) → is-contr (prime-decomposition-list-ℕ x) pr1 (fundamental-theorem-arithmetic-list-ℕ x H) = prime-decomposition-fundamental-theorem-arithmetic-ℕ x H pr2 (fundamental-theorem-arithmetic-list-ℕ x H) d = @@ -736,7 +741,7 @@ pr2 (fundamental-theorem-arithmetic-list-ℕ x H) d = ### The sorted list associated with the concatenation of the prime decomposition of `n` and the prime decomposition of `m` is the prime decomposition of `n *ℕ m` -```text +```agda is-prime-list-concat-list-ℕ : (p q : list ℕ) → is-prime-list-ℕ p → is-prime-list-ℕ q → is-prime-list-ℕ (concat-list p q) @@ -744,42 +749,42 @@ is-prime-list-concat-list-ℕ nil q Pp Pq = Pq is-prime-list-concat-list-ℕ (cons x p) q Pp Pq = pr1 Pp , is-prime-list-concat-list-ℕ p q (pr2 Pp) Pq -all-elements-is-prime-list-ℕ : +is-prime-is-in-list-ℕ : (p : list ℕ) → UU lzero -all-elements-is-prime-list-ℕ p = (x : ℕ) → x ∈-list p → is-prime-ℕ x +is-prime-is-in-list-ℕ p = (x : ℕ) → x ∈-list p → is-prime-ℕ x -all-elements-is-prime-list-tail-ℕ : - (p : list ℕ) (x : ℕ) (P : all-elements-is-prime-list-ℕ (cons x p)) → - all-elements-is-prime-list-ℕ p -all-elements-is-prime-list-tail-ℕ p x P y I = P y (is-in-tail y x p I) +is-prime-is-in-list-tail-ℕ : + (p : list ℕ) (x : ℕ) (P : is-prime-is-in-list-ℕ (cons x p)) → + is-prime-is-in-list-ℕ p +is-prime-is-in-list-tail-ℕ p x P y I = P y (is-in-tail y x p I) -all-elements-is-prime-list-is-prime-list-ℕ : - (p : list ℕ) → is-prime-list-ℕ p → all-elements-is-prime-list-ℕ p -all-elements-is-prime-list-is-prime-list-ℕ (cons x p) P .x (is-head .x .p) = +is-prime-is-in-list-is-prime-list-ℕ : + (p : list ℕ) → is-prime-list-ℕ p → is-prime-is-in-list-ℕ p +is-prime-is-in-list-is-prime-list-ℕ (cons x p) P .x (is-head .x .p) = pr1 P -all-elements-is-prime-list-is-prime-list-ℕ +is-prime-is-in-list-is-prime-list-ℕ ( cons x p) ( P) ( y) ( is-in-tail .y .x .p I) = - all-elements-is-prime-list-is-prime-list-ℕ p (pr2 P) y I + is-prime-is-in-list-is-prime-list-ℕ p (pr2 P) y I -is-prime-list-all-elements-is-prime-list-ℕ : - (p : list ℕ) → all-elements-is-prime-list-ℕ p → is-prime-list-ℕ p -is-prime-list-all-elements-is-prime-list-ℕ nil P = raise-star -is-prime-list-all-elements-is-prime-list-ℕ (cons x p) P = +is-prime-list-is-prime-is-in-list-ℕ : + (p : list ℕ) → is-prime-is-in-list-ℕ p → is-prime-list-ℕ p +is-prime-list-is-prime-is-in-list-ℕ nil P = raise-star +is-prime-list-is-prime-is-in-list-ℕ (cons x p) P = P x (is-head x p) , - is-prime-list-all-elements-is-prime-list-ℕ + is-prime-list-is-prime-is-in-list-ℕ ( p) - ( all-elements-is-prime-list-tail-ℕ p x P) + ( is-prime-is-in-list-tail-ℕ p x P) is-prime-list-permute-list-ℕ : (p : list ℕ) (t : Permutation (length-list p)) → is-prime-list-ℕ p → is-prime-list-ℕ (permute-list p t) is-prime-list-permute-list-ℕ p t P = - is-prime-list-all-elements-is-prime-list-ℕ + is-prime-list-is-prime-is-in-list-ℕ ( permute-list p t) - ( λ x I → all-elements-is-prime-list-is-prime-list-ℕ + ( λ x I → is-prime-is-in-list-is-prime-list-ℕ ( p) ( P) ( x) @@ -789,22 +794,12 @@ is-prime-list-permute-list-ℕ p t P = ( x) ( I))) -is-decomposition-list-concat-list-ℕ : - (n m : ℕ) (p q : list ℕ) → - is-multiplicative-decomposition-list-ℕ n p → - is-multiplicative-decomposition-list-ℕ m q → - is-multiplicative-decomposition-list-ℕ (n *ℕ m) (concat-list p q) -is-decomposition-list-concat-list-ℕ n m p q Dp Dq = - ( eq-mul-list-concat-list-ℕ p q ∙ - ( ap (mul-ℕ (mul-list-ℕ p)) Dq ∙ - ap (λ n → n *ℕ m) Dp)) - is-decomposition-list-permute-list-ℕ : (n : ℕ) (p : list ℕ) (t : Permutation (length-list p)) → is-multiplicative-decomposition-list-ℕ n p → is-multiplicative-decomposition-list-ℕ n (permute-list p t) is-decomposition-list-permute-list-ℕ n p t D = - inv (invariant-permutation-mul-list-ℕ p t) ∙ D + {!!} -- inv (invariant-permutation-mul-list-ℕ p t) ∙ D is-prime-decomposition-list-sort-concatenation-ℕ : (x y : ℕ) (H : 1 ≤-ℕ x) (I : 1 ≤-ℕ y) (p q : list ℕ) → @@ -851,19 +846,20 @@ pr2 ( permutation-insertion-sort-list ( ℕ-Decidable-Total-Order) ( concat-list p q)) - ( is-decomposition-list-concat-list-ℕ + ( is-multiplicative-decomposition-concat-list-ℕ ( x) ( y) ( p) ( q) ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ x p Dp) - ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ y q Dq))) + ( is-multiplicative-decomposition-is-prime-decomposition-list-ℕ + y q Dq))) prime-decomposition-list-sort-concatenation-ℕ : (x y : ℕ) (H : 1 ≤-ℕ x) (I : 1 ≤-ℕ y) (p q : list ℕ) → is-prime-decomposition-list-ℕ x p → is-prime-decomposition-list-ℕ y q → - prime-decomposition-list-ℕ (x *ℕ y) (preserves-leq-mul-ℕ 1 x 1 y H I) + prime-decomposition-list-ℕ (x *ℕ y) pr1 (prime-decomposition-list-sort-concatenation-ℕ x y H I p q Dp Dq) = insertion-sort-list ℕ-Decidable-Total-Order (concat-list p q) pr2 (prime-decomposition-list-sort-concatenation-ℕ x y H I p q Dp Dq) = diff --git a/src/elementary-number-theory/lists-of-prime-numbers.lagda.md b/src/elementary-number-theory/lists-of-prime-numbers.lagda.md index 9aabd05901..f47377c5b1 100644 --- a/src/elementary-number-theory/lists-of-prime-numbers.lagda.md +++ b/src/elementary-number-theory/lists-of-prime-numbers.lagda.md @@ -17,14 +17,16 @@ open import foundation.universe-levels open import lists.functoriality-lists open import lists.lists -open import lists.predicates-on-lists +open import lists.universal-quantification-lists ```
## Idea -A {{#concept "list of prime numbers"}} is a [list](lists.lists.md) of [natural numbers](elementary-number-theory.natural-numbers.md) such that each number in the list is [prime](elementary-number-theory.prime-numbers.md). +A {{#concept "list of prime numbers"}} is a [list](lists.lists.md) of +[natural numbers](elementary-number-theory.natural-numbers.md) such that each +number in the list is [prime](elementary-number-theory.prime-numbers.md). ## Definitions @@ -33,11 +35,11 @@ A {{#concept "list of prime numbers"}} is a [list](lists.lists.md) of [natural n ```agda is-prime-list-ℕ : list ℕ → UU lzero -is-prime-list-ℕ = for-all-list ℕ is-prime-ℕ-Prop +is-prime-list-ℕ l = for-all-list l is-prime-ℕ is-prop-is-prime-list-ℕ : (l : list ℕ) → is-prop (is-prime-list-ℕ l) -is-prop-is-prime-list-ℕ = is-prop-for-all-list ℕ is-prime-ℕ-Prop +is-prop-is-prime-list-ℕ l = is-prop-for-all-list l is-prime-ℕ-Prop ``` ## Properties @@ -51,4 +53,3 @@ is-prime-list-list-Prime-ℕ nil = raise-star is-prime-list-list-Prime-ℕ (cons x l) = ( is-prime-Prime-ℕ x , is-prime-list-list-Prime-ℕ l) ``` - diff --git a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md index c6589caff0..3888742b7f 100644 --- a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md @@ -21,7 +21,7 @@ open import foundation.identity-types open import lists.concatenation-lists open import lists.lists open import lists.permutation-lists -open import lists.predicates-on-lists +open import lists.universal-quantification-lists ```
@@ -45,8 +45,7 @@ mul-list-ℕ = fold-list 1 mul-ℕ ```agda leq-one-mul-list-ℕ : - (l : list ℕ) (H : for-all-list ℕ (leq-ℕ-Prop 1) l) → - 1 ≤-ℕ mul-list-ℕ l + (l : list ℕ) (H : for-all-list l (leq-ℕ 1)) → 1 ≤-ℕ mul-list-ℕ l leq-one-mul-list-ℕ nil H = refl-leq-ℕ 1 leq-one-mul-list-ℕ (cons n l) (H , K) = preserves-leq-mul-ℕ 1 n 1 (mul-list-ℕ l) H (leq-one-mul-list-ℕ l K) diff --git a/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md index 8f10ecc72c..487e0c5afe 100644 --- a/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md @@ -17,6 +17,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.nontrivial-divisors-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.action-on-identifications-functions open import foundation.cartesian-product-types open import foundation.dependent-pair-types open import foundation.identity-types @@ -24,15 +25,24 @@ open import foundation.propositions open import foundation.unit-type open import foundation.universe-levels +open import lists.concatenation-lists open import lists.lists -open import lists.predicates-on-lists +open import lists.universal-quantification-lists ```
## Idea -A {{#concept "multiplicative decomposition" Disambiguation="natural numbers" Agda=is-multiplicative-decomposition-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a [list](lists.lists.md) `l` of natural numbers [strictly greater than](elementary-number-theory.strict-inequality-natural-numbers.md) `1`, such that the [product](elementary-number-theory.multiplication-lists-of-natural-numbers.md) of the numbers in this list is [equal to](foundation-core.identity-types.md) `n`. +A +{{#concept "multiplicative decomposition" Disambiguation="natural numbers" Agda=is-multiplicative-decomposition-ℕ}} +of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a +[list](lists.lists.md) `l` of natural numbers +[strictly greater than](elementary-number-theory.strict-inequality-natural-numbers.md) +`1`, such that the +[product](elementary-number-theory.multiplication-lists-of-natural-numbers.md) +of the numbers in this list is [equal to](foundation-core.identity-types.md) +`n`. ## Definitions @@ -47,19 +57,18 @@ module _ is-multiplicative-decomposition-list-ℕ : UU lzero is-multiplicative-decomposition-list-ℕ = - for-all-list ℕ (le-ℕ-Prop 1) l × - (mul-list-ℕ l = n) + ( for-all-list l (le-ℕ 1)) × (mul-list-ℕ l = n) is-prop-is-multiplicative-decomposition-list-ℕ : is-prop (is-multiplicative-decomposition-list-ℕ) is-prop-is-multiplicative-decomposition-list-ℕ = is-prop-product - ( is-prop-for-all-list ℕ (le-ℕ-Prop 1) l) + ( is-prop-for-all-list l (le-ℕ-Prop 1)) ( is-set-ℕ (mul-list-ℕ l) n) le-one-list-is-multiplicative-decomposition-list-ℕ : is-multiplicative-decomposition-list-ℕ → - for-all-list ℕ (le-ℕ-Prop 1) l + for-all-list l (le-ℕ 1) le-one-list-is-multiplicative-decomposition-list-ℕ = pr1 eq-is-multiplicative-decomposition-list-ℕ : @@ -75,7 +84,7 @@ module _ ```agda is-multiplicative-decomposition-mul-list-ℕ : (l : list ℕ) → - for-all-list ℕ (le-ℕ-Prop 1) l → + for-all-list l (le-ℕ 1) → is-multiplicative-decomposition-list-ℕ (mul-list-ℕ l) l pr1 (is-multiplicative-decomposition-mul-list-ℕ l H) = H pr2 (is-multiplicative-decomposition-mul-list-ℕ l H) = refl @@ -96,7 +105,7 @@ module _ le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ : is-multiplicative-decomposition-list-ℕ n (cons x l) → - for-all-list ℕ (le-ℕ-Prop 1) l + for-all-list l (le-ℕ 1) le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ H = pr2 (le-one-list-is-multiplicative-decomposition-list-ℕ n (cons x l) H) ``` @@ -160,3 +169,23 @@ is-list-of-nontrivial-divisors-is-multiplicative-decomposition-list-ℕ ( le-one-tail-list-is-multiplicative-decomposition-cons-list-ℕ n x l ( H))))) ``` + +### The concatenation of two multiplicative decompositions is a multiplicative decomposition + +```agda +is-multiplicative-decomposition-concat-list-ℕ : + (n m : ℕ) (p q : list ℕ) → + is-multiplicative-decomposition-list-ℕ n p → + is-multiplicative-decomposition-list-ℕ m q → + is-multiplicative-decomposition-list-ℕ (n *ℕ m) (concat-list p q) +pr1 (is-multiplicative-decomposition-concat-list-ℕ n m l k H K) = + for-all-concat-list l k + ( le-ℕ 1) + ( le-one-list-is-multiplicative-decomposition-list-ℕ n l H) + ( le-one-list-is-multiplicative-decomposition-list-ℕ m k K) +pr2 (is-multiplicative-decomposition-concat-list-ℕ n m l k H K) = + eq-mul-list-concat-list-ℕ l k ∙ + ap-mul-ℕ + ( eq-is-multiplicative-decomposition-list-ℕ n l H) + ( eq-is-multiplicative-decomposition-list-ℕ m k K) +``` diff --git a/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md index b49688b451..ba53142781 100644 --- a/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md @@ -25,16 +25,23 @@ open import foundation.unit-type open import foundation.universe-levels open import lists.lists -open import lists.predicates-on-lists +open import lists.universal-quantification-lists ```
## Idea -A {{#concept "nontrivial divisor"}} of a [natural number](elementary-number-theory.natural-numbers.md) is a [divisor](elementary-number-theory.divisibility-natural-numbers.md) [strictly greater](elementary-number-theory.strict-inequality-natural-numbers.md) than `1`. +A {{#concept "nontrivial divisor"}} of a +[natural number](elementary-number-theory.natural-numbers.md) is a +[divisor](elementary-number-theory.divisibility-natural-numbers.md) +[strictly greater](elementary-number-theory.strict-inequality-natural-numbers.md) +than `1`. -Every number strictly greater than `1` has a least nontrivial divisor, which is a [prime number](elementary-number-theory.prime-numbers.md). This fact is essential in the [Fundamental Theorem of Arithmetic](elementary-number-theory.fundamental-theorem-of-arithmetic.md). +Every number strictly greater than `1` has a least nontrivial divisor, which is +a [prime number](elementary-number-theory.prime-numbers.md). This fact is +essential in the +[Fundamental Theorem of Arithmetic](elementary-number-theory.fundamental-theorem-of-arithmetic.md). ## Definitions @@ -46,10 +53,10 @@ Nontrivial divisors of a natural number are divisors strictly greater than `1`. module _ (n x : ℕ) where - + is-nontrivial-divisor-ℕ : UU lzero is-nontrivial-divisor-ℕ = (1 <-ℕ x) × (div-ℕ x n) - + is-prop-is-nontrivial-divisor-ℕ : is-prop is-nontrivial-divisor-ℕ is-prop-is-nontrivial-divisor-ℕ = is-prop-Σ @@ -74,8 +81,8 @@ module _ ```agda is-list-of-nontrivial-divisors-ℕ : ℕ → list ℕ → UU lzero -is-list-of-nontrivial-divisors-ℕ n = - for-all-list ℕ (is-nontrivial-divisor-ℕ-Prop n) +is-list-of-nontrivial-divisors-ℕ n l = + for-all-list l (is-nontrivial-divisor-ℕ n) ``` ## Properties @@ -118,7 +125,7 @@ is-list-of-nontrivial-divisors-div-ℕ : is-list-of-nontrivial-divisors-ℕ x l → is-list-of-nontrivial-divisors-ℕ y l is-list-of-nontrivial-divisors-div-ℕ x y d nil H = raise-star pr1 (is-list-of-nontrivial-divisors-div-ℕ x y d (cons z l) (H , K)) = - is-nontrivial-divisor-div-ℕ z x y d H + is-nontrivial-divisor-div-ℕ z x y d H pr2 (is-list-of-nontrivial-divisors-div-ℕ x y d (cons z l) (H , K)) = is-list-of-nontrivial-divisors-div-ℕ x y d l K ``` @@ -237,6 +244,23 @@ quotient-div-least-prime-divisor-ℕ x H = ( x) ( div-least-prime-divisor-ℕ x H) +div-quotient-div-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ x) → div-ℕ (quotient-div-least-prime-divisor-ℕ x H) x +div-quotient-div-least-prime-divisor-ℕ x H = + div-quotient-div-ℕ + ( nat-least-prime-divisor-ℕ x H) + ( x) + ( div-least-prime-divisor-ℕ x H) + +leq-one-quotient-div-least-prime-divisor-ℕ : + (x : ℕ) (H : 1 <-ℕ x) → 1 ≤-ℕ quotient-div-least-prime-divisor-ℕ x H +leq-one-quotient-div-least-prime-divisor-ℕ x H = + leq-one-quotient-div-ℕ + ( nat-least-prime-divisor-ℕ x H) + ( x) + ( div-least-prime-divisor-ℕ x H) + ( leq-le-ℕ 1 x H) + upper-bound-quotient-div-least-prime-divisor-ℕ : (x : ℕ) (H : 1 <-ℕ succ-ℕ x) → quotient-div-least-prime-divisor-ℕ (succ-ℕ x) H ≤-ℕ x diff --git a/src/lists.lagda.md b/src/lists.lagda.md index 0263108995..ca2b96a680 100644 --- a/src/lists.lagda.md +++ b/src/lists.lagda.md @@ -13,7 +13,7 @@ open import lists.lists public open import lists.lists-discrete-types public open import lists.permutation-lists public open import lists.permutation-vectors public -open import lists.predicates-on-lists public +open import lists.universal-quantification-lists public open import lists.quicksort-lists public open import lists.reversing-lists public open import lists.sort-by-insertion-lists public diff --git a/src/lists/predicates-on-lists.lagda.md b/src/lists/predicates-on-lists.lagda.md deleted file mode 100644 index 56e2eeb219..0000000000 --- a/src/lists/predicates-on-lists.lagda.md +++ /dev/null @@ -1,64 +0,0 @@ -# Predicates on lists - -```agda -module lists.predicates-on-lists where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.propositions -open import foundation.unit-type -open import foundation.universe-levels - -open import lists.lists -``` - -
- -## Definitions - -### Universal quantification over the elements of a list - -```agda -module _ - {l1 l2 : Level} (X : UU l1) (P : X → Prop l2) - where - - for-all-list-Prop : - (l : list X) → Prop l2 - for-all-list-Prop nil = raise-unit-Prop l2 - for-all-list-Prop (cons x l) = product-Prop (P x) (for-all-list-Prop l) - - for-all-list : - (l : list X) → UU l2 - for-all-list l = type-Prop (for-all-list-Prop l) - - is-prop-for-all-list : - (l : list X) → is-prop (for-all-list l) - is-prop-for-all-list l = is-prop-type-Prop (for-all-list-Prop l) -``` - -## Properties - -### Universal quantification over the elements of the empty list - -```agda -for-all-nil-list : - {l1 l2 : Level} {X : UU l1} (P : X → Prop l2) → - for-all-list X P nil -for-all-nil-list P = raise-star -``` - -### If a predicate implies another predicate, then its universal quantification over the elements of a list implies the universal-quantification of the other predicate - -```agda -map-for-all-list : - {l1 l2 l3 : Level} {X : UU l1} (P : X → Prop l2) (Q : X → Prop l3) → - ((x : X) → type-Prop (P x) → type-Prop (Q x)) → - (l : list X) → for-all-list X P l → for-all-list X Q l -map-for-all-list P Q f nil H = raise-star -pr1 (map-for-all-list P Q f (cons x l) (H , K)) = f _ H -pr2 (map-for-all-list P Q f (cons x l) (H , K)) = map-for-all-list P Q f l K -``` diff --git a/src/lists/sorted-lists.lagda.md b/src/lists/sorted-lists.lagda.md index dbd5f1f37e..1ece7eb3c0 100644 --- a/src/lists/sorted-lists.lagda.md +++ b/src/lists/sorted-lists.lagda.md @@ -58,6 +58,10 @@ module _ ### The proposition that an element is less or equal than every element in a list ```agda +module _ + {l1 l2 : Level} (X : Decidable-Total-Order l1 l2) + where + is-least-element-list-Prop : type-Decidable-Total-Order X → list (type-Decidable-Total-Order X) → Prop l2 @@ -78,9 +82,13 @@ module _ ### If a list is sorted, then its tail is also sorted ```agda +module _ + {l1 l2 : Level} (X : Decidable-Total-Order l1 l2) + where + is-sorted-tail-is-sorted-list : (l : list (type-Decidable-Total-Order X)) → - is-sorted-list l → is-sorted-list (tail-list l) + is-sorted-list X l → is-sorted-list X (tail-list l) is-sorted-tail-is-sorted-list nil _ = raise-star is-sorted-tail-is-sorted-list (cons x nil) s = raise-star is-sorted-tail-is-sorted-list (cons x (cons y l)) s = pr2 s @@ -89,10 +97,14 @@ module _ ### If a list is sorted then its head is less or equal than every element in the list ```agda +module _ + {l1 l2 : Level} (X : Decidable-Total-Order l1 l2) + where + leq-element-in-list-leq-head-is-sorted-list : (x y z : type-Decidable-Total-Order X) (l : list (type-Decidable-Total-Order X)) → - is-sorted-list (cons y l) → + is-sorted-list X (cons y l) → z ∈-list (cons y l) → leq-Decidable-Total-Order X x y → leq-Decidable-Total-Order X x z @@ -119,12 +131,16 @@ module _ ### An equivalent definition of being sorted ```agda +module _ + {l1 l2 : Level} (X : Decidable-Total-Order l1 l2) + where + is-sorted-least-element-list-Prop : list (type-Decidable-Total-Order X) → Prop l2 is-sorted-least-element-list-Prop nil = raise-unit-Prop l2 is-sorted-least-element-list-Prop (cons x l) = product-Prop - ( is-least-element-list-Prop x l) + ( is-least-element-list-Prop X x l) ( is-sorted-least-element-list-Prop l) is-sorted-least-element-list : @@ -134,7 +150,7 @@ module _ is-sorted-list-is-sorted-least-element-list : (l : list (type-Decidable-Total-Order X)) → - is-sorted-least-element-list l → is-sorted-list l + is-sorted-least-element-list l → is-sorted-list X l is-sorted-list-is-sorted-least-element-list nil _ = raise-star is-sorted-list-is-sorted-least-element-list (cons x nil) _ = @@ -149,10 +165,14 @@ module _ ### If a vector `v` of length `n` is sorted, then the list `list-vec n v` is also sorted ```agda +module _ + {l1 l2 : Level} (X : Decidable-Total-Order l1 l2) + where + is-sorted-list-is-sorted-vec : (n : ℕ) (v : vec (type-Decidable-Total-Order X) n) → is-sorted-vec X v → - is-sorted-list (list-vec n v) + is-sorted-list X (list-vec n v) is-sorted-list-is-sorted-vec 0 v S = raise-star is-sorted-list-is-sorted-vec 1 (x ∷ v) S = raise-star is-sorted-list-is-sorted-vec (succ-ℕ (succ-ℕ n)) (x ∷ y ∷ v) S = diff --git a/src/lists/universal-quantification-lists.lagda.md b/src/lists/universal-quantification-lists.lagda.md new file mode 100644 index 0000000000..3e9381f2f4 --- /dev/null +++ b/src/lists/universal-quantification-lists.lagda.md @@ -0,0 +1,141 @@ +# Universal quantification over the elements of lists + +```agda +module lists.universal-quantification-lists where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.propositions +open import foundation.unit-type +open import foundation.universe-levels + +open import lists.concatenation-lists +open import lists.lists +open import lists.permutation-lists + +open import finite-group-theory.permutations-standard-finite-types +``` + +
+ +## Idea + +Consider a predicate `P` on a type `X`, and consider a [list](lists.lists.md) `l` of elements of `X`. The {{#concept "universal quantification" Disambiguation="elements of a list" Agda=for-all-list}} of `P` over the elements of `l` is the type of choices of elements `P x` for each element `x ∈ l`. More specifically, the universal quantification `∀ l P` of `P` over `l` is defined inductively by + +```text + ∀ nil P := unit + ∀ (cons x l) P := (P x) × (∀ l P) +``` + +Alternatively, the universal quantification over the elements of a list can be defined directly by + +```text + ∀ l P := (x : X) → x ∈ l → P x. +``` + +These definitions are [equivalent](foundation-core.equivalences.md). However, note that the inductive definition of the universal quantification has the same universe level as `P`, while the direct definition is of universe level `l1 ⊔ l2`, where `l1` is the universe level of `X` and `l2` is the universe level of `P`. + +## Definitions + +### Universal quantification over the elements of a list + +```agda +module _ + {l1 l2 : Level} {X : UU l1} + where + + for-all-list : + (l : list X) (P : X → UU l2) → UU l2 + for-all-list nil P = raise-unit l2 + for-all-list (cons x l) P = (P x) × (for-all-list l P) + + for-all-nil-list : + (P : X → UU l2) → for-all-list nil P + for-all-nil-list P = raise-star + + head-for-all-cons-list : + (x : X) (l : list X) (P : X → UU l2) → for-all-list (cons x l) P → P x + head-for-all-cons-list x l P H = pr1 H + + tail-for-all-cons-list : + (x : X) (l : list X) (P : X → UU l2) → + for-all-list (cons x l) P → for-all-list l P + tail-for-all-cons-list x l P H = pr2 H +``` + +```agda +module _ + {l1 l2 : Level} {X : UU l1} + where + + is-prop-for-all-list : + (l : list X) (P : X → Prop l2) → is-prop (for-all-list l (type-Prop ∘ P)) + is-prop-for-all-list nil P = + is-prop-raise-unit + is-prop-for-all-list (cons x l) P = + is-prop-product (is-prop-type-Prop (P x)) (is-prop-for-all-list l P) + + for-all-list-Prop : + (l : list X) (P : X → Prop l2) → Prop l2 + pr1 (for-all-list-Prop l P) = for-all-list l (type-Prop ∘ P) + pr2 (for-all-list-Prop l P) = is-prop-for-all-list l P +``` + +### Universal quantification over the elements of a list + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (l : list X) (P : X → UU l2) + where + + for-all-elements-list : UU (l1 ⊔ l2) + for-all-elements-list = (x : X) → x ∈-list l → P x +``` + +## Properties + +### The two definitions of universal quantification over the elements of a list are equivalent + +```agda + +``` + +### If a predicate implies another predicate, then its universal quantification over the elements of a list implies the universal-quantification of the other predicate + +```agda +map-for-all-list : + {l1 l2 l3 : Level} {X : UU l1} {P : X → UU l2} {Q : X → UU l3} → + (l : list X) → ((x : X) → P x → Q x) → + for-all-list l P → for-all-list l Q +map-for-all-list nil f H = raise-star +pr1 (map-for-all-list (cons x l) f (H , K)) = f _ H +pr2 (map-for-all-list (cons x l) f (H , K)) = map-for-all-list l f K +``` + +### If two lists satisfy a universally quantified predicate, then so does their concatenation + +```agda +for-all-concat-list : + {l1 l2 : Level} {X : UU l1} (l k : list X) (P : X → UU l2) → + for-all-list l P → for-all-list k P → for-all-list (concat-list l k) P +for-all-concat-list nil k P H K = K +pr1 (for-all-concat-list (cons x l) k P H K) = + head-for-all-cons-list x l P H +pr2 (for-all-concat-list (cons x l) k P H K) = + for-all-concat-list l k P (tail-for-all-cons-list x l P H) K +``` + +### If a lists satisfies a universally quantified predicate, then so does any permutation of the list + +```text +for-all-permute-list : + {l1 l2 : Level} {X : UU l1} (l : list X) (P : X → UU l2) + (e : Permutation (length-list l)) → + for-all-list l P → for-all-list (permute-list l e) P +for-all-permute-list l P e H = {!!} +``` From 71f3b76a199d45eb45ed6a439117dba67d6fc3c7 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Thu, 2 Jan 2025 18:23:12 -0500 Subject: [PATCH 25/72] well-ordering principles for the integers --- references.bib | 29 +++ .../absolute-value-integers.lagda.md | 4 +- ...-strong-induction-natural-numbers.lagda.md | 6 +- .../congruence-natural-numbers.lagda.md | 194 +++++++++----- .../difference-integers.lagda.md | 11 +- .../divisibility-integers.lagda.md | 21 +- .../divisibility-natural-numbers.lagda.md | 21 +- ...uclidean-division-natural-numbers.lagda.md | 10 +- ...fundamental-theorem-of-arithmetic.lagda.md | 10 +- ...st-common-divisor-natural-numbers.lagda.md | 4 +- .../inequality-integer-fractions.lagda.md | 4 +- .../inequality-integers.lagda.md | 131 +++++++++- .../inequality-natural-numbers.lagda.md | 64 ++--- .../inequality-rational-numbers.lagda.md | 14 +- .../inequality-standard-finite-types.lagda.md | 22 +- .../irrationality-square-root-of-2.lagda.md | 3 +- .../kolakoski-sequence.lagda.md | 2 +- ...-arithmetic-standard-finite-types.lagda.md | 32 +-- ...lication-lists-of-natural-numbers.lagda.md | 2 +- .../parity-integers.lagda.md | 12 + .../parity-natural-numbers.lagda.md | 54 ++++ .../squares-integers.lagda.md | 21 +- .../squares-natural-numbers.lagda.md | 127 ++++++++- .../sums-of-natural-numbers.lagda.md | 31 ++- ...rdering-principle-natural-numbers.lagda.md | 4 +- ...well-ordering-principles-integers.lagda.md | 217 +++++++++++++++ src/linear-algebra/vectors.lagda.md | 71 ++++- src/lists/arrays.lagda.md | 24 +- src/lists/dependent-lists.lagda.md | 44 ++++ src/lists/elementhood-relation-lists.lagda.md | 246 ++++++++++++++++++ src/lists/lists.lagda.md | 14 +- src/lists/permutation-lists.lagda.md | 1 + .../universal-quantification-lists.lagda.md | 122 ++++++++- .../metric-space-of-rational-numbers.lagda.md | 2 +- .../products-ideals-rings.lagda.md | 16 +- .../products-left-ideals-rings.lagda.md | 16 +- .../products-right-ideals-rings.lagda.md | 16 +- ...ime-decomposition-natural-numbers.lagda.md | 6 +- .../inequality-types-with-counting.lagda.md | 8 +- .../standard-finite-types.lagda.md | 2 +- 40 files changed, 1386 insertions(+), 252 deletions(-) create mode 100644 src/elementary-number-theory/well-ordering-principles-integers.lagda.md create mode 100644 src/lists/dependent-lists.lagda.md create mode 100644 src/lists/elementhood-relation-lists.lagda.md diff --git a/references.bib b/references.bib index 9a25a11681..00398ca814 100644 --- a/references.bib +++ b/references.bib @@ -44,6 +44,17 @@ @article{AL19 langid = {english} } +@book{Andrews94, + title = {Number Theory}, + author = {Andrews, G.E.}, + isbn = {9780486682525}, + lccn = {lc94005243}, + series = {Dover Books on Mathematics}, + year = {1994}, + publisher = {Dover Publications} +} + + @misc{Awodey22, author = {{Awodey}, Steve}, title = "{On Hofmann-Streicher universes}", @@ -386,6 +397,24 @@ @inproceedings{KvR19 eprintclass = {cs, math} } +@book{Leveque12volI, + title = {Topics in Number Theory, Volume I}, + author = {LeVeque, W.J.}, + isbn = {9780486152080}, + series = {Dover Books on Mathematics}, + year = {2012}, + publisher = {Dover Publications} +} + +@book{Leveque12volII, + title = {Topics in Number Theory, Volume II}, + author = {LeVeque, W.J.}, + isbn = {9780486152080}, + series = {Dover Books on Mathematics}, + year = {2012}, + publisher = {Dover Publications} +} + @incollection{Makkai98, author = {Makkai, M.}, title = {Towards a categorical foundation of mathematics}, diff --git a/src/elementary-number-theory/absolute-value-integers.lagda.md b/src/elementary-number-theory/absolute-value-integers.lagda.md index 949bcc95db..d0297d3ee5 100644 --- a/src/elementary-number-theory/absolute-value-integers.lagda.md +++ b/src/elementary-number-theory/absolute-value-integers.lagda.md @@ -94,7 +94,7 @@ predecessor-law-abs-ℤ (inr (inl star)) = predecessor-law-abs-ℤ (inr (inr zero-ℕ)) = star predecessor-law-abs-ℤ (inr (inr (succ-ℕ x))) = - preserves-leq-succ-ℕ x (succ-ℕ x) (succ-leq-ℕ x) + leq-succ-leq-ℕ x (succ-ℕ x) (succ-leq-ℕ x) ``` ### `|x + 1| ≤ |x| + 1` @@ -105,7 +105,7 @@ successor-law-abs-ℤ : successor-law-abs-ℤ (inl zero-ℕ) = star successor-law-abs-ℤ (inl (succ-ℕ x)) = - preserves-leq-succ-ℕ x (succ-ℕ x) (succ-leq-ℕ x) + leq-succ-leq-ℕ x (succ-ℕ x) (succ-leq-ℕ x) successor-law-abs-ℤ (inr (inl star)) = refl-leq-ℕ zero-ℕ successor-law-abs-ℤ (inr (inr x)) = diff --git a/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md b/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md index 52541dc0ab..d2478c6ce3 100644 --- a/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md +++ b/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md @@ -412,7 +412,7 @@ based-strong-ind-ℕ' : (p0 : P k (refl-leq-ℕ k)) → (pS : (x : ℕ) → (H : k ≤-ℕ x) → based-□-≤-ℕ' k P x → - P (succ-ℕ x) (preserves-leq-succ-ℕ k x H)) + P (succ-ℕ x) (leq-succ-leq-ℕ k x H)) (n : ℕ) → (H : k ≤-ℕ n) → P n H based-strong-ind-ℕ' {l} k P p0 pS n H = based-strong-ind-ℕ @@ -425,10 +425,10 @@ based-strong-ind-ℕ' {l} k P p0 pS n H = ( p0)) ( λ x H p → apply-dependent-universal-property-contr - ( preserves-leq-succ-ℕ k x H) + ( leq-succ-leq-ℕ k x H) ( is-proof-irrelevant-is-prop ( is-prop-leq-ℕ k (succ-ℕ x)) - ( preserves-leq-succ-ℕ k x H)) + ( leq-succ-leq-ℕ k x H)) ( P (succ-ℕ x)) ( pS x H ( compute-base-□-≤-ℕ' k P x p))) ( n) diff --git a/src/elementary-number-theory/congruence-natural-numbers.lagda.md b/src/elementary-number-theory/congruence-natural-numbers.lagda.md index 57781d53a3..adde88a456 100644 --- a/src/elementary-number-theory/congruence-natural-numbers.lagda.md +++ b/src/elementary-number-theory/congruence-natural-numbers.lagda.md @@ -27,6 +27,18 @@ open import univalent-combinatorics.standard-finite-types
+## Idea + +Two [natural numbers](elementary-number-theory.natural-numbers.md) `x` and `y` are said to be {{#concept "congruent" Disambiguation="natural numbers" Agda=cong-ℕ WDID=Q3773677 WD="congruence of integers"}} modulo `k` if their [distance](elementary-number-theory.distance-natural-numbers.md) `dist-ℕ x y` is [divisible](elementary-number-theory.divisibility-natural-numbers.md) by `k`, i.e., if + +```text + k | dist-ℕ x y. +``` + +For each natural number `k`, the congruence relation modulo `k` defines an [equivalence relation](foundation.equivalence-relations.md). Furthermore, the congruence relations respect [addition](elementary-number-theory.addition-natural-numbers.md) and [multiplication](elementary-number-theory.multiplication-natural-numbers.md). + +[Quotienting](foundation.set-quotients.md) by the congruence relation leads to [modular arithmetic](elementary-number-theory.modular-arithmetic.md). Properties of the congruence relation with respect to the [standard finite types](univalent-combinatorics.standard-finite-types.md) are formalized in the file [`modular-arithmetic-standard-finite-types`](elementary-number-theory.modular-arithmetic-standard-finite-types.md). + ## Properties ### The congruence relations on the natural numbers @@ -36,56 +48,44 @@ cong-ℕ : ℕ → ℕ → ℕ → UU lzero cong-ℕ k x y = div-ℕ k (dist-ℕ x y) -_≡_mod_ : ℕ → ℕ → ℕ → UU lzero -x ≡ y mod k = cong-ℕ k x y - -concatenate-eq-cong-eq-ℕ : - (k : ℕ) {x1 x2 x3 x4 : ℕ} → - x1 = x2 → cong-ℕ k x2 x3 → x3 = x4 → cong-ℕ k x1 x4 -concatenate-eq-cong-eq-ℕ k refl H refl = H +_≡_mod-ℕ_ : ℕ → ℕ → ℕ → UU lzero +x ≡ y mod-ℕ k = cong-ℕ k x y +``` -concatenate-eq-cong-ℕ : - (k : ℕ) {x1 x2 x3 : ℕ} → - x1 = x2 → cong-ℕ k x2 x3 → cong-ℕ k x1 x3 -concatenate-eq-cong-ℕ k refl H = H +## Properties -concatenate-cong-eq-ℕ : - (k : ℕ) {x1 x2 x3 : ℕ} → - cong-ℕ k x1 x2 → x2 = x3 → cong-ℕ k x1 x3 -concatenate-cong-eq-ℕ k H refl = H +### Any two natural numbers are congruent to each other modulo `1` +```agda is-indiscrete-cong-one-ℕ : - (x y : ℕ) → cong-ℕ 1 x y + (x y : ℕ) → x ≡ y mod-ℕ 1 is-indiscrete-cong-one-ℕ x y = div-one-ℕ (dist-ℕ x y) +``` -is-discrete-cong-zero-ℕ : - (x y : ℕ) → cong-ℕ zero-ℕ x y → x = y -is-discrete-cong-zero-ℕ x y (pair k p) = - eq-dist-ℕ x y ((inv p) ∙ (right-zero-law-mul-ℕ k)) - -cong-zero-ℕ : - (k : ℕ) → cong-ℕ k k zero-ℕ -pr1 (cong-zero-ℕ k) = 1 -pr2 (cong-zero-ℕ k) = - (left-unit-law-mul-ℕ k) ∙ (inv (right-unit-law-dist-ℕ k)) +### Congruence of natural numbers is a reflexive relation +```agda refl-cong-ℕ : (k : ℕ) → is-reflexive (cong-ℕ k) pr1 (refl-cong-ℕ k x) = zero-ℕ pr2 (refl-cong-ℕ k x) = (left-zero-law-mul-ℕ (succ-ℕ k)) ∙ (inv (dist-eq-ℕ x x refl)) cong-identification-ℕ : - (k : ℕ) {x y : ℕ} → x = y → cong-ℕ k x y + (k : ℕ) {x y : ℕ} → x = y → x ≡ y mod-ℕ k cong-identification-ℕ k {x} refl = refl-cong-ℕ k x +``` + +### Congruence of natural numbers is a symmetric relation +```agda symmetric-cong-ℕ : (k : ℕ) → is-symmetric (cong-ℕ k) -pr1 (symmetric-cong-ℕ k x y (pair d p)) = d -pr2 (symmetric-cong-ℕ k x y (pair d p)) = p ∙ (symmetric-dist-ℕ x y) +pr1 (symmetric-cong-ℕ k x y (d , p)) = d +pr2 (symmetric-cong-ℕ k x y (d , p)) = p ∙ (symmetric-dist-ℕ x y) +``` -cong-zero-ℕ' : (k : ℕ) → cong-ℕ k zero-ℕ k -cong-zero-ℕ' k = - symmetric-cong-ℕ k k zero-ℕ (cong-zero-ℕ k) +### Congruence of natural numbers is a transitive relation +```agda transitive-cong-ℕ : (k : ℕ) → is-transitive (cong-ℕ k) transitive-cong-ℕ k x y z e d with is-total-dist-ℕ x y z transitive-cong-ℕ k x y z e d | inl α = @@ -96,39 +96,60 @@ transitive-cong-ℕ k x y z e d | inr (inl α) = transitive-cong-ℕ k x y z e d | inr (inr α) = div-left-summand-ℕ k (dist-ℕ x z) (dist-ℕ x y) d ( concatenate-div-eq-ℕ e (inv α)) +``` + +### Concatenating congruence and equality + +```agda +concatenate-eq-cong-eq-ℕ : + (k : ℕ) {x1 x2 x3 x4 : ℕ} → + x1 = x2 → x2 ≡ x3 mod-ℕ k → x3 = x4 → x1 ≡ x4 mod-ℕ k +concatenate-eq-cong-eq-ℕ k refl H refl = H + +concatenate-eq-cong-ℕ : + (k : ℕ) {x1 x2 x3 : ℕ} → + x1 = x2 → x2 ≡ x3 mod-ℕ k → x1 ≡ x3 mod-ℕ k +concatenate-eq-cong-ℕ k refl H = H + +concatenate-cong-eq-ℕ : + (k : ℕ) {x1 x2 x3 : ℕ} → + x1 ≡ x2 mod-ℕ k → x2 = x3 → x1 ≡ x3 mod-ℕ k +concatenate-cong-eq-ℕ k H refl = H concatenate-cong-eq-cong-ℕ : {k x1 x2 x3 x4 : ℕ} → - cong-ℕ k x1 x2 → x2 = x3 → cong-ℕ k x3 x4 → cong-ℕ k x1 x4 + x1 ≡ x2 mod-ℕ k → x2 = x3 → x3 ≡ x4 mod-ℕ k → x1 ≡ x4 mod-ℕ k concatenate-cong-eq-cong-ℕ {k} {x} {y} {.y} {z} H refl K = transitive-cong-ℕ k x y z K H concatenate-eq-cong-eq-cong-eq-ℕ : (k : ℕ) {x1 x2 x3 x4 x5 x6 : ℕ} → - x1 = x2 → cong-ℕ k x2 x3 → x3 = x4 → - cong-ℕ k x4 x5 → x5 = x6 → cong-ℕ k x1 x6 + x1 = x2 → x2 ≡ x3 mod-ℕ k → x3 = x4 → + x4 ≡ x5 mod-ℕ k → x5 = x6 → x1 ≡ x6 mod-ℕ k concatenate-eq-cong-eq-cong-eq-ℕ k {x} {.x} {y} {.y} {z} {.z} refl H refl K refl = transitive-cong-ℕ k x y z K H ``` +### If the distance between two natural numbers is strictly less than `k`, then they are congruent modulo `k` if and only if they are equal + ```agda eq-cong-le-dist-ℕ : - (k x y : ℕ) → le-ℕ (dist-ℕ x y) k → cong-ℕ k x y → x = y + (k x y : ℕ) → le-ℕ (dist-ℕ x y) k → x ≡ y mod-ℕ k → x = y eq-cong-le-dist-ℕ k x y H K = eq-dist-ℕ x y (is-zero-div-ℕ k (dist-ℕ x y) H K) ``` +### Any two numbers strictly less than `k` are congruent modulo `k` if and only if they are equal + ```agda eq-cong-le-ℕ : - (k x y : ℕ) → le-ℕ x k → le-ℕ y k → cong-ℕ k x y → x = y + (k x y : ℕ) → le-ℕ x k → le-ℕ y k → x ≡ y mod-ℕ k → x = y eq-cong-le-ℕ k x y H K = eq-cong-le-dist-ℕ k x y (strict-upper-bound-dist-ℕ k x y H K) -``` -```agda eq-cong-nat-Fin : - (k : ℕ) (x y : Fin k) → cong-ℕ k (nat-Fin k x) (nat-Fin k y) → x = y + (k : ℕ) (x y : Fin k) → nat-Fin k x ≡ nat-Fin k y mod-ℕ k → x = y eq-cong-nat-Fin (succ-ℕ k) x y H = is-injective-nat-Fin (succ-ℕ k) ( eq-cong-le-ℕ (succ-ℕ k) (nat-Fin (succ-ℕ k) x) (nat-Fin (succ-ℕ k) y) @@ -137,36 +158,75 @@ eq-cong-nat-Fin (succ-ℕ k) x y H = ( H)) ``` +### Any natural number is congruent to `0` modulo itself + +```agda +cong-zero-ℕ : + (k : ℕ) → k ≡ 0 mod-ℕ k +pr1 (cong-zero-ℕ k) = 1 +pr2 (cong-zero-ℕ k) = + (left-unit-law-mul-ℕ k) ∙ (inv (right-unit-law-dist-ℕ k)) + +cong-zero-ℕ' : (k : ℕ) → 0 ≡ k mod-ℕ k +cong-zero-ℕ' k = + symmetric-cong-ℕ k k zero-ℕ (cong-zero-ℕ k) +``` + +### The inclusion of `0` from `Fin (k + 1)` into `ℕ` is congruent to `0` modulo `k + 1` + ```agda cong-is-zero-nat-zero-Fin : - {k : ℕ} → cong-ℕ (succ-ℕ k) (nat-Fin (succ-ℕ k) (zero-Fin k)) zero-ℕ + {k : ℕ} → nat-Fin (succ-ℕ k) (zero-Fin k) ≡ 0 mod-ℕ succ-ℕ k cong-is-zero-nat-zero-Fin {k} = cong-identification-ℕ (succ-ℕ k) (is-zero-nat-zero-Fin {k}) ``` +### Any two natural numbers that are congruent modulo `0` are equal + +In other words, the congruence relation modulo `0` is discrete. + ```agda -eq-cong-zero-ℕ : (x y : ℕ) → cong-ℕ zero-ℕ x y → x = y -eq-cong-zero-ℕ x y H = - eq-dist-ℕ x y (is-zero-div-zero-ℕ (dist-ℕ x y) H) +is-discrete-cong-zero-ℕ : + (x y : ℕ) → x ≡ y mod-ℕ 0 → x = y +is-discrete-cong-zero-ℕ x y (k , p) = + eq-dist-ℕ x y (inv p ∙ right-zero-law-mul-ℕ k) +``` + +### If `x` and `x + 1` are congruent modulo `k`, then `k = 1` -is-one-cong-succ-ℕ : {k : ℕ} (x : ℕ) → cong-ℕ k x (succ-ℕ x) → is-one-ℕ k +```agda +is-one-cong-succ-ℕ : {k : ℕ} (x : ℕ) → x ≡ succ-ℕ x mod-ℕ k → is-one-ℕ k is-one-cong-succ-ℕ {k} x H = is-one-div-one-ℕ k (tr (div-ℕ k) (is-one-dist-succ-ℕ x) H) ``` +### A number is divisible by `k` if and only if it is congruent to `0` modulo `k` + +```agda +div-cong-zero-ℕ : + (k x : ℕ) → x ≡ 0 mod-ℕ k → div-ℕ k x +div-cong-zero-ℕ k x = + tr (div-ℕ k) (right-unit-law-dist-ℕ x) + +cong-zero-div-ℕ : + (k x : ℕ) → div-ℕ k x → x ≡ 0 mod-ℕ k +cong-zero-div-ℕ k x = + tr (div-ℕ k) (inv (right-unit-law-dist-ℕ x)) +``` + ### Congruence is invariant under scalar multiplication ```agda scalar-invariant-cong-ℕ : - (k x y z : ℕ) → cong-ℕ k x y → cong-ℕ k (z *ℕ x) (z *ℕ y) -pr1 (scalar-invariant-cong-ℕ k x y z (pair d p)) = z *ℕ d -pr2 (scalar-invariant-cong-ℕ k x y z (pair d p)) = + (k x y z : ℕ) → x ≡ y mod-ℕ k → z *ℕ x ≡ z *ℕ y mod-ℕ k +pr1 (scalar-invariant-cong-ℕ k x y z (d , p)) = z *ℕ d +pr2 (scalar-invariant-cong-ℕ k x y z (d , p)) = ( associative-mul-ℕ z d k) ∙ ( ( ap (z *ℕ_) p) ∙ ( left-distributive-mul-dist-ℕ x y z)) scalar-invariant-cong-ℕ' : - (k x y z : ℕ) → cong-ℕ k x y → cong-ℕ k (x *ℕ z) (y *ℕ z) + (k x y z : ℕ) → x ≡ y mod-ℕ k → x *ℕ z ≡ y *ℕ z mod-ℕ k scalar-invariant-cong-ℕ' k x y z H = concatenate-eq-cong-eq-ℕ k ( commutative-mul-ℕ x z) @@ -179,7 +239,7 @@ scalar-invariant-cong-ℕ' k x y z H = ```agda congruence-mul-ℕ : (k : ℕ) {x y x' y' : ℕ} → - cong-ℕ k x x' → cong-ℕ k y y' → cong-ℕ k (x *ℕ y) (x' *ℕ y') + x ≡ x' mod-ℕ k → y ≡ y' mod-ℕ k → x *ℕ y ≡ x' *ℕ y' mod-ℕ k congruence-mul-ℕ k {x} {y} {x'} {y'} H K = transitive-cong-ℕ k (x *ℕ y) (x *ℕ y') (x' *ℕ y') ( scalar-invariant-cong-ℕ' k x x' y' H) @@ -190,13 +250,13 @@ congruence-mul-ℕ k {x} {y} {x'} {y'} H K = ```agda translation-invariant-cong-ℕ : - (k x y z : ℕ) → cong-ℕ k x y → cong-ℕ k (z +ℕ x) (z +ℕ y) -pr1 (translation-invariant-cong-ℕ k x y z (pair d p)) = d -pr2 (translation-invariant-cong-ℕ k x y z (pair d p)) = + (k x y z : ℕ) → x ≡ y mod-ℕ k → z +ℕ x ≡ z +ℕ y mod-ℕ k +pr1 (translation-invariant-cong-ℕ k x y z (d , p)) = d +pr2 (translation-invariant-cong-ℕ k x y z (d , p)) = p ∙ inv (translation-invariant-dist-ℕ z x y) translation-invariant-cong-ℕ' : - (k x y z : ℕ) → cong-ℕ k x y → cong-ℕ k (x +ℕ z) (y +ℕ z) + (k x y z : ℕ) → x ≡ y mod-ℕ k → x +ℕ z ≡ y +ℕ z mod-ℕ k translation-invariant-cong-ℕ' k x y z H = concatenate-eq-cong-eq-ℕ k ( commutative-add-ℕ x z) @@ -204,17 +264,17 @@ translation-invariant-cong-ℕ' k x y z H = ( commutative-add-ℕ z y) step-invariant-cong-ℕ : - (k x y : ℕ) → cong-ℕ k x y → cong-ℕ k (succ-ℕ x) (succ-ℕ y) + (k x y : ℕ) → x ≡ y mod-ℕ k → succ-ℕ x ≡ succ-ℕ y mod-ℕ k step-invariant-cong-ℕ k x y = translation-invariant-cong-ℕ' k x y 1 reflects-cong-add-ℕ : - {k : ℕ} (x : ℕ) {y z : ℕ} → cong-ℕ k (x +ℕ y) (x +ℕ z) → cong-ℕ k y z -pr1 (reflects-cong-add-ℕ {k} x {y} {z} (pair d p)) = d -pr2 (reflects-cong-add-ℕ {k} x {y} {z} (pair d p)) = + {k : ℕ} (x : ℕ) {y z : ℕ} → x +ℕ y ≡ x +ℕ z mod-ℕ k → y ≡ z mod-ℕ k +pr1 (reflects-cong-add-ℕ {k} x {y} {z} (d , p)) = d +pr2 (reflects-cong-add-ℕ {k} x {y} {z} (d , p)) = p ∙ translation-invariant-dist-ℕ x y z reflects-cong-add-ℕ' : - {k : ℕ} (x : ℕ) {y z : ℕ} → cong-ℕ k (add-ℕ' x y) (add-ℕ' x z) → cong-ℕ k y z + {k : ℕ} (x : ℕ) {y z : ℕ} → y +ℕ x ≡ z +ℕ x mod-ℕ k → y ≡ z mod-ℕ k reflects-cong-add-ℕ' {k} x {y} {z} H = reflects-cong-add-ℕ x {y} {z} ( concatenate-eq-cong-eq-ℕ k @@ -222,3 +282,19 @@ reflects-cong-add-ℕ' {k} x {y} {z} H = ( H) ( commutative-add-ℕ z x)) ``` + +### The congruence relations respect addition + +```agda +congruence-add-ℕ : + (k : ℕ) {x y x' y' : ℕ} → + x ≡ x' mod-ℕ k → y ≡ y' mod-ℕ k → x +ℕ y ≡ x' +ℕ y' mod-ℕ k +congruence-add-ℕ k {x} {y} {x'} {y'} H K = + transitive-cong-ℕ k (x +ℕ y) (x +ℕ y') (x' +ℕ y') + ( translation-invariant-cong-ℕ' k x x' y' H) + ( translation-invariant-cong-ℕ k y y' x K) +``` + +## See also + +- [Congruence relations on the integers](elementary-number-theory.congruence-integers.md) diff --git a/src/elementary-number-theory/difference-integers.lagda.md b/src/elementary-number-theory/difference-integers.lagda.md index 6cd796526c..78b8eee55e 100644 --- a/src/elementary-number-theory/difference-integers.lagda.md +++ b/src/elementary-number-theory/difference-integers.lagda.md @@ -27,9 +27,11 @@ are derived there. ## Definition +### The difference of two integers + ```agda diff-ℤ : ℤ → ℤ → ℤ -diff-ℤ x y = x +ℤ (neg-ℤ y) +diff-ℤ x y = x +ℤ neg-ℤ y infixl 36 _-ℤ_ _-ℤ_ = diff-ℤ @@ -38,6 +40,13 @@ ap-diff-ℤ : {x x' y y' : ℤ} → x = x' → y = y' → x -ℤ y = x' - ap-diff-ℤ p q = ap-binary diff-ℤ p q ``` +### The difference of two integers with the negative integer on the left + +```agda +left-diff-ℤ : ℤ → ℤ → ℤ +left-diff-ℤ x y = neg-ℤ x +ℤ y +``` + ## Properties ### Two integers with a difference equal to zero are equal diff --git a/src/elementary-number-theory/divisibility-integers.lagda.md b/src/elementary-number-theory/divisibility-integers.lagda.md index 231f62e274..4a6a6d3bfe 100644 --- a/src/elementary-number-theory/divisibility-integers.lagda.md +++ b/src/elementary-number-theory/divisibility-integers.lagda.md @@ -293,13 +293,20 @@ div-left-summand-ℤ x y z H K = ```agda preserves-div-mul-ℤ : + (k l x y : ℤ) → div-ℤ k x → div-ℤ l y → div-ℤ (k *ℤ l) (x *ℤ y) +pr1 (preserves-div-mul-ℤ k l x y (q , α) (p , β)) = q *ℤ p +pr2 (preserves-div-mul-ℤ k l x y (q , α) (p , β)) = + interchange-law-mul-mul-ℤ q p k l ∙ ap-mul-ℤ α β + +preserves-div-left-mul-ℤ : (k x y : ℤ) → div-ℤ x y → div-ℤ (k *ℤ x) (k *ℤ y) -pr1 (preserves-div-mul-ℤ k x y (pair q p)) = q -pr2 (preserves-div-mul-ℤ k x y (pair q p)) = - ( inv (associative-mul-ℤ q k x)) ∙ - ( ( ap (_*ℤ x) (commutative-mul-ℤ q k)) ∙ - ( ( associative-mul-ℤ k q x) ∙ - ( ap (k *ℤ_) p))) +preserves-div-left-mul-ℤ k x y H = + preserves-div-mul-ℤ k x k y (refl-div-ℤ k) H + +preserves-div-right-mul-ℤ : + (k x y : ℤ) → div-ℤ x y → div-ℤ (x *ℤ k) (y *ℤ k) +preserves-div-right-mul-ℤ k x y H = + preserves-div-mul-ℤ x k y k H (refl-div-ℤ k) ``` ### Multiplication by a nonzero number reflects divisibility @@ -335,7 +342,7 @@ div-div-quotient-div-ℤ x y d H K = tr ( div-ℤ (d *ℤ x)) ( eq-quotient-div-ℤ' d y H) - ( preserves-div-mul-ℤ d x (quotient-div-ℤ d y H) K) + ( preserves-div-left-mul-ℤ d x (quotient-div-ℤ d y H) K) ``` ### `sim-unit-ℤ x y` holds if and only if `x|y` and `y|x` diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 913a825491..99dc51bc47 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -474,13 +474,20 @@ is-one-div-ℕ x y H K = is-one-div-one-ℕ x (div-right-summand-ℕ x y 1 H K) ```agda preserves-div-mul-ℕ : + (k l x y : ℕ) → div-ℕ k x → div-ℕ l y → div-ℕ (k *ℕ l) (x *ℕ y) +pr1 (preserves-div-mul-ℕ k l x y (q , α) (p , β)) = q *ℕ p +pr2 (preserves-div-mul-ℕ k l x y (q , α) (p , β)) = + interchange-law-mul-mul-ℕ q p k l ∙ ap-mul-ℕ α β + +preserves-div-left-mul-ℕ : (k x y : ℕ) → div-ℕ x y → div-ℕ (k *ℕ x) (k *ℕ y) -pr1 (preserves-div-mul-ℕ k x y (q , p)) = q -pr2 (preserves-div-mul-ℕ k x y (q , p)) = - ( inv (associative-mul-ℕ q k x)) ∙ - ( ( ap (_*ℕ x) (commutative-mul-ℕ q k)) ∙ - ( ( associative-mul-ℕ k q x) ∙ - ( ap (k *ℕ_) p))) +preserves-div-left-mul-ℕ k x y = + preserves-div-mul-ℕ k x k y (refl-div-ℕ k) + +preserves-div-right-mul-ℕ : + (k x y : ℕ) → div-ℕ x y → div-ℕ (x *ℕ k) (y *ℕ k) +preserves-div-right-mul-ℕ k x y H = + preserves-div-mul-ℕ x k y k H (refl-div-ℕ k) ``` ### Multiplication by a nonzero number reflects divisibility @@ -516,7 +523,7 @@ div-div-quotient-div-ℕ x y d H K = tr ( div-ℕ (d *ℕ x)) ( eq-quotient-div-ℕ' d y H) - ( preserves-div-mul-ℕ d x (quotient-div-ℕ d y H) K) + ( preserves-div-left-mul-ℕ d x (quotient-div-ℕ d y H) K) ``` ### If `d` divides a nonzero number `x`, then the quotient `x/d` is also nonzero diff --git a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md index 028788652f..2478117a58 100644 --- a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md +++ b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md @@ -29,15 +29,15 @@ open import univalent-combinatorics.standard-finite-types ## Idea -{{#concept "Euclidean division" Agda=euclidean-division-ℕ}} is a binary operation on the [natural numbers](elementary-number-theory.natural-numbers.md) given by division with remainder. In other words, the Euclidean division of -`n` by `d` gives the unique pair of natural numbers `q` and `r +## Idea + +The {{#concept "square" Disambiguation="integer" Agda=square-ℤ}} `a²` of +an [integer](elementary-number-theory.integers.md) `a` is the +[product](elementary-number-theory.multiplication-natural-numbers.md) + +```text + a² := a * a +``` + +of `a` with itself. + ## Definitions ### The square of an integer @@ -216,6 +229,12 @@ Consider a integer `n`. The following are equivalent: - Its square is even. - Its square is divisible by 4. +```agda +div-four-square-is-even-ℤ : + (a : ℤ) → is-even-ℤ a → div-ℤ 4 (square-ℤ a) +div-four-square-is-even-ℤ a H = ? +``` + ```text div-four-square-is-even-ℤ : (a : ℤ) → is-even-ℤ a → div-ℤ 4 (square-ℤ a) diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index d78891eed1..d920c0faa4 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -1,4 +1,4 @@ -# Squares in the natural numbers +# Squares of natural numbers ```agda module elementary-number-theory.squares-natural-numbers where @@ -8,7 +8,9 @@ module elementary-number-theory.squares-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.congruence-natural-numbers open import elementary-number-theory.decidable-types +open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers @@ -132,7 +134,7 @@ square-succ-succ-ℕ n = preserves-order-square-ℕ : (m n : ℕ) → m ≤-ℕ n → square-ℕ m ≤-ℕ square-ℕ n preserves-order-square-ℕ m n H = - preserves-leq-mul-ℕ m n m n H H + preserves-order-mul-ℕ m n m n H H preserves-strict-order-square-ℕ : (m n : ℕ) → m <-ℕ n → square-ℕ m <-ℕ square-ℕ n @@ -327,6 +329,11 @@ Consider a natural number `n`. The following are equivalent: - Its square is divisible by 4. ```agda +compute-square-even-number-ℕ : + (n : ℕ) → square-ℕ (even-number-ℕ n) = 4 *ℕ square-ℕ n +compute-square-even-number-ℕ n = + distributive-square-mul-ℕ 2 n + div-four-square-is-even-ℕ : (n : ℕ) → is-even-ℕ n → div-ℕ 4 (square-ℕ n) pr1 (div-four-square-is-even-ℕ .(m *ℕ 2) (m , refl)) = @@ -363,6 +370,122 @@ is-even-div-four-square-ℕ n H = is-even-is-even-square-ℕ n (is-even-div-4-ℕ (square-ℕ n) H) ``` +### Equivalent characterizations of a number being odd in terms of its square + +Consider a natural number `n`. The following are equivalent: + +- The number `n` is odd. +- Its square is odd. +- Its square is congruent to `1` modulo `4`. +- Its square is congruent to `1` modulo `8`. + +**Proof.** If `n` is of the form `2k + 1`, then its square is of the form +`4k(k+1) + 1`. Since the number `k(k + 1)` is even, it follows that the square +of an odd number is congruent to `1` modulo `8`. Further more, since squares of +even numbers are even, and hence not congruent to `1` modulo `8`, we get a +logical equivalence. + +```agda +square-odd-number-ℕ : ℕ → ℕ +square-odd-number-ℕ n = 4 *ℕ square-ℕ n +ℕ 4 *ℕ n +ℕ 1 + +square-odd-number-ℕ' : ℕ → ℕ +square-odd-number-ℕ' n = 4 *ℕ (n *ℕ (n +ℕ 1)) +ℕ 1 + +compute-square-odd-number-ℕ : + (n : ℕ) → square-ℕ (odd-number-ℕ n) = square-odd-number-ℕ n +compute-square-odd-number-ℕ n = + square-succ-ℕ (2 *ℕ n) ∙ + ap + ( succ-ℕ) + ( ap-add-ℕ + ( compute-square-even-number-ℕ n) + ( inv (associative-mul-ℕ 2 2 n))) + +compute-square-odd-number-ℕ' : + (n : ℕ) → square-ℕ (odd-number-ℕ n) = square-odd-number-ℕ' n +compute-square-odd-number-ℕ' n = + compute-square-odd-number-ℕ n ∙ + inv + ( ap + ( succ-ℕ) + ( ( ap (4 *ℕ_) ( right-successor-law-mul-ℕ n n)) ∙ + ( left-distributive-mul-add-ℕ 4 (square-ℕ n) n))) + +compute-distance-to-1-square-odd-number-ℕ : + (n : ℕ) → dist-ℕ (square-ℕ (odd-number-ℕ n)) 1 = 4 *ℕ (n *ℕ (n +ℕ 1)) +compute-distance-to-1-square-odd-number-ℕ n = + ( ap (λ x → dist-ℕ x 1) (compute-square-odd-number-ℕ' n)) ∙ + ( right-unit-law-dist-ℕ _) + +has-odd-expansion-square-has-odd-expansion-ℕ : + (n : ℕ) → has-odd-expansion-ℕ n → has-odd-expansion-ℕ (square-ℕ n) +pr1 (has-odd-expansion-square-has-odd-expansion-ℕ ._ (k , refl)) = + 2 *ℕ (k *ℕ (k +ℕ 1)) +pr2 (has-odd-expansion-square-has-odd-expansion-ℕ ._ (k , refl)) = + inv + ( ( compute-square-odd-number-ℕ' k) ∙ + ( ap succ-ℕ (associative-mul-ℕ 2 2 (k *ℕ (k +ℕ 1))))) + +is-odd-square-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → is-odd-ℕ (square-ℕ n) +is-odd-square-is-odd-ℕ n H = + is-odd-has-odd-expansion-ℕ + ( square-ℕ n) + ( has-odd-expansion-square-has-odd-expansion-ℕ n + ( has-odd-expansion-is-odd-ℕ n H)) + +is-odd-is-odd-square-ℕ : + (n : ℕ) → is-odd-ℕ (square-ℕ n) → is-odd-ℕ n +is-odd-is-odd-square-ℕ n = + map-neg (is-even-square-is-even-ℕ n) + +is-1-mod-4-square-has-odd-expansion-ℕ : + (n : ℕ) → has-odd-expansion-ℕ n → square-ℕ n ≡ 1 mod-ℕ 4 +is-1-mod-4-square-has-odd-expansion-ℕ ._ (k , refl) = + tr + ( div-ℕ 4) + ( inv (compute-distance-to-1-square-odd-number-ℕ k)) + ( div-mul-ℕ' (k *ℕ (k +ℕ 1)) 4 4 (refl-div-ℕ 4)) + +is-1-mod-8-square-has-odd-expansion-ℕ : + (n : ℕ) → has-odd-expansion-ℕ n → square-ℕ n ≡ 1 mod-ℕ 8 +is-1-mod-8-square-has-odd-expansion-ℕ ._ (k , refl) = + tr + ( div-ℕ 8) + ( inv (compute-distance-to-1-square-odd-number-ℕ k)) + ( preserves-div-mul-ℕ 4 2 4 + ( k *ℕ (k +ℕ 1)) + ( refl-div-ℕ 4) + ( is-even-mul-succ-ℕ k)) + +is-1-mod-8-square-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → square-ℕ n ≡ 1 mod-ℕ 8 +is-1-mod-8-square-is-odd-ℕ n H = + is-1-mod-8-square-has-odd-expansion-ℕ n (has-odd-expansion-is-odd-ℕ n H) +``` + +### Any two odd squares are congruent modulo `8` + +This solves exercise 6 of section 2.1 in {{#cite "andrews94"}}. + +```agda +cong-8-square-odd-number-ℕ : + (m n : ℕ) → is-odd-ℕ m → is-odd-ℕ n → square-ℕ m ≡ square-ℕ n mod-ℕ 8 +cong-8-square-odd-number-ℕ m n H K = + transitive-cong-ℕ 8 + ( square-ℕ m) + ( 1) + ( square-ℕ n) + ( symmetric-cong-ℕ 8 (square-ℕ n) 1 (is-1-mod-8-square-is-odd-ℕ n K)) + ( is-1-mod-8-square-is-odd-ℕ m H) +``` + ## See also - [Cubes of natural numbers](elementary-number-theory.cubes-natural-numbers.md) +- [Squares of integers](elementary-number-theory.squares-integers.md) + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md index f44a8d6600..b2562fc4ee 100644 --- a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md @@ -35,17 +35,32 @@ open import univalent-combinatorics.standard-finite-types ## Idea -Consider a family of [natural numbers](elementary-number-theory.natural-numbers.md) $a : \mathsf{Fin}(n) → \mathbb{N}$ indexed by a [standard finite type](univalent-combinatorics.standard-finite-types.md) $\mathsf{Fin}(n)$. The {{#concept "sum" Disambiguation="natural numbers" Agda=sum-Fin-ℕ}} +Consider a family of +[natural numbers](elementary-number-theory.natural-numbers.md) +$a : \mathsf{Fin}(n) → \mathbb{N}$ indexed by a +[standard finite type](univalent-combinatorics.standard-finite-types.md) +$\mathsf{Fin}(n)$. The +{{#concept "sum" Disambiguation="natural numbers" Agda=sum-Fin-ℕ}} $$ \sum_{0\leq iImports + +```agda +open import elementary-number-theory.addition-integers +open import elementary-number-theory.difference-integers +open import elementary-number-theory.inequality-integers +open import elementary-number-theory.integers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.nonnegative-integers +open import elementary-number-theory.positive-and-negative-integers +open import elementary-number-theory.positive-integers +open import elementary-number-theory.well-ordering-principle-natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.transport-along-identifications +open import foundation.universe-levels +``` + +
+ +## Idea + +The [poset of integers](elementary-number-theory.inequality-integers.md) does not satisfy the [well-ordering principle of the natural numbers](elementary-number-theory.well-ordering-principle-natural-numbers.md). However, there are several useful approximations of the well-ordering principle for the [integers](elementary-number-theory.integers.md) that are worth recording: + +- For every family `A` of [decidable types](foundation.decidable-types.md) over `ℤ` that +is bounded from below and comes equipped with an integer `a` and an element of type `A a`, we can find a least integer `a₀` equipped with an element of type `A a₀`. Here, a type family `A` over the integers is bounded from below if it comes equipped with a [lower bound](order-theory.lower-bounds-posets.md) `b` such that `b ≤ x` for all `x : ℤ` equipped with an element of type `A x`. +- For every family `A` of [decidable types](foundation.decidable-types.md) over `ℤ` that +is bounded from above and comes equipped with an integer `a` and an element of type `A a`, we can find a largest integer `a₀` equipped with an element of type `A a₀`. Here, a type family `A` over the integers is bounded from above if it comes equipped with an [upper bound](order-theory.lower-bounds-posets.md) `b` such that `x ≤ b` for all `x : ℤ` equipped with an element of type `A x`. + +## Definitions + +### Minimal elements + +```agda +module _ + {l1 : Level} (A : ℤ → UU l1) + where + + minimal-element-ℤ : UU l1 + minimal-element-ℤ = Σ ℤ (λ a → A a × ((b : ℤ) → A b → a ≤-ℤ b)) +``` + +### Translations of minimal elements + +```agda +module _ + {l1 : Level} (A : ℤ → UU l1) (a : ℤ) + where + + translation-minimal-element-ℤ : + minimal-element-ℤ (λ x → A (x +ℤ a)) → minimal-element-ℤ A + translation-minimal-element-ℤ = + map-Σ + ( λ x → A x × ((y : ℤ) → A y → x ≤-ℤ y)) + ( λ x → x +ℤ a) + ( λ x (b , H) → + ( b , + λ y c → + inv-transpose-right-summand-leq-ℤ' x y a + ( H ( diff-ℤ y a) + ( tr A (inv (is-section-right-add-neg-ℤ a y)) c)))) +``` + +### Maximal elements + +```agda +module _ + {l1 : Level} (A : ℤ → UU l1) + where + + maximal-element-ℤ : UU l1 + maximal-element-ℤ = Σ ℤ (λ a → A a × ((b : ℤ) → A b → b ≤-ℤ a)) +``` + +### Reflections of minimal and maximal elements + +```agda +module _ + {l1 : Level} (A : ℤ → UU l1) + where + + reflect-minimal-element-ℤ : + minimal-element-ℤ (A ∘ neg-ℤ) → maximal-element-ℤ A + reflect-minimal-element-ℤ = + map-Σ + ( λ x → A x × ((y : ℤ) → A y → y ≤-ℤ x)) + ( neg-ℤ) + ( λ x (a , H) → + ( a , + λ y b → + transpose-right-neg-leq-ℤ y x + ( H (neg-ℤ y) (tr A (inv (neg-neg-ℤ y)) b)))) + + reflect-maximal-element-ℤ : + maximal-element-ℤ (A ∘ neg-ℤ) → minimal-element-ℤ A + reflect-maximal-element-ℤ = + map-Σ + ( λ x → A x × ((y : ℤ) → A y → x ≤-ℤ y)) + ( neg-ℤ) + λ x (a , H) → + ( a , + λ y b → + transpose-left-neg-leq-ℤ x y + ( H (neg-ℤ y) (tr A (inv (neg-neg-ℤ y)) b))) +``` + +### The well-ordering principle for the nonnegative integers + +The {{#concept "well-ordering principle for the nonnegative integers" Agda=well-ordering-principle-nonnegative-ℤ}} states that for any family `A` of decidable types over `ℤ` such that the implication `A x → is-nonnegative-ℤ x` holds for any `x : ℤ`, if we are given an element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an element of type `A a₀`. + +```agda +module _ + {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + where + + well-ordering-principle-nonnegative-ℤ' : + ((x : ℤ) → A x → is-nonnegative-ℤ x) → + minimal-element-ℕ (A ∘ int-ℕ) → minimal-element-ℤ A + pr1 (well-ordering-principle-nonnegative-ℤ' H (n , x , K)) = + int-ℕ n + pr1 (pr2 (well-ordering-principle-nonnegative-ℤ' H (n , x , K))) = + x + pr2 (pr2 (well-ordering-principle-nonnegative-ℤ' H (n , x , K))) b y = + concatenate-leq-eq-ℤ (int-ℕ n) (leq-int-ℕ n m (K m (tr A (inv α) y))) α + where + + m : ℕ + m = nat-nonnegative-ℤ (b , H b y) + + α : int-ℕ m = b + α = ap pr1 (is-section-nat-nonnegative-ℤ (b , H b y)) + + well-ordering-principle-nonnegative-ℤ : + ((x : ℤ) → A x → is-nonnegative-ℤ x) → Σ ℤ A → minimal-element-ℤ A + well-ordering-principle-nonnegative-ℤ H (a , x) = + well-ordering-principle-nonnegative-ℤ' H + ( well-ordering-principle-ℕ + ( A ∘ int-ℕ) + ( d ∘ int-ℕ) + ( map-Σ + ( A ∘ int-ℕ) + ( nat-nonnegative-ℤ) + ( λ x → id) + ( ( a , H a x) , + ( tr A + ( inv (ap pr1 (is-section-nat-nonnegative-ℤ (a , H a x)))) + ( x))))) +``` + +### The well-ordering principle for the positive integers + +The {{#concept "well-ordering principle for the positive integers" Agda=well-ordering-principle-positive-ℤ}} states that for any family `A` of decidable types over `ℤ` such that the implication `A x → is-positive-ℤ x` holds for any `x : ℤ`, if we are given an element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an element of type `A a₀`. + +```agda +module _ + {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + where + + well-ordering-principle-positive-ℤ : + ((x : ℤ) → A x → is-positive-ℤ x) → Σ ℤ A → minimal-element-ℤ A + well-ordering-principle-positive-ℤ H (a , x) = + well-ordering-principle-nonnegative-ℤ d + ( λ b y → is-nonnegative-is-positive-ℤ (H b y)) + ( a , x) +``` + +### The lower bounded well-ordering principle for the integers + +```agda +module _ + {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + where + + lower-bounded-well-ordering-principle-ℤ : + (b : ℤ) (H : (a : ℤ) → A a → b ≤-ℤ a) → Σ ℤ A → minimal-element-ℤ A + lower-bounded-well-ordering-principle-ℤ b H (a , x) = + translation-minimal-element-ℤ A b + ( well-ordering-principle-nonnegative-ℤ + ( λ c → d (c +ℤ b)) + ( λ c z → + tr + ( is-nonnegative-ℤ) + ( is-retraction-right-add-neg-ℤ b c) + ( H (c +ℤ b) z) ) + ( diff-ℤ a b , tr A (inv (is-section-right-add-neg-ℤ b a)) x)) +``` + +### The upper bounded well-ordering principle for the integers + +```agda +module _ + {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + where + + upper-bounded-well-ordering-principle-ℤ : + (b : ℤ) (H : (a : ℤ) → A a → a ≤-ℤ b) → Σ ℤ A → maximal-element-ℤ A + upper-bounded-well-ordering-principle-ℤ b H (a , x) = + reflect-minimal-element-ℤ A + ( lower-bounded-well-ordering-principle-ℤ + ( λ x → d (neg-ℤ x)) + ( neg-ℤ b) + ( λ x y → transpose-left-neg-leq-ℤ b x (H (neg-ℤ x) y)) + ( neg-ℤ a , tr A (inv (neg-neg-ℤ a)) x)) +``` diff --git a/src/linear-algebra/vectors.lagda.md b/src/linear-algebra/vectors.lagda.md index a870ba74c9..75ab5fbf56 100644 --- a/src/linear-algebra/vectors.lagda.md +++ b/src/linear-algebra/vectors.lagda.md @@ -29,6 +29,8 @@ open import foundation.unit-type open import foundation.universe-levels open import foundation.whiskering-higher-homotopies-composition +open import lists.lists + open import univalent-combinatorics.involution-standard-finite-types open import univalent-combinatorics.standard-finite-types ``` @@ -71,9 +73,9 @@ module _ revert-vec empty-vec = empty-vec revert-vec (x ∷ v) = snoc-vec (revert-vec v) x - all-vec : {l2 : Level} {n : ℕ} → (P : A → UU l2) → vec A n → UU l2 - all-vec P empty-vec = raise-unit _ - all-vec P (x ∷ v) = P x × all-vec P v + for-all-vec : {l2 : Level} {n : ℕ} → (P : A → UU l2) → vec A n → UU l2 + for-all-vec P empty-vec = raise-unit _ + for-all-vec P (x ∷ v) = P x × for-all-vec P v component-vec : (n : ℕ) → vec A n → Fin n → A @@ -85,7 +87,7 @@ module _ is-head : {n : ℕ} (a : A) (l : vec A n) → a ∈-vec (a ∷ l) is-in-tail : {n : ℕ} (a x : A) (l : vec A n) → a ∈-vec l → a ∈-vec (x ∷ l) - index-in-vec : (n : ℕ) → (a : A) → (v : vec A n) → a ∈-vec v → Fin n + index-in-vec : (n : ℕ) (a : A) (v : vec A n) → a ∈-vec v → Fin n index-in-vec (succ-ℕ n) a (.a ∷ v) (is-head .a .v) = inr star index-in-vec (succ-ℕ n) a (x ∷ v) (is-in-tail .a .x .v I) = @@ -149,6 +151,16 @@ module _ eq-component-functional-vec-index-in-functional-vec n x v I = pr2 I ``` +### The definition of `fold-vec` + +```agda +fold-vec : + {l1 l2 : Level} {A : UU l1} {B : UU l2} (b : B) (μ : A → (B → B)) → + {n : ℕ} → vec A n → B +fold-vec b μ {0} _ = b +fold-vec b μ (a ∷ l) = μ a (fold-vec b μ l) +``` + ## Properties ### Characterizing equality of listed vectors @@ -436,3 +448,54 @@ compute-tr-vec : (x ∷ tr (vec A) (is-injective-succ-ℕ p) v) compute-tr-vec refl v x = refl ``` + +### Back and forth between vectors and lists + +```agda +module _ + {l : Level} {A : UU l} + where + + list-vec : (n : ℕ) → vec A n → list A + list-vec zero-ℕ _ = nil + list-vec (succ-ℕ n) (x ∷ l) = cons x (list-vec n l) + + vec-list : (l : list A) → vec A (length-list l) + vec-list nil = empty-vec + vec-list (cons x l) = x ∷ vec-list l + + is-section-vec-list : (λ l → list-vec (length-list l) (vec-list l)) ~ id + is-section-vec-list nil = refl + is-section-vec-list (cons x l) = ap (cons x) (is-section-vec-list l) + + is-retraction-vec-list : + ( λ (x : Σ ℕ (λ n → vec A n)) → + ( length-list (list-vec (pr1 x) (pr2 x)) , + vec-list (list-vec (pr1 x) (pr2 x)))) ~ + id + is-retraction-vec-list (zero-ℕ , empty-vec) = refl + is-retraction-vec-list (succ-ℕ n , (x ∷ v)) = + ap + ( λ v → succ-ℕ (pr1 v) , (x ∷ (pr2 v))) + ( is-retraction-vec-list (n , v)) +``` + +### Back and forth between functional vectors and lists + +```agda +module _ + {l : Level} {A : UU l} + where + + list-functional-vec : (n : ℕ) → functional-vec A n → list A + list-functional-vec zero-ℕ v = + nil + list-functional-vec (succ-ℕ n) v = + cons (v (inr star)) (list-functional-vec n (v ∘ inl)) + + functional-vec-list : (l : list A) → functional-vec A (length-list l) + functional-vec-list nil = + empty-functional-vec + functional-vec-list (cons x l) = + cons-functional-vec (length-list l) x (functional-vec-list l) +``` diff --git a/src/lists/arrays.lagda.md b/src/lists/arrays.lagda.md index 054ea39af4..d1c983031a 100644 --- a/src/lists/arrays.lagda.md +++ b/src/lists/arrays.lagda.md @@ -24,6 +24,7 @@ open import foundation.universe-levels open import linear-algebra.vectors +open import lists.elementhood-relation-lists open import lists.lists open import univalent-combinatorics.involution-standard-finite-types @@ -34,8 +35,11 @@ open import univalent-combinatorics.standard-finite-types ## Idea -An array is a pair of a natural number `n`, and a function from `Fin n` to `A`. -We show that arrays and lists are equivalent. +An {{#concept "array" Agda=array}} is a pair consisting of a [natural number](elementary-number-theory.natural-numbers.md) `n`, and a [functional vector](linear-algebra.vectors.md) of `n` elements of `A`. The concept of array is [equivalent](foundation-core.equivalences.md) to the concept of [list](lists.lists.md). + +## Definitions + +### Arrays ```agda array : {l : Level} → UU l → UU l @@ -84,16 +88,6 @@ module _ revert-array (n , t) = (n , λ k → t (opposite-Fin n k)) ``` -### The definition of `fold-vec` - -```agda -fold-vec : - {l1 l2 : Level} {A : UU l1} {B : UU l2} (b : B) (μ : A → (B → B)) → - {n : ℕ} → vec A n → B -fold-vec b μ {0} _ = b -fold-vec b μ (a ∷ l) = μ a (fold-vec b μ l) -``` - ## Properties ### The types of lists and arrays are equivalent @@ -125,6 +119,12 @@ module _ ap ( λ v → succ-ℕ (pr1 v) , (x ∷ (pr2 v))) ( is-retraction-vec-list (n , v)) +``` + +```agda +module _ + {l : Level} {A : UU l} + where list-array : array A → list A list-array (n , t) = list-vec n (listed-vec-functional-vec n t) diff --git a/src/lists/dependent-lists.lagda.md b/src/lists/dependent-lists.lagda.md new file mode 100644 index 0000000000..c90245f348 --- /dev/null +++ b/src/lists/dependent-lists.lagda.md @@ -0,0 +1,44 @@ +# Dependent lists + +```agda +module lists.dependent-lists where +``` + +
Imports + +```agda +open import foundation.universe-levels + +open import lists.lists +``` + +
+ +## Idea + +Consider a dependent type `B` over a type `A`, and let `l` be a [list](lists.lists.md)` of elements of `A`. A {{#concept "dependent list" Agda=dependent-list}} of elements of `B` over `l` consists of a choice of an element `B x` for every element of `l`. More formally, the type `dependent-list B l` of dependent lists is a data type with constructors + +```text + nil-dependent-list : dependent-list B nil + cons-dependent-list : B x → dependent-list B l → dependent-list B (cons x l) +``` + +Note that the type `dependent-list B l` is [equivalent](foundation-core.equivalences.md) to the [universal quantification](lists.universal-quantification-lists.md) over the elements of `l` at `B`. + +## Definitions + +### Dependent lists + +```agda +module _ + {l1 l2 : Level} {A : UU l2} (B : A → UU l2) + where + + data dependent-list : (l : list A) → UU (l1 ⊔ l2) + where + nil-dependent-list : + dependent-list nil + cons-dependent-list : + (a : A) (l : list A) → B a → dependent-list l → dependent-list (cons a l) +``` + diff --git a/src/lists/elementhood-relation-lists.lagda.md b/src/lists/elementhood-relation-lists.lagda.md new file mode 100644 index 0000000000..3c63b7fec3 --- /dev/null +++ b/src/lists/elementhood-relation-lists.lagda.md @@ -0,0 +1,246 @@ +# The elementhood relation on lists + +```agda +module lists.elementhood-relation-lists where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.commuting-squares-of-maps +open import foundation.commuting-triangles-of-maps +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.raising-universe-levels +open import foundation.retractions +open import foundation.sections +open import foundation.unit-type +open import foundation.universe-levels + +open import linear-algebra.vectors + +open import lists.lists + +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +Consider a [list](lists.lists.md) `l` of elements of a type `A`. We say that an element `a : A` is an {{#concept "element" Disambiguation="lists" Agda=_∈-list_}} of `l` if there is an element of type `a ∈-list l`, which is the data type with constructors + +```text + is-head : (a : A) (l : list A) → a ∈-list cons a l + is-in-tail : (a x : A) (l : list) → a ∈-list l → a ∈-list cons x l +``` + +Equivalently, the elementhood relation can be defined inductively by + +```text + a ∈-list nil := empty + a ∈-list (cons x l) := (a = x) + (a ∈-list l). +``` + +The elementhood relation is a special case of the [elementhood relation](trees.elementhood-relation-w-types.md) on [W-types](trees.w-types.md) and the [elementhood relation](trees.elementhood-relation-coalgebras-polynomial-endofunctors.md) of [coalgebras](trees.coalgebras-polynomial-endofunctors.md) of [polynomial endofunctors](trees.polynomial-endofunctors.md). + +## Definitions + +### The elementhood predicate on lists + +```agda +infix 6 _∈-list_ + +data _∈-list_ {l1 : Level} {A : UU l1} : A → list A → UU l1 where + is-head : (a : A) (l : list A) → a ∈-list cons a l + is-in-tail : (a x : A) (l : list A) → a ∈-list l → a ∈-list cons x l +``` + +### The recursive definition of the elementhood relation + +```agda +elementhood-list : {l1 : Level} {A : UU l1} → A → list A → UU l1 +elementhood-list a nil = raise-empty _ +elementhood-list a (cons x l) = (a = x) + (elementhood-list a l) +``` + +## Properties + +### The `nil` list has no elements + +```agda +module _ + {l1 : Level} {A : UU l1} + where + + is-empty-∈-nil-list : (x : A) → is-empty (x ∈-list nil) + is-empty-∈-nil-list x () +``` + +### The two definitions of the elementhood relation are equivalent + +```agda +module _ + {l1 : Level} {A : UU l1} + where + + map-compute-elementhood-list : + (a : A) (l : list A) → elementhood-list a l → a ∈-list l + map-compute-elementhood-list a nil (map-raise ()) + map-compute-elementhood-list a (cons .a l) (inl refl) = + is-head a l + map-compute-elementhood-list a (cons x l) (inr H) = + is-in-tail a x l (map-compute-elementhood-list a l H) + + map-inv-compute-elementhood-list : + (a : A) (l : list A) → a ∈-list l → elementhood-list a l + map-inv-compute-elementhood-list a .(cons a l) (is-head .a l) = + inl refl + map-inv-compute-elementhood-list a .(cons x l) (is-in-tail .a x l H) = + inr (map-inv-compute-elementhood-list a l H) + + is-section-map-inv-compute-elementhood-list : + (a : A) (l : list A) → + is-section + ( map-compute-elementhood-list a l) + ( map-inv-compute-elementhood-list a l) + is-section-map-inv-compute-elementhood-list a .(cons a l) (is-head .a l) = + refl + is-section-map-inv-compute-elementhood-list a ._ (is-in-tail .a x l H) = + ap (is-in-tail a x l) (is-section-map-inv-compute-elementhood-list a l H) + + is-retraction-map-inv-compute-elementhood-list : + (a : A) (l : list A) → + is-retraction + ( map-compute-elementhood-list a l) + ( map-inv-compute-elementhood-list a l) + is-retraction-map-inv-compute-elementhood-list a nil (map-raise ()) + is-retraction-map-inv-compute-elementhood-list a (cons .a l) (inl refl) = + refl + is-retraction-map-inv-compute-elementhood-list a (cons x l) (inr H) = + ap inr (is-retraction-map-inv-compute-elementhood-list a l H) + + is-equiv-map-compute-elementhood-list : + (a : A) (l : list A) → is-equiv (map-compute-elementhood-list a l) + is-equiv-map-compute-elementhood-list a l = + is-equiv-is-invertible + ( map-inv-compute-elementhood-list a l) + ( is-section-map-inv-compute-elementhood-list a l) + ( is-retraction-map-inv-compute-elementhood-list a l) + + compute-elementhood-list : + (a : A) (l : list A) → elementhood-list a l ≃ (a ∈-list l) + pr1 (compute-elementhood-list a l) = map-compute-elementhood-list a l + pr2 (compute-elementhood-list a l) = is-equiv-map-compute-elementhood-list a l +``` + +### The type of all elements of a list is equivalent to the standard finite type with the length of the list as its number of elements + +Furthermore, there is a commuting triangle + +```text + Fin (length l) + / \ + / \ vec-list l + ∨ ∨ + type-of-elements l ----> X + pr1 +``` + +```agda +module _ + {l1 : Level} {A : UU l1} + where + + type-of-elements-list : list A → UU l1 + type-of-elements-list l = Σ A (_∈-list l) + + element-type-of-elements-list : + (l : list A) → type-of-elements-list l → A + element-type-of-elements-list l = pr1 + + cons-type-of-elements-list : + (x : A) (l : list A) → + type-of-elements-list l → type-of-elements-list (cons x l) + cons-type-of-elements-list x l = + tot (λ z → is-in-tail z x l) + + map-compute-type-of-elements-list : + (l : list A) → Fin (length-list l) → type-of-elements-list l + map-compute-type-of-elements-list (cons x l) (inl i) = + cons-type-of-elements-list x l (map-compute-type-of-elements-list l i) + map-compute-type-of-elements-list (cons x l) (inr _) = + (x , is-head x l) + + coherence-square-cons-type-of-elements-list : + (x : A) (l : list A) → + coherence-square-maps + ( map-compute-type-of-elements-list l) + ( inl-Fin (length-list l)) + ( cons-type-of-elements-list x l) + ( map-compute-type-of-elements-list (cons x l)) + coherence-square-cons-type-of-elements-list x l a = refl + + map-inv-compute-type-of-elements-list : + (l : list A) → type-of-elements-list l → Fin (length-list l) + map-inv-compute-type-of-elements-list (cons x l) (.x , is-head .x .l) = + inr star + map-inv-compute-type-of-elements-list (cons x l) (a , is-in-tail .a .x .l H) = + inl (map-inv-compute-type-of-elements-list l (a , H)) + + is-section-map-inv-compute-type-of-elements-list : + (l : list A) → + is-section + ( map-compute-type-of-elements-list l) + ( map-inv-compute-type-of-elements-list l) + is-section-map-inv-compute-type-of-elements-list + ( cons x l) + ( .x , is-head .x .l) = + refl + is-section-map-inv-compute-type-of-elements-list + ( cons x l) + ( a , is-in-tail .a .x .l H) = + ap + ( cons-type-of-elements-list x l) + ( is-section-map-inv-compute-type-of-elements-list l (a , H)) + + is-retraction-map-inv-compute-type-of-elements-list : + (l : list A) → + is-retraction + ( map-compute-type-of-elements-list l) + ( map-inv-compute-type-of-elements-list l) + is-retraction-map-inv-compute-type-of-elements-list (cons x l) (inl i) = + ap inl (is-retraction-map-inv-compute-type-of-elements-list l i) + is-retraction-map-inv-compute-type-of-elements-list (cons x l) (inr star) = + refl + + is-equiv-map-compute-type-of-elements-list : + (l : list A) → is-equiv (map-compute-type-of-elements-list l) + is-equiv-map-compute-type-of-elements-list l = + is-equiv-is-invertible + ( map-inv-compute-type-of-elements-list l) + ( is-section-map-inv-compute-type-of-elements-list l) + ( is-retraction-map-inv-compute-type-of-elements-list l) + + compute-type-of-elements-list : + (l : list A) → Fin (length-list l) ≃ type-of-elements-list l + pr1 (compute-type-of-elements-list l) = + map-compute-type-of-elements-list l + pr2 (compute-type-of-elements-list l) = + is-equiv-map-compute-type-of-elements-list l + + triangle-compute-type-of-elements-list : + (l : list A) → + coherence-triangle-maps + {!functional-vec-list l!} + {!!} + {!!} + triangle-compute-type-of-elements-list = {!!} +``` diff --git a/src/lists/lists.lagda.md b/src/lists/lists.lagda.md index 3a019086f1..f0fc415772 100644 --- a/src/lists/lists.lagda.md +++ b/src/lists/lists.lagda.md @@ -98,16 +98,6 @@ length-list : {l : Level} {A : UU l} → list A → ℕ length-list = fold-list 0 (λ a → succ-ℕ) ``` -### The elementhood predicate on lists - -```agda -infix 6 _∈-list_ - -data _∈-list_ {l : Level} {A : UU l} : A → list A → UU l where - is-head : (a : A) (l : list A) → a ∈-list (cons a l) - is-in-tail : (a x : A) (l : list A) → a ∈-list l → a ∈-list (cons x l) -``` - ## Properties ### A list that uses cons is not nil @@ -391,3 +381,7 @@ is-equiv-map-algebra-list A = ( is-section-map-inv-algebra-list A) ( is-retraction-map-inv-algebra-list A) ``` + +## See also + +- [The elementhood relation on lists](lists.elementhood-relation-lists.md) diff --git a/src/lists/permutation-lists.lagda.md b/src/lists/permutation-lists.lagda.md index 26ca630086..857ca511c1 100644 --- a/src/lists/permutation-lists.lagda.md +++ b/src/lists/permutation-lists.lagda.md @@ -24,6 +24,7 @@ open import linear-algebra.functoriality-vectors open import linear-algebra.vectors open import lists.arrays +open import lists.elementhood-relation-lists open import lists.functoriality-lists open import lists.lists open import lists.permutation-vectors diff --git a/src/lists/universal-quantification-lists.lagda.md b/src/lists/universal-quantification-lists.lagda.md index 3e9381f2f4..6a27b991eb 100644 --- a/src/lists/universal-quantification-lists.lagda.md +++ b/src/lists/universal-quantification-lists.lagda.md @@ -8,13 +8,23 @@ module lists.universal-quantification-lists where ```agda open import foundation.cartesian-product-types +open import foundation.contractible-types open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.equality-cartesian-product-types +open import foundation.equivalences +open import foundation.function-extensionality open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.identity-types open import foundation.propositions +open import foundation.retractions +open import foundation.sections open import foundation.unit-type open import foundation.universe-levels open import lists.concatenation-lists +open import lists.elementhood-relation-lists open import lists.lists open import lists.permutation-lists @@ -40,6 +50,8 @@ Alternatively, the universal quantification over the elements of a list can be d These definitions are [equivalent](foundation-core.equivalences.md). However, note that the inductive definition of the universal quantification has the same universe level as `P`, while the direct definition is of universe level `l1 ⊔ l2`, where `l1` is the universe level of `X` and `l2` is the universe level of `P`. +Note that the universal quantification over the elements of a list `l` is equivalent to the type of [dependent lists](lists.dependent-lists.md) of elements of `P` over `l`. + ## Definitions ### Universal quantification over the elements of a list @@ -52,7 +64,7 @@ module _ for-all-list : (l : list X) (P : X → UU l2) → UU l2 for-all-list nil P = raise-unit l2 - for-all-list (cons x l) P = (P x) × (for-all-list l P) + for-all-list (cons x l) P = P x × for-all-list l P for-all-nil-list : (P : X → UU l2) → for-all-list nil P @@ -68,6 +80,8 @@ module _ tail-for-all-cons-list x l P H = pr2 H ``` +### Universal quantification by a predicate valued in propositions + ```agda module _ {l1 l2 : Level} {X : UU l1} @@ -95,6 +109,26 @@ module _ for-all-elements-list : UU (l1 ⊔ l2) for-all-elements-list = (x : X) → x ∈-list l → P x + +for-all-elements-nil-list : + {l1 l2 : Level} {X : UU l1} (P : X → UU l2) → for-all-elements-list nil P +for-all-elements-nil-list P x () + +for-all-elements-cons-list : + {l1 l2 : Level} {X : UU l1} (x : X) (l : list X) (P : X → UU l2) → + P x → for-all-elements-list l P → for-all-elements-list (cons x l) P +for-all-elements-cons-list x l P a b .x (is-head .x .l) = a +for-all-elements-cons-list x l P a b y (is-in-tail .y .x .l H) = b y H + +head-for-all-elements-cons-list : + {l1 l2 : Level} {X : UU l1} (x : X) (l : list X) (P : X → UU l2) → + for-all-elements-list (cons x l) P → P x +head-for-all-elements-cons-list x l P H = H x (is-head x l) + +tail-for-all-elements-cons-list : + {l1 l2 : Level} {X : UU l1} (x : X) (l : list X) (P : X → UU l2) → + for-all-elements-list (cons x l) P → for-all-elements-list l P +tail-for-all-elements-cons-list x l P H y K = H y (is-in-tail y x l K) ``` ## Properties @@ -102,7 +136,93 @@ module _ ### The two definitions of universal quantification over the elements of a list are equivalent ```agda +module _ + {l1 l2 : Level} {X : UU l1} + where + is-contr-for-all-elements-nil-list : + (P : X → UU l2) → is-contr (for-all-elements-list nil P) + pr1 (is-contr-for-all-elements-nil-list P) = + for-all-elements-nil-list P + pr2 (is-contr-for-all-elements-nil-list P) H = + eq-htpy (λ x → eq-htpy (λ K → ex-falso (is-empty-∈-nil-list x K))) + + compute-for-all-elements-nil-list : + (P : X → UU l2) → for-all-elements-list nil P ≃ raise-unit l2 + compute-for-all-elements-nil-list P = + equiv-is-contr + ( is-contr-for-all-elements-nil-list P) + ( is-contr-raise-unit) + + module _ + (x : X) (l : list X) (P : X → UU l2) + where + + map-compute-for-all-elements-cons-list : + P x × for-all-elements-list l P → for-all-elements-list (cons x l) P + map-compute-for-all-elements-cons-list (a , b) = + for-all-elements-cons-list x l P a b + + map-inv-compute-for-all-elements-cons-list : + for-all-elements-list (cons x l) P → P x × for-all-elements-list l P + pr1 (map-inv-compute-for-all-elements-cons-list H) = + head-for-all-elements-cons-list x l P H + pr2 (map-inv-compute-for-all-elements-cons-list H) = + tail-for-all-elements-cons-list x l P H + + is-section-map-inv-compute-for-all-elements-cons-list' : + (H : for-all-elements-list (cons x l) P) → + (y : X) (K : y ∈-list cons x l) → + map-compute-for-all-elements-cons-list + ( map-inv-compute-for-all-elements-cons-list H) + ( y) + ( K) = + H y K + is-section-map-inv-compute-for-all-elements-cons-list' H y (is-head .y .l) = + refl + is-section-map-inv-compute-for-all-elements-cons-list' H y + ( is-in-tail .y .x .l K) = + refl + + is-section-map-inv-compute-for-all-elements-cons-list : + is-section + map-compute-for-all-elements-cons-list + map-inv-compute-for-all-elements-cons-list + is-section-map-inv-compute-for-all-elements-cons-list H = + eq-htpy + ( λ y → + eq-htpy + ( is-section-map-inv-compute-for-all-elements-cons-list' H y)) + + is-retraction-map-inv-compute-for-all-elements-cons-list : + is-retraction + map-compute-for-all-elements-cons-list + map-inv-compute-for-all-elements-cons-list + is-retraction-map-inv-compute-for-all-elements-cons-list H = + refl + + is-equiv-map-compute-for-all-elements-cons-list : + is-equiv map-compute-for-all-elements-cons-list + is-equiv-map-compute-for-all-elements-cons-list = + is-equiv-is-invertible + map-inv-compute-for-all-elements-cons-list + is-section-map-inv-compute-for-all-elements-cons-list + is-retraction-map-inv-compute-for-all-elements-cons-list + + compute-for-all-elements-cons-list : + P x × for-all-elements-list l P ≃ for-all-elements-list (cons x l) P + pr1 compute-for-all-elements-cons-list = + map-compute-for-all-elements-cons-list + pr2 compute-for-all-elements-cons-list = + is-equiv-map-compute-for-all-elements-cons-list + + compute-for-all-list : + (l : list X) (P : X → UU l2) → + for-all-elements-list l P ≃ for-all-list l P + compute-for-all-list nil P = compute-for-all-elements-nil-list P + compute-for-all-list (cons x l) P = + ( equiv-product id-equiv (compute-for-all-list l P)) ∘e + ( inv-equiv (compute-for-all-elements-cons-list x l P)) ``` ### If a predicate implies another predicate, then its universal quantification over the elements of a list implies the universal-quantification of the other predicate diff --git a/src/metric-spaces/metric-space-of-rational-numbers.lagda.md b/src/metric-spaces/metric-space-of-rational-numbers.lagda.md index 8ad75841c4..bac11d40c4 100644 --- a/src/metric-spaces/metric-space-of-rational-numbers.lagda.md +++ b/src/metric-spaces/metric-space-of-rational-numbers.lagda.md @@ -217,7 +217,7 @@ module _ lower-neighborhood-leq-ℚ d (x +ℚ u) (x +ℚ v) → lower-neighborhood-leq-ℚ d u v reflects-lower-neighborhood-leq-add-ℚ = - ( reflects-leq-right-add-ℚ x v (u +ℚ rational-ℚ⁺ d)) ∘ + ( reflects-order-right-add-ℚ x v (u +ℚ rational-ℚ⁺ d)) ∘ ( tr (leq-ℚ (x +ℚ v)) (associative-add-ℚ x u (rational-ℚ⁺ d))) ``` diff --git a/src/ring-theory/products-ideals-rings.lagda.md b/src/ring-theory/products-ideals-rings.lagda.md index 73cb6e9253..b63eb9886c 100644 --- a/src/ring-theory/products-ideals-rings.lagda.md +++ b/src/ring-theory/products-ideals-rings.lagda.md @@ -113,13 +113,13 @@ module _ (L : ideal-Ring l5 A) where - preserves-leq-product-ideal-Ring : + preserves-order-product-ideal-Ring : leq-ideal-Ring A I J → leq-ideal-Ring A K L → leq-ideal-Ring A ( product-ideal-Ring A I K) ( product-ideal-Ring A J L) - preserves-leq-product-ideal-Ring p q = + preserves-order-product-ideal-Ring p q = is-product-product-ideal-Ring A I K ( product-ideal-Ring A J L) ( λ x y u v → @@ -134,21 +134,21 @@ module _ (K : ideal-Ring l4 A) where - preserves-leq-left-product-ideal-Ring : + preserves-order-left-product-ideal-Ring : leq-ideal-Ring A I J → leq-ideal-Ring A ( product-ideal-Ring A I K) ( product-ideal-Ring A J K) - preserves-leq-left-product-ideal-Ring p = - preserves-leq-product-ideal-Ring A I J K K p + preserves-order-left-product-ideal-Ring p = + preserves-order-product-ideal-Ring A I J K K p ( refl-leq-ideal-Ring A K) - preserves-leq-right-product-ideal-Ring : + preserves-order-right-product-ideal-Ring : leq-ideal-Ring A J K → leq-ideal-Ring A ( product-ideal-Ring A I J) ( product-ideal-Ring A I K) - preserves-leq-right-product-ideal-Ring = - preserves-leq-product-ideal-Ring A I I J K + preserves-order-right-product-ideal-Ring = + preserves-order-product-ideal-Ring A I I J K ( refl-leq-ideal-Ring A I) ``` diff --git a/src/ring-theory/products-left-ideals-rings.lagda.md b/src/ring-theory/products-left-ideals-rings.lagda.md index 553c0eacf9..e4655f1c4e 100644 --- a/src/ring-theory/products-left-ideals-rings.lagda.md +++ b/src/ring-theory/products-left-ideals-rings.lagda.md @@ -118,13 +118,13 @@ module _ (L : left-ideal-Ring l5 A) where - preserves-leq-product-left-ideal-Ring : + preserves-order-product-left-ideal-Ring : leq-left-ideal-Ring A I J → leq-left-ideal-Ring A K L → leq-left-ideal-Ring A ( product-left-ideal-Ring A I K) ( product-left-ideal-Ring A J L) - preserves-leq-product-left-ideal-Ring p q = + preserves-order-product-left-ideal-Ring p q = is-product-product-left-ideal-Ring A I K ( product-left-ideal-Ring A J L) ( λ x y u v → @@ -139,21 +139,21 @@ module _ (K : left-ideal-Ring l4 A) where - preserves-leq-left-product-left-ideal-Ring : + preserves-order-left-product-left-ideal-Ring : leq-left-ideal-Ring A I J → leq-left-ideal-Ring A ( product-left-ideal-Ring A I K) ( product-left-ideal-Ring A J K) - preserves-leq-left-product-left-ideal-Ring p = - preserves-leq-product-left-ideal-Ring A I J K K p + preserves-order-left-product-left-ideal-Ring p = + preserves-order-product-left-ideal-Ring A I J K K p ( refl-leq-left-ideal-Ring A K) - preserves-leq-right-product-left-ideal-Ring : + preserves-order-right-product-left-ideal-Ring : leq-left-ideal-Ring A J K → leq-left-ideal-Ring A ( product-left-ideal-Ring A I J) ( product-left-ideal-Ring A I K) - preserves-leq-right-product-left-ideal-Ring = - preserves-leq-product-left-ideal-Ring A I I J K + preserves-order-right-product-left-ideal-Ring = + preserves-order-product-left-ideal-Ring A I I J K ( refl-leq-left-ideal-Ring A I) ``` diff --git a/src/ring-theory/products-right-ideals-rings.lagda.md b/src/ring-theory/products-right-ideals-rings.lagda.md index 22cf1b4cfe..cf6227ccc1 100644 --- a/src/ring-theory/products-right-ideals-rings.lagda.md +++ b/src/ring-theory/products-right-ideals-rings.lagda.md @@ -118,13 +118,13 @@ module _ (L : right-ideal-Ring l5 A) where - preserves-leq-product-right-ideal-Ring : + preserves-order-product-right-ideal-Ring : leq-right-ideal-Ring A I J → leq-right-ideal-Ring A K L → leq-right-ideal-Ring A ( product-right-ideal-Ring A I K) ( product-right-ideal-Ring A J L) - preserves-leq-product-right-ideal-Ring p q = + preserves-order-product-right-ideal-Ring p q = is-product-product-right-ideal-Ring A I K ( product-right-ideal-Ring A J L) ( λ x y u v → @@ -139,21 +139,21 @@ module _ (K : right-ideal-Ring l4 A) where - preserves-leq-left-product-right-ideal-Ring : + preserves-order-left-product-right-ideal-Ring : leq-right-ideal-Ring A I J → leq-right-ideal-Ring A ( product-right-ideal-Ring A I K) ( product-right-ideal-Ring A J K) - preserves-leq-left-product-right-ideal-Ring p = - preserves-leq-product-right-ideal-Ring A I J K K p + preserves-order-left-product-right-ideal-Ring p = + preserves-order-product-right-ideal-Ring A I J K K p ( refl-leq-right-ideal-Ring A K) - preserves-leq-right-product-right-ideal-Ring : + preserves-order-right-product-right-ideal-Ring : leq-right-ideal-Ring A J K → leq-right-ideal-Ring A ( product-right-ideal-Ring A I J) ( product-right-ideal-Ring A I K) - preserves-leq-right-product-right-ideal-Ring = - preserves-leq-product-right-ideal-Ring A I I J K + preserves-order-right-product-right-ideal-Ring = + preserves-order-product-right-ideal-Ring A I I J K ( refl-leq-right-ideal-Ring A I) ``` diff --git a/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md b/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md index f1c5ab49c3..b30db00541 100644 --- a/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md +++ b/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md @@ -82,7 +82,7 @@ to the cycle prime decomposition of `n *ℕ m`. equiv-product-cycle-prime-decomposition-ℕ : (n m : ℕ) → (H : leq-ℕ 1 n) → (I : leq-ℕ 1 m) → ( cycle-prime-decomposition-ℕ n H × cycle-prime-decomposition-ℕ m I) ≃ - cycle-prime-decomposition-ℕ (n *ℕ m) (preserves-leq-mul-ℕ 1 n 1 m H I) + cycle-prime-decomposition-ℕ (n *ℕ m) (preserves-order-mul-ℕ 1 n 1 m H I) equiv-product-cycle-prime-decomposition-ℕ n m H I = ( ( equiv-eq ( ap @@ -98,7 +98,7 @@ equiv-product-cycle-prime-decomposition-ℕ n m H I = ( eq-is-contr' ( fundamental-theorem-arithmetic-list-ℕ ( n *ℕ m) - ( preserves-leq-mul-ℕ 1 n 1 m H I)) + ( preserves-order-mul-ℕ 1 n 1 m H I)) ( prime-decomposition-list-sort-concatenation-ℕ ( n) ( m) @@ -114,7 +114,7 @@ equiv-product-cycle-prime-decomposition-ℕ n m H I = I)) ( prime-decomposition-fundamental-theorem-arithmetic-list-ℕ (n *ℕ m) - ( preserves-leq-mul-ℕ 1 n 1 m H I))))))) ∘e + ( preserves-order-mul-ℕ 1 n 1 m H I))))))) ∘e ( equiv-eq ( ap ( iterated-product-lists) diff --git a/src/univalent-combinatorics/inequality-types-with-counting.lagda.md b/src/univalent-combinatorics/inequality-types-with-counting.lagda.md index 137bb6bd70..8e3e667f47 100644 --- a/src/univalent-combinatorics/inequality-types-with-counting.lagda.md +++ b/src/univalent-combinatorics/inequality-types-with-counting.lagda.md @@ -63,24 +63,24 @@ transitive-leq-count (pair k e) {x} {y} {z} = ( map-inv-equiv e y) ( map-inv-equiv-count (pair k e) z) -preserves-leq-equiv-count : +preserves-order-equiv-count : {l : Level} {X : UU l} (e : count X) {x y : Fin (number-of-elements-count e)} → leq-Fin (number-of-elements-count e) x y → leq-count e (map-equiv-count e x) (map-equiv-count e y) -preserves-leq-equiv-count e {x} {y} H = +preserves-order-equiv-count e {x} {y} H = concatenate-eq-leq-eq-Fin ( number-of-elements-count e) ( is-retraction-map-inv-equiv (equiv-count e) x) ( H) ( inv (is-retraction-map-inv-equiv (equiv-count e) y)) -reflects-leq-equiv-count : +reflects-order-equiv-count : {l : Level} {X : UU l} (e : count X) {x y : Fin (number-of-elements-count e)} → leq-count e (map-equiv-count e x) (map-equiv-count e y) → leq-Fin (number-of-elements-count e) x y -reflects-leq-equiv-count e {x} {y} H = +reflects-order-equiv-count e {x} {y} H = concatenate-eq-leq-eq-Fin ( number-of-elements-count e) ( inv (is-retraction-map-inv-equiv (equiv-count e) x)) diff --git a/src/univalent-combinatorics/standard-finite-types.lagda.md b/src/univalent-combinatorics/standard-finite-types.lagda.md index b581beeac5..8ee224d56a 100644 --- a/src/univalent-combinatorics/standard-finite-types.lagda.md +++ b/src/univalent-combinatorics/standard-finite-types.lagda.md @@ -247,7 +247,7 @@ upper-bound-nat-Fin : (k : ℕ) (x : Fin (succ-ℕ k)) → leq-ℕ (nat-Fin (succ-ℕ k) x) k upper-bound-nat-Fin zero-ℕ (inr star) = star upper-bound-nat-Fin (succ-ℕ k) (inl x) = - preserves-leq-succ-ℕ (nat-Fin (succ-ℕ k) x) k (upper-bound-nat-Fin k x) + leq-succ-leq-ℕ (nat-Fin (succ-ℕ k) x) k (upper-bound-nat-Fin k x) upper-bound-nat-Fin (succ-ℕ k) (inr star) = refl-leq-ℕ (succ-ℕ k) upper-bound-nat-Fin' : From 6270496e7b50271c9008d49748ed841419a4e6c1 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 3 Jan 2025 21:28:08 -0500 Subject: [PATCH 26/72] Nicomachus's Theorem --- references.bib | 10 + ...nded-divisibility-natural-numbers.lagda.md | 15 ++ .../divisibility-natural-numbers.lagda.md | 13 ++ .../inequality-integers.lagda.md | 189 ++++++++++++++++-- .../multiplication-natural-numbers.lagda.md | 2 + .../natural-numbers.lagda.md | 6 + .../nicomachuss-theorem.lagda.md | 128 ++++++++++++ .../parity-integers.lagda.md | 30 ++- .../parity-natural-numbers.lagda.md | 24 +-- .../positive-and-negative-integers.lagda.md | 2 +- .../positive-integers.lagda.md | 8 + .../pronic-numbers.lagda.md | 85 ++++++++ .../squares-natural-numbers.lagda.md | 54 ++++- .../triangular-numbers.lagda.md | 41 +++- ...well-ordering-principles-integers.lagda.md | 38 +++- src/literature/oeis.lagda.md | 18 ++ 16 files changed, 596 insertions(+), 67 deletions(-) create mode 100644 src/elementary-number-theory/nicomachuss-theorem.lagda.md create mode 100644 src/elementary-number-theory/pronic-numbers.lagda.md diff --git a/references.bib b/references.bib index 00398ca814..ef8570c226 100644 --- a/references.bib +++ b/references.bib @@ -354,6 +354,16 @@ @online{GGMS24 keywords = {Computer Science - Logic in Computer Science,Mathematics - Logic} } +@book{HW08, + title = {{An Introduction to the Theory of Numbers (6th edition)}}, + author = {Hardy, G. H. and Wright, Edward M. and Heath-Brown, D. R. and Silverman, Joseph H.}, + abstract = {{The sixth edition of the classic undergraduate text in elementary number theory includes a new chapter on elliptic curves and their role in the proof of Fermat's Last Theorem, a foreword by Andrew Wiles and extensively revised and updated end-of-chapter notes.}}, + isbn = {0199219869}, + keywords = {congruences, primitive roots, residue systems, instructional exposition}, + publisher = {Oxford University Press}, + year = 2008 +} + @article{KECA17, title = {{Notions of Anonymous Existence in {{Martin-L\"of}} Type Theory}}, author = {Nicolai Kraus and Martín Escardó and Thierry Coquand and Thorsten Altenkirch}, diff --git a/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md index 2663d2642c..6af203f470 100644 --- a/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md @@ -224,6 +224,21 @@ pr1 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-prop-bounded-div-ℕ k n pr2 (pr2 (bounded-div-ℕ-Decidable-Prop k n)) = is-decidable-bounded-div-ℕ k n ``` +### The quotient of two numbers by any two proofs bounded divisibility are the same + +```agda +compute-quotient-bounded-div-ℕ : + {m m' n n' : ℕ} (q : m = m') (p : n = n') + (H : bounded-div-ℕ m n) (K : bounded-div-ℕ m' n') → + quotient-bounded-div-ℕ m n H = quotient-bounded-div-ℕ m' n' K +compute-quotient-bounded-div-ℕ refl refl H K = + ap + ( quotient-bounded-div-ℕ _ _) + { H} + { K} + ( eq-is-prop (is-prop-bounded-div-ℕ _ _)) +``` + ### `0` is bounded divisible by any natural number ```agda diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 99dc51bc47..10c6c595f0 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -207,6 +207,19 @@ module _ ( bounded-div-quotient-bounded-div-ℕ m n (bounded-div-div-ℕ m n H)) ``` +### The quotient by divisibility of two equal dividends by two equal divisors are equal + +```agda +compute-quotient-div-ℕ : + {m m' n n' : ℕ} (q : m = m') (p : n = n') + (H : div-ℕ m n) (K : div-ℕ m' n') → + quotient-div-ℕ m n H = quotient-div-ℕ m' n' K +compute-quotient-div-ℕ q p H K = + compute-quotient-bounded-div-ℕ q p + ( bounded-div-div-ℕ _ _ H) + ( bounded-div-div-ℕ _ _ K) +``` + ### Concatenating equality and divisibility ```agda diff --git a/src/elementary-number-theory/inequality-integers.lagda.md b/src/elementary-number-theory/inequality-integers.lagda.md index 8f03265c56..50112a73ac 100644 --- a/src/elementary-number-theory/inequality-integers.lagda.md +++ b/src/elementary-number-theory/inequality-integers.lagda.md @@ -28,13 +28,20 @@ open import foundation.functoriality-coproduct-types open import foundation.identity-types open import foundation.negated-equality open import foundation.negation +open import foundation.propositional-truncations open import foundation.propositions open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels +open import order-theory.decidable-posets +open import order-theory.decidable-preorders +open import order-theory.decidable-total-preorders +open import order-theory.decidable-total-orders open import order-theory.posets open import order-theory.preorders +open import order-theory.total-preorders +open import order-theory.total-orders ``` @@ -49,6 +56,15 @@ defines the {{#concept "standard ordering" Disambiguation="integers" Agda=leq-ℤ}} on the integers. +Alternatively, the standard ordering can be defined as a `data` type by specifying the constructors + +```text + r : (x : ℤ) → x ≤ x + s : (x y : ℤ) → x ≤ y → x ≤ y + 1. +``` + +We will introduce both orderings and prove that they are equivalent. + ## Definition ### Inequality on the integers @@ -67,6 +83,29 @@ infix 30 _≤-ℤ_ _≤-ℤ_ = leq-ℤ ``` +### Inductive definition of inequality on the integers + +```agda +data + inductive-leq-ℤ : + ℤ → ℤ → UU lzero + where + + refl-inductive-leq-ℤ : + (a : ℤ) → inductive-leq-ℤ a a + + succ-inductive-leq-ℤ : + (a b : ℤ) → inductive-leq-ℤ a b → inductive-leq-ℤ a (succ-ℤ b) + +refl-inductive-leq-ℤ' : + (a b : ℤ) → a = b → inductive-leq-ℤ a b +refl-inductive-leq-ℤ' a .a refl = refl-inductive-leq-ℤ a + +succ-inductive-leq-ℤ' : + (a b c : ℤ) (p : succ-ℤ b = c) → inductive-leq-ℤ a b → inductive-leq-ℤ a c +succ-inductive-leq-ℤ' a b .(succ-ℤ b) refl H = succ-inductive-leq-ℤ a b H +``` + ## Properties ### Inequality on the integers is reflexive, antisymmetric and transitive @@ -116,19 +155,44 @@ linear-leq-ℤ x y = ( decide-is-negative-is-nonnegative-ℤ) ``` -### An integer is lesser than its successor +### The partially ordered set of integers ordered by inequality ```agda -succ-leq-ℤ : (k : ℤ) → leq-ℤ k (succ-ℤ k) -succ-leq-ℤ k = - is-nonnegative-eq-ℤ - ( inv - ( ( left-successor-law-add-ℤ k (neg-ℤ k)) ∙ - ( ap succ-ℤ (right-inverse-law-add-ℤ k)))) - ( star) +ℤ-Preorder : Preorder lzero lzero +pr1 ℤ-Preorder = ℤ +pr1 (pr2 ℤ-Preorder) = leq-ℤ-Prop +pr1 (pr2 (pr2 ℤ-Preorder)) = refl-leq-ℤ +pr2 (pr2 (pr2 ℤ-Preorder)) = transitive-leq-ℤ -leq-ℤ-succ-leq-ℤ : (k l : ℤ) → leq-ℤ k l → leq-ℤ k (succ-ℤ l) -leq-ℤ-succ-leq-ℤ k l = transitive-leq-ℤ k l (succ-ℤ l) (succ-leq-ℤ l) +ℤ-Decidable-Preorder : Decidable-Preorder lzero lzero +pr1 ℤ-Decidable-Preorder = ℤ-Preorder +pr2 ℤ-Decidable-Preorder = is-decidable-leq-ℤ + +ℤ-Poset : Poset lzero lzero +pr1 ℤ-Poset = ℤ-Preorder +pr2 ℤ-Poset x y = antisymmetric-leq-ℤ + +ℤ-Decidable-Poset : Decidable-Poset lzero lzero +pr1 ℤ-Decidable-Poset = ℤ-Poset +pr2 ℤ-Decidable-Poset = is-decidable-leq-ℤ + +ℤ-Total-Preorder : Total-Preorder lzero lzero +pr1 ℤ-Total-Preorder = ℤ-Preorder +pr2 ℤ-Total-Preorder x y = unit-trunc-Prop (linear-leq-ℤ x y) + +ℤ-Decidable-Total-Preorder : Decidable-Total-Preorder lzero lzero +pr1 ℤ-Decidable-Total-Preorder = ℤ-Preorder +pr1 (pr2 ℤ-Decidable-Total-Preorder) x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr2 (pr2 ℤ-Decidable-Total-Preorder) = is-decidable-leq-ℤ + +ℤ-Total-Order : Total-Order lzero lzero +pr1 ℤ-Total-Order = ℤ-Poset +pr2 ℤ-Total-Order x y = unit-trunc-Prop (linear-leq-ℤ x y) + +ℤ-Decidable-Total-Order : Decidable-Total-Order lzero lzero +pr1 ℤ-Decidable-Total-Order = ℤ-Poset +pr1 (pr2 ℤ-Decidable-Total-Order) x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr2 (pr2 ℤ-Decidable-Total-Order) = is-decidable-leq-ℤ ``` ### Chaining rules for equality and inequality @@ -147,6 +211,87 @@ concatenate-eq-leq-ℤ : concatenate-eq-leq-ℤ y refl H = H ``` +### An integer is lesser than its successor + +```agda +succ-leq-ℤ : (k : ℤ) → leq-ℤ k (succ-ℤ k) +succ-leq-ℤ k = + is-nonnegative-eq-ℤ + ( inv + ( ( left-successor-law-add-ℤ k (neg-ℤ k)) ∙ + ( ap succ-ℤ (right-inverse-law-add-ℤ k)))) + ( star) + +leq-succ-leq-ℤ : (k l : ℤ) → leq-ℤ k l → leq-ℤ k (succ-ℤ l) +leq-succ-leq-ℤ k l = transitive-leq-ℤ k l (succ-ℤ l) (succ-leq-ℤ l) +``` + +### The inductively defined inequality on the integers is valued in propositions + +```agda +contradiction-inductive-leq-ℤ : + (x y : ℤ) (p : x = succ-ℤ y) → ¬ (inductive-leq-ℤ x y) +contradiction-inductive-leq-ℤ x .x p (refl-inductive-leq-ℤ .x) = + has-no-fixed-points-succ-ℤ x (inv p) +contradiction-inductive-leq-ℤ x .(succ-ℤ b) p (succ-inductive-leq-ℤ .x b H) = {!!} + +all-elements-equal-inductive-leq-ℤ : + (x y z : ℤ) (p : y = z) (H : inductive-leq-ℤ x y) (K : inductive-leq-ℤ x z) → + tr (inductive-leq-ℤ x) p H = K +all-elements-equal-inductive-leq-ℤ x .x .x p + ( refl-inductive-leq-ℤ .x) + ( refl-inductive-leq-ℤ .x) = + ap + ( λ t → tr (inductive-leq-ℤ x) t (refl-inductive-leq-ℤ x)) + ( eq-is-prop (is-set-ℤ x x) {p} {refl}) +all-elements-equal-inductive-leq-ℤ x .x ._ p + ( refl-inductive-leq-ℤ .x) + ( succ-inductive-leq-ℤ .x b K) = {!!} +all-elements-equal-inductive-leq-ℤ x .(succ-ℤ b) .x p (succ-inductive-leq-ℤ .x b H) (refl-inductive-leq-ℤ .x) = {!!} +all-elements-equal-inductive-leq-ℤ x .(succ-ℤ b) .(succ-ℤ b₁) p (succ-inductive-leq-ℤ .x b H) (succ-inductive-leq-ℤ .x b₁ K) = {!!} +``` + +### Inequality on the integers is equivalent to the inductively defined inequality on the integers + +```agda +leq-inductive-leq-ℤ : + (a b : ℤ) → inductive-leq-ℤ a b → a ≤-ℤ b +leq-inductive-leq-ℤ a .a (refl-inductive-leq-ℤ .a) = + refl-leq-ℤ a +leq-inductive-leq-ℤ a .(succ-ℤ b) (succ-inductive-leq-ℤ .a b H) = + leq-succ-leq-ℤ a b (leq-inductive-leq-ℤ a b H) + +inductive-leq-leq-ℤ' : + (a b c : ℤ) (p : diff-ℤ b a = c) → is-nonnegative-ℤ c → inductive-leq-ℤ a b +inductive-leq-leq-ℤ' a b (inr (inl star)) p H = + refl-inductive-leq-ℤ' a b (inv (eq-diff-ℤ p)) +inductive-leq-leq-ℤ' a b (inr (inr zero-ℕ)) p star = + succ-inductive-leq-ℤ' a a b + ( is-left-add-one-succ-ℤ a ∙ + ap (_+ℤ a) (inv p) ∙ + is-section-right-add-neg-ℤ a b) + ( refl-inductive-leq-ℤ a) +inductive-leq-leq-ℤ' a b (inr (inr (succ-ℕ x))) p star = + succ-inductive-leq-ℤ' a + ( a +ℤ inr (inr x)) + ( b) + ( inv (right-successor-law-add-ℤ a (inr (inr x))) ∙ + ap (a +ℤ_) (inv p) ∙ + commutative-add-ℤ a (diff-ℤ b a) ∙ + is-section-right-add-neg-ℤ a b) + ( inductive-leq-leq-ℤ' a + ( a +ℤ inr (inr x)) + ( inr (inr x)) + ( commutative-add-ℤ (a +ℤ inr (inr x)) (neg-ℤ a) ∙ + is-retraction-left-add-neg-ℤ a (inr (inr x))) + ( star)) + +inductive-leq-leq-ℤ : + (a b : ℤ) → a ≤-ℤ b → inductive-leq-ℤ a b +inductive-leq-leq-ℤ a b = + inductive-leq-leq-ℤ' a b (diff-ℤ b a) refl +``` + ### Addition on the integers preserves inequality ```agda @@ -266,7 +411,9 @@ inv-transpose-right-summand-leq-ℤ' a b c H = ### The operation taking integers to their negatives reverses their order -**Proof.** If `a ≤ b`, then `b - a` is nonnegative. Note that `b - a = (-a) - (-b)`, which is therefore also nonnegative, implying that `-b ≤ -a`. +**Proof.** If `a ≤ b`, then `b - a` is nonnegative. Note that +`b - a = (-a) - (-b)`, which is therefore also nonnegative, implying that +`-b ≤ -a`. ```agda reverses-order-neg-ℤ : @@ -308,17 +455,6 @@ transpose-left-neg-leq-ℤ a b H = ( concatenate-leq-eq-ℤ (neg-ℤ b) H (inv (neg-neg-ℤ a))) ``` -### The partially ordered set of integers ordered by inequality - -```agda -ℤ-Preorder : Preorder lzero lzero -ℤ-Preorder = - (ℤ , leq-ℤ-Prop , refl-leq-ℤ , transitive-leq-ℤ) - -ℤ-Poset : Poset lzero lzero -ℤ-Poset = (ℤ-Preorder , λ x y → antisymmetric-leq-ℤ) -``` - ### An integer `x` is nonnegative if and only if `0 ≤ x` ```agda @@ -392,6 +528,15 @@ module _ ( I)) ``` +### A positive integer is greater than or equal to `0` + +```agda +leq-zero-is-positive-ℤ : + (x : ℤ) → is-positive-ℤ x → zero-ℤ ≤-ℤ x +leq-zero-is-positive-ℤ x H = + leq-zero-is-nonnegative-ℤ x (is-nonnegative-is-positive-ℤ H) +``` + ## See also - The decidable total order on the integers is defined in diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index cb7bd10022..e912ec5289 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -335,3 +335,5 @@ neq-mul-ℕ m n p = - [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) - [Cubes of natural numbers](elementary-number-theory.cubes-natural-numbers.md) +- [Pronic numbers](elementary-number-theory.pronic-numbers.md), i.e., numbers of + the form $n(n+1)$. diff --git a/src/elementary-number-theory/natural-numbers.lagda.md b/src/elementary-number-theory/natural-numbers.lagda.md index 34008bc237..602392102e 100644 --- a/src/elementary-number-theory/natural-numbers.lagda.md +++ b/src/elementary-number-theory/natural-numbers.lagda.md @@ -141,6 +141,12 @@ is-nonzero-two-ℕ () is-not-one-two-ℕ : is-not-one-ℕ 2 is-not-one-two-ℕ () + +is-nonzero-three-ℕ : is-nonzero-ℕ 3 +is-nonzero-three-ℕ () + +is-nonzero-four-ℕ : is-nonzero-ℕ 4 +is-nonzero-four-ℕ () ``` ## See also diff --git a/src/elementary-number-theory/nicomachuss-theorem.lagda.md b/src/elementary-number-theory/nicomachuss-theorem.lagda.md new file mode 100644 index 0000000000..c90b48f279 --- /dev/null +++ b/src/elementary-number-theory/nicomachuss-theorem.lagda.md @@ -0,0 +1,128 @@ +# Nicomachus's Theorem + +```agda +module elementary-number-theory.nicomachuss-theorem where +``` + +
Imports + +```agda +open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.cubes-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.pronic-numbers +open import elementary-number-theory.squares-natural-numbers +open import elementary-number-theory.sums-of-natural-numbers +open import elementary-number-theory.triangular-numbers + +open import foundation.action-on-identifications-functions +open import foundation.identity-types +open import foundation.transport-along-identifications + +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +{{#concept "Nicomachus's Theorem" Agda=nicomachuss-theorem-ℕ WDID=Q2197859 WD="squared triangular number"}} +asserts that the [sum](elementary-number-theory.sums-natural-numbers.md) from +`k = 0` to `n` of the +[cube](elemenatary-number-theory.cubes-natural-numbers.md) `k³` is +[equal to](foundation-core.identity-types.md) the +[square](elementary-number-theory.squares-natural-numbers.md) of the `n`th +[triangular number](elementary-number-theory.triangular-numbers.md), i.e., it +asserts that + +$$ + \sum_{k=0}^n k^3 = \left(\sum_{k=0}^n k \right)^2. +$$ + +This identity is a common exercise in introductory sections on mathematical +induction. As such, it can be found as exercise 2 in section 1.2 of +{{#cite "Andrews94}}, or exercise 2(c) of section 1.2 in +{{#cite "Leveque12volI"}}. + +**Proof.** Recall that the $n$th triangular number is $n(n+1)/2$. We will prove by induction that + +$$ + 4\sum_{k=0}^n k^3 = 4\left(\sum_{k=0}^n k \right)^2. +$$ + +This allows us to avoid carrying along division by 4 throughout the proof. In the base case we have + +$$ + 4\sum_{k=0}^0 k^3 = 4·0^3 = 4·0^2 = 4(0·1/2)^2. +$$ + +For the inductive step, assume that the identity holds for `n`. Then we have + +$$ + 4\sum_{k=0}^{n+1} k^3 = (n(n+1))^2 + 4(n+1)^3 = ((n+1)(n+2))^2, +$$ + +where the last step follows by algebra. This completes the proof. + +## Properties + +### The square of the pronic number of a successor + +We have `((n + 1) * (n + 2))² = (n * (n + 1))² + 4 * (n + 1)³ + +```agda +compute-square-pronic-number-succ-ℕ : + (n : ℕ) → + square-ℕ (pronic-number-ℕ (succ-ℕ n)) = + square-ℕ (pronic-number-ℕ n) +ℕ 4 *ℕ cube-ℕ (succ-ℕ n) +compute-square-pronic-number-succ-ℕ n = + ( ap square-ℕ (commutative-mul-ℕ (succ-ℕ n) (succ-ℕ (succ-ℕ n)))) ∙ + ( distributive-square-mul-ℕ (succ-ℕ (succ-ℕ n)) (succ-ℕ n)) ∙ + ( ap (_*ℕ (square-ℕ (succ-ℕ n))) (square-succ-succ-ℕ' n)) ∙ + ( right-distributive-mul-add-ℕ + ( square-ℕ n) + ( 4 *ℕ succ-ℕ n) + ( square-ℕ (succ-ℕ n))) ∙ + ( ap-add-ℕ + ( inv (distributive-square-mul-ℕ n (succ-ℕ n))) + ( associative-mul-ℕ 4 (succ-ℕ n) (square-ℕ (succ-ℕ n)) ∙ + ap (mul-ℕ 4) (commutative-mul-ℕ (succ-ℕ n) (square-ℕ (succ-ℕ n))))) +``` + +### Nicomachus's theorem + +```agda +nicomachuss-theorem-ℕ' : + (n : ℕ) → + 4 *ℕ sum-Fin-ℕ (succ-ℕ n) (λ i → cube-ℕ (nat-Fin (succ-ℕ n) i)) = + 4 *ℕ square-ℕ (triangular-number-ℕ n) +nicomachuss-theorem-ℕ' zero-ℕ = refl +nicomachuss-theorem-ℕ' (succ-ℕ n) = + ( left-distributive-mul-add-ℕ 4 + ( sum-Fin-ℕ (succ-ℕ n) (λ i → cube-ℕ (nat-Fin (succ-ℕ n) i))) + ( cube-ℕ (succ-ℕ n))) ∙ + ( ap + ( _+ℕ 4 *ℕ cube-ℕ (succ-ℕ n)) + ( ( nicomachuss-theorem-ℕ' n) ∙ + ( inv (distributive-square-mul-ℕ 2 (triangular-number-ℕ n))) ∙ + ( ap square-ℕ (compute-triangular-number-ℕ' n)))) ∙ + ( inv (compute-square-pronic-number-succ-ℕ n)) ∙ + ( ap square-ℕ (inv (compute-triangular-number-ℕ' (succ-ℕ n)))) ∙ + ( distributive-square-mul-ℕ 2 (triangular-number-ℕ (succ-ℕ n))) + +nicomachuss-theorem-ℕ : + (n : ℕ) → + sum-Fin-ℕ (succ-ℕ n) (λ i → cube-ℕ (nat-Fin (succ-ℕ n) i)) = + square-ℕ (triangular-number-ℕ n) +nicomachuss-theorem-ℕ n = + is-injective-left-mul-ℕ 4 + ( is-nonzero-four-ℕ) + ( nicomachuss-theorem-ℕ' n) +``` + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/parity-integers.lagda.md b/src/elementary-number-theory/parity-integers.lagda.md index a1916bcf2d..79a423c4ae 100644 --- a/src/elementary-number-theory/parity-integers.lagda.md +++ b/src/elementary-number-theory/parity-integers.lagda.md @@ -12,12 +12,15 @@ open import elementary-number-theory.addition-integers open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.difference-integers open import elementary-number-theory.divisibility-integers +open import elementary-number-theory.inequality-integers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-integers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.positive-and-negative-integers open import elementary-number-theory.unit-similarity-integers +open import foundation.cartesian-product-types open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types @@ -357,17 +360,26 @@ has-odd-expansion-neg-one-ℤ : has-odd-expansion-ℤ neg-one-ℤ pr1 has-odd-expansion-neg-one-ℤ = neg-one-ℤ pr2 has-odd-expansion-neg-one-ℤ = refl +is-even-integer-below-ℤ : + (a b : ℤ) → UU lzero +is-even-integer-below-ℤ a b = is-even-ℤ b × (b ≤-ℤ a) + +is-decidable-is-even-integer-below-ℤ : + (a b : ℤ) → is-decidable (is-even-integer-below-ℤ a b) +is-decidable-is-even-integer-below-ℤ a b = + is-decidable-product (is-decidable-is-even-ℤ b) (is-decidable-leq-ℤ b a) + +instance-even-integer-below-ℤ : + (a : ℤ) → Σ ℤ (is-even-integer-below-ℤ a) +instance-even-integer-below-ℤ (inl x) = {!!} +instance-even-integer-below-ℤ (inr (inl x)) = + ( zero-ℤ , is-even-zero-ℤ , refl-leq-ℤ zero-ℤ) +instance-even-integer-below-ℤ (inr (inr x)) = + ( zero-ℤ , is-even-zero-ℤ , leq-zero-is-positive-ℤ (inr (inr x)) star) + has-odd-expansion-is-odd-ℤ : (a : ℤ) → is-odd-ℤ a → has-odd-expansion-ℤ a -has-odd-expansion-is-odd-ℤ (inl zero-ℕ) H = - has-odd-expansion-neg-one-ℤ -has-odd-expansion-is-odd-ℤ (inl (succ-ℕ zero-ℕ)) H = - ex-falso (H (neg-one-ℤ , refl)) -has-odd-expansion-is-odd-ℤ (inl (succ-ℕ (succ-ℕ a))) H = - {!!} - where - t = has-odd-expansion-is-odd-ℤ (inl a) {!!} -has-odd-expansion-is-odd-ℤ (inr a) H = {!!} +has-odd-expansion-is-odd-ℤ a H = {!!} ``` ```text diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index d2a3e66d6a..e6d0913690 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -404,31 +404,11 @@ is-odd-add-succ-self-ℕ n = is-odd-succ-is-even-ℕ (n +ℕ n) (is-even-add-self-ℕ n) ``` -### The product `n(n + 1)` is even - -```agda -abstract - is-even-mul-succ-ℕ : - (n : ℕ) → is-even-ℕ (n *ℕ succ-ℕ n) - is-even-mul-succ-ℕ n - with is-even-or-is-even-succ-ℕ n - ... | inl H = - is-even-div-is-even-ℕ - ( n *ℕ succ-ℕ n) - ( n) - ( H) - ( succ-ℕ n , commutative-mul-ℕ (succ-ℕ n) n) - ... | inr H = - is-even-div-is-even-ℕ - ( n *ℕ succ-ℕ n) - ( succ-ℕ n) - ( H) - ( n , refl) -``` - ## See also Further laws of parity are proven in other files, e.g.: - [Parity of integers](elementary-number-theory.parity-integers.md) - [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) +- The fact that the pronic numbers $n(n+1)$ are even is proven in + [Pronic numbers](elementary-number-theory.pronic-numbers.md). diff --git a/src/elementary-number-theory/positive-and-negative-integers.lagda.md b/src/elementary-number-theory/positive-and-negative-integers.lagda.md index 844b95a8b8..5080475c92 100644 --- a/src/elementary-number-theory/positive-and-negative-integers.lagda.md +++ b/src/elementary-number-theory/positive-and-negative-integers.lagda.md @@ -172,7 +172,7 @@ is-nonpositive-neg-is-nonnegative-ℤ {inr (inl x)} H = H is-nonpositive-neg-is-nonnegative-ℤ {inr (inr x)} H = H neg-nonnegative-ℤ : nonnegative-ℤ → nonpositive-ℤ -neg-nonnegative-ℤ (x , H) = neg-ℤ x , is-nonpositive-neg-is-nonnegative-ℤ H +neg-nonnegative-ℤ (x , H) = (neg-ℤ x , is-nonpositive-neg-is-nonnegative-ℤ H) ``` #### The negative of a nonpositive integer is nonnegative diff --git a/src/elementary-number-theory/positive-integers.lagda.md b/src/elementary-number-theory/positive-integers.lagda.md index 5ac772e16d..122af48408 100644 --- a/src/elementary-number-theory/positive-integers.lagda.md +++ b/src/elementary-number-theory/positive-integers.lagda.md @@ -131,6 +131,14 @@ positive-ℤ-Set : Set lzero positive-ℤ-Set = positive-ℤ , is-set-positive-ℤ ``` +### Integers of the form `in-pos-ℤ n` are positive + +```agda +is-positive-in-pos-ℤ : + (n : ℕ) → is-positive-ℤ (in-pos-ℤ n) +is-positive-in-pos-ℤ n = star +``` + ### The successor of a positive integer is positive ```agda diff --git a/src/elementary-number-theory/pronic-numbers.lagda.md b/src/elementary-number-theory/pronic-numbers.lagda.md new file mode 100644 index 0000000000..3cd8ce2475 --- /dev/null +++ b/src/elementary-number-theory/pronic-numbers.lagda.md @@ -0,0 +1,85 @@ +# The pronic numbers + +```agda +module elementary-number-theory.pronic-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers + +open import foundation.dependent-pair-types +open import foundation.coproduct-types +open import foundation.identity-types +``` + +
+ +## Idea + +The {{#concept "pronic numbers" Agda=pronic-number-ℕ WD="pronic number" WDID=Q1486643}} are the natural numbers +of the form + +$$ + n(n+1). +$$ + +The sequence `0, 2, 6, 12, 20, 30, 42, ⋯` of pronic numbers is listed as +[A002378](https://oeis.org/A002378) in the OEIS {{#cite "OEIS"}}. The $n$th +pronic number is [even](elementary-number-theory.parity-natural-numbers.md) for +every $n$, and it is twice the $n$th +[triangular number](elementary-number-theory.triangular-numbers.md). + +## Definitions + +### The pronic numbers + +```agda +pronic-number-ℕ : ℕ → ℕ +pronic-number-ℕ n = n *ℕ succ-ℕ n +``` + +## Properties + +### The $n$th pronic number $n(n + 1)$ is even + +```agda +abstract + is-even-pronic-number-ℕ : + (n : ℕ) → is-even-ℕ (pronic-number-ℕ n) + is-even-pronic-number-ℕ n + with is-even-or-is-even-succ-ℕ n + ... | inl H = + is-even-div-is-even-ℕ + ( pronic-number-ℕ n) + ( n) + ( H) + ( succ-ℕ n , commutative-mul-ℕ (succ-ℕ n) n) + ... | inr H = + is-even-div-is-even-ℕ + ( pronic-number-ℕ n) + ( succ-ℕ n) + ( H) + ( n , refl) +``` + +### The $(n+1)$st pronic number + +We have `(n + 1) * (n + 2) = n * (n + 1) + 2 * (n + 1)` + +```agda +compute-pronic-number-succ-ℕ : + (n : ℕ) → pronic-number-ℕ (succ-ℕ n) = pronic-number-ℕ n +ℕ 2 *ℕ succ-ℕ n +compute-pronic-number-succ-ℕ n = + commutative-mul-ℕ (succ-ℕ n) (succ-ℕ (succ-ℕ n)) ∙ + right-distributive-mul-add-ℕ n 2 (succ-ℕ n) +``` + +## See also + +- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) +- [Nicomachus's Theorem](elementary-number-theory.nicomachuss-theorem.md) diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index d920c0faa4..47911558c6 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -8,6 +8,7 @@ module elementary-number-theory.squares-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.congruence-natural-numbers open import elementary-number-theory.decidable-types open import elementary-number-theory.distance-natural-numbers @@ -17,6 +18,7 @@ open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.pronic-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions @@ -126,6 +128,12 @@ square-succ-succ-ℕ n = ( add-ℕ' 4) ( ( inv (associative-add-ℕ (2 *ℕ n) (0 +ℕ n) n)) ∙ ( inv (ap (add-ℕ' n) (associative-add-ℕ (2 *ℕ n) 0 n))))) + +square-succ-succ-ℕ' : + (n : ℕ) → square-ℕ (succ-ℕ (succ-ℕ n)) = square-ℕ n +ℕ 4 *ℕ (n +ℕ 1) +square-succ-succ-ℕ' n = + square-succ-succ-ℕ n ∙ + ap (square-ℕ n +ℕ_) (inv (left-distributive-mul-add-ℕ 4 n 1)) ``` ### The square function is order preserving @@ -457,7 +465,7 @@ is-1-mod-8-square-has-odd-expansion-ℕ ._ (k , refl) = ( preserves-div-mul-ℕ 4 2 4 ( k *ℕ (k +ℕ 1)) ( refl-div-ℕ 4) - ( is-even-mul-succ-ℕ k)) + ( is-even-pronic-number-ℕ k)) is-1-mod-8-square-is-odd-ℕ : (n : ℕ) → is-odd-ℕ n → square-ℕ n ≡ 1 mod-ℕ 8 @@ -481,6 +489,50 @@ cong-8-square-odd-number-ℕ m n H K = ( is-1-mod-8-square-is-odd-ℕ m H) ``` +### Squaring preserves divisibility + +```agda +module _ + (m n : ℕ) (H : div-ℕ m n) + where + + quotient-div-square-ℕ : ℕ + quotient-div-square-ℕ = square-ℕ (quotient-div-ℕ m n H) + + upper-bound-quotient-div-square-ℕ : + quotient-div-square-ℕ ≤-ℕ square-ℕ n + upper-bound-quotient-div-square-ℕ = + preserves-order-square-ℕ + ( quotient-div-ℕ m n H) + ( n) + ( upper-bound-quotient-div-ℕ m n H) + + eq-quotient-div-square-ℕ : + quotient-div-square-ℕ *ℕ square-ℕ m = square-ℕ n + eq-quotient-div-square-ℕ = + inv (distributive-square-mul-ℕ (quotient-div-ℕ m n H) m) ∙ + ap square-ℕ (eq-quotient-div-ℕ m n H) + + bounded-div-square-ℕ : bounded-div-ℕ (square-ℕ m) (square-ℕ n) + pr1 bounded-div-square-ℕ = quotient-div-square-ℕ + pr1 (pr2 bounded-div-square-ℕ) = upper-bound-quotient-div-square-ℕ + pr2 (pr2 bounded-div-square-ℕ) = eq-quotient-div-square-ℕ + + preserves-div-square-ℕ : div-ℕ (square-ℕ m) (square-ℕ n) + pr1 preserves-div-square-ℕ = quotient-div-square-ℕ + pr2 preserves-div-square-ℕ = eq-quotient-div-square-ℕ + + compute-quotient-div-square-ℕ : + (K : div-ℕ (square-ℕ m) (square-ℕ n)) → + quotient-div-ℕ (square-ℕ m) (square-ℕ n) K = quotient-div-square-ℕ + compute-quotient-div-square-ℕ K = + compute-quotient-bounded-div-ℕ + ( refl) + ( refl) + ( bounded-div-div-ℕ (square-ℕ m) (square-ℕ n) K) + ( bounded-div-square-ℕ) +``` + ## See also - [Cubes of natural numbers](elementary-number-theory.cubes-natural-numbers.md) diff --git a/src/elementary-number-theory/triangular-numbers.lagda.md b/src/elementary-number-theory/triangular-numbers.lagda.md index 41c27c4114..a9997e4d62 100644 --- a/src/elementary-number-theory/triangular-numbers.lagda.md +++ b/src/elementary-number-theory/triangular-numbers.lagda.md @@ -8,9 +8,11 @@ module elementary-number-theory.triangular-numbers where ```agda open import elementary-number-theory.addition-natural-numbers -open import elementary-number-theory.euclidean-division-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.pronic-numbers open import elementary-number-theory.sums-of-natural-numbers open import foundation.action-on-identifications-functions @@ -21,19 +23,26 @@ open import foundation.identity-types ## Idea -The $n$th {{#concept "triangular number" Agda=triangular-number-ℕ}} is the [sum](elementary-number-theory.sums-of-natural-numbers.md) of the [natural numbers](elementary-number-theory.natural-numbers.md) up to $n$. Alternatively, the triangular numbers can be defined inductively by +The $n$th {{#concept "triangular number" Agda=triangular-number-ℕ}} is the +[sum](elementary-number-theory.sums-of-natural-numbers.md) of the +[natural numbers](elementary-number-theory.natural-numbers.md) up to $n$. +Alternatively, the triangular numbers can be defined inductively by ```text T 0 := 0 T (n + 1) := T n + (n + 1). ``` -A classic inductive proof, which is often used as the first example when introducing students to mathematical induction, establishes the famous fact that +A classic inductive proof, which is often used as the first example when +introducing students to mathematical induction, establishes the famous fact that ```text T n = n(n + 1)/2. ``` +In other words, the $n$th triangular number is half the $n$th +[pronic number](elementary-number-theory.pronic-numbers.md). + ## Definition ### The sum definition of the triangular numbers @@ -72,15 +81,22 @@ $$ 2 \cdot \left(\sum_{i\leq n} i\right) = n(n+1). $$ -We prove this claim to avoid an early division by two, and the proof is by induction on $n$. In the base case both sides of the equality are $0$. For the inductive step, assume that $2 \cdot \sum_{i\leq n} i = n(n+1)$. Then we can compute +We prove this claim to avoid an early division by two, and the proof is by +induction on $n$. In the base case both sides of the equality are $0$. For the +inductive step, assume that $2 \cdot \sum_{i\leq n} i = n(n+1)$. Then we can +compute $$ 2 \cdot \sum_{i\leq n+1} i = 2 \cdot \left(\sum_{i\leq n} i\right)+ 2(n+1) = n(n+1) + ((n+1) + (n+1)) = (n+1)(n+2). $$ ```agda +value-triangular-number-ℕ : ℕ → ℕ +value-triangular-number-ℕ n = + quotient-div-ℕ 2 (pronic-number-ℕ n) (is-even-pronic-number-ℕ n) + compute-triangular-number-ℕ' : - (n : ℕ) → 2 *ℕ triangular-number-ℕ n = n *ℕ succ-ℕ n + (n : ℕ) → 2 *ℕ triangular-number-ℕ n = pronic-number-ℕ n compute-triangular-number-ℕ' zero-ℕ = refl compute-triangular-number-ℕ' (succ-ℕ n) = left-distributive-mul-add-ℕ 2 (triangular-number-ℕ n) (succ-ℕ n) ∙ @@ -90,7 +106,16 @@ compute-triangular-number-ℕ' (succ-ℕ n) = compute-triangular-number-ℕ : (n : ℕ) → - triangular-number-ℕ n = quotient-euclidean-division-ℕ 2 (n *ℕ succ-ℕ n) -compute-triangular-number-ℕ zero-ℕ = refl -compute-triangular-number-ℕ (succ-ℕ n) = {!!} + triangular-number-ℕ n = + quotient-div-ℕ 2 (pronic-number-ℕ n) (is-even-pronic-number-ℕ n) +compute-triangular-number-ℕ n = + is-injective-left-mul-ℕ 2 + ( is-nonzero-two-ℕ) + ( compute-triangular-number-ℕ' n ∙ + inv (eq-quotient-div-ℕ 2 (n *ℕ succ-ℕ n) (is-even-pronic-number-ℕ n)) ∙ + commutative-mul-ℕ (value-triangular-number-ℕ n) 2) ``` + +## See also + +- [Nicomachus's Theorem](elementary-number-theory.nicomachuss-theorem.md) diff --git a/src/elementary-number-theory/well-ordering-principles-integers.lagda.md b/src/elementary-number-theory/well-ordering-principles-integers.lagda.md index 49e12b75be..6c8b215669 100644 --- a/src/elementary-number-theory/well-ordering-principles-integers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principles-integers.lagda.md @@ -182,11 +182,12 @@ module _ ```agda module _ {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + (b : ℤ) (H : (a : ℤ) → A a → b ≤-ℤ a) ((a , x) : Σ ℤ A) where lower-bounded-well-ordering-principle-ℤ : - (b : ℤ) (H : (a : ℤ) → A a → b ≤-ℤ a) → Σ ℤ A → minimal-element-ℤ A - lower-bounded-well-ordering-principle-ℤ b H (a , x) = + minimal-element-ℤ A + lower-bounded-well-ordering-principle-ℤ = translation-minimal-element-ℤ A b ( well-ordering-principle-nonnegative-ℤ ( λ c → d (c +ℤ b)) @@ -196,6 +197,20 @@ module _ ( is-retraction-right-add-neg-ℤ b c) ( H (c +ℤ b) z) ) ( diff-ℤ a b , tr A (inv (is-section-right-add-neg-ℤ b a)) x)) + + integer-lower-bounded-well-ordering-principle-ℤ : ℤ + integer-lower-bounded-well-ordering-principle-ℤ = + pr1 lower-bounded-well-ordering-principle-ℤ + + structure-lower-bounded-well-ordering-principle-ℤ : + A integer-lower-bounded-well-ordering-principle-ℤ + structure-lower-bounded-well-ordering-principle-ℤ = + pr1 (pr2 lower-bounded-well-ordering-principle-ℤ) + + is-least-element-integer-lower-bounded-well-ordering-principle-ℤ : + (a : ℤ) → A a → integer-lower-bounded-well-ordering-principle-ℤ ≤-ℤ a + is-least-element-integer-lower-bounded-well-ordering-principle-ℤ = + pr2 (pr2 lower-bounded-well-ordering-principle-ℤ) ``` ### The upper bounded well-ordering principle for the integers @@ -203,15 +218,30 @@ module _ ```agda module _ {l1 : Level} {A : ℤ → UU l1} (d : (x : ℤ) → is-decidable (A x)) + (b : ℤ) (H : (a : ℤ) → A a → a ≤-ℤ b) ((a , x) : Σ ℤ A) where upper-bounded-well-ordering-principle-ℤ : - (b : ℤ) (H : (a : ℤ) → A a → a ≤-ℤ b) → Σ ℤ A → maximal-element-ℤ A - upper-bounded-well-ordering-principle-ℤ b H (a , x) = + maximal-element-ℤ A + upper-bounded-well-ordering-principle-ℤ = reflect-minimal-element-ℤ A ( lower-bounded-well-ordering-principle-ℤ ( λ x → d (neg-ℤ x)) ( neg-ℤ b) ( λ x y → transpose-left-neg-leq-ℤ b x (H (neg-ℤ x) y)) ( neg-ℤ a , tr A (inv (neg-neg-ℤ a)) x)) + + integer-upper-bounded-well-ordering-principle-ℤ : ℤ + integer-upper-bounded-well-ordering-principle-ℤ = + pr1 upper-bounded-well-ordering-principle-ℤ + + structure-upper-bounded-well-ordering-principle-ℤ : + A integer-upper-bounded-well-ordering-principle-ℤ + structure-upper-bounded-well-ordering-principle-ℤ = + pr1 (pr2 upper-bounded-well-ordering-principle-ℤ) + + is-largest-element-integer-upper-bounded-well-ordering-principle-ℤ : + (a : ℤ) → A a → a ≤-ℤ integer-upper-bounded-well-ordering-principle-ℤ + is-largest-element-integer-upper-bounded-well-ordering-principle-ℤ = + pr2 (pr2 upper-bounded-well-ordering-principle-ℤ) ``` diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index ba0c1f6f3f..2d17a51a26 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -147,6 +147,15 @@ open import elementary-number-theory.fermat-numbers using OEIS: [A000215](https://oeis.org/A000215) +### A000217 -- The triangular numbers + +```agda +open import elementary-number-theory.triangular-numbers using + ( triangular-number-ℕ) +``` + +OEIS: [A000217](https://oeis.org/A000217) + ### A000244 -- Powers of `3` ```agda @@ -201,6 +210,15 @@ A001477 = id OEIS: [A001477](https://oeis.org/A001477) +### A002378 -- The pronic numbers + +```agda +open import elementary-number-theory.pronic-numbers using + ( pronic-number-ℕ) +``` + +OEIS: [A002378](https://oeis.org/A002378) + ### A003090 -- The number of main classes of Latin squares of order `n` ```agda From 405c30a4c78f0f3ca28bc1d6499ae95a93b68d6e Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 4 Jan 2025 14:01:30 -0500 Subject: [PATCH 27/72] square pyramidal numbers --- .../multiplication-natural-numbers.lagda.md | 18 ++ .../pronic-numbers.lagda.md | 6 +- .../square-pyramidal-numbers.lagda.md | 170 ++++++++++++++++++ .../squares-natural-numbers.lagda.md | 1 + src/literature/oeis.lagda.md | 9 + 5 files changed, 202 insertions(+), 2 deletions(-) create mode 100644 src/elementary-number-theory/square-pyramidal-numbers.lagda.md diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index e912ec5289..0cc4824158 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -230,6 +230,24 @@ interchange-law-mul-mul-ℕ = associative-mul-ℕ ``` +### Swapping the order of multiplication + +```agda +left-swap-mul-ℕ : + (m n x : ℕ) → m *ℕ (n *ℕ x) = n *ℕ (m *ℕ x) +left-swap-mul-ℕ m n x = + ( inv (associative-mul-ℕ m n x)) ∙ + ( ap (_*ℕ x) (commutative-mul-ℕ m n)) ∙ + ( associative-mul-ℕ n m x) + +right-swap-mul-ℕ : + (x m n : ℕ) → (x *ℕ m) *ℕ n = (x *ℕ n) *ℕ m +right-swap-mul-ℕ x m n = + ( associative-mul-ℕ x m n) ∙ + ( ap (x *ℕ_) (commutative-mul-ℕ m n)) ∙ + ( inv (associative-mul-ℕ x n m)) +``` + ### Multiplication by a nonzero natural number is injective ```agda diff --git a/src/elementary-number-theory/pronic-numbers.lagda.md b/src/elementary-number-theory/pronic-numbers.lagda.md index 3cd8ce2475..9763698627 100644 --- a/src/elementary-number-theory/pronic-numbers.lagda.md +++ b/src/elementary-number-theory/pronic-numbers.lagda.md @@ -21,7 +21,7 @@ open import foundation.identity-types ## Idea -The {{#concept "pronic numbers" Agda=pronic-number-ℕ WD="pronic number" WDID=Q1486643}} are the natural numbers +The {{#concept "pronic numbers" Agda=pronic-number-ℕ WD="pronic number" WDID=Q1486643}} are the [natural numbers](elementary-number-theory.natural-numbers.md) of the form $$ @@ -81,5 +81,7 @@ compute-pronic-number-succ-ℕ n = ## See also -- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) - [Nicomachus's Theorem](elementary-number-theory.nicomachuss-theorem.md) +- [Square pyramidal numbers](elementary-number-theory.square-pyramidal-numbers.md) +- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) +- [Triangular numbers](elementary-number-theory.triangular-numbers.md) diff --git a/src/elementary-number-theory/square-pyramidal-numbers.lagda.md b/src/elementary-number-theory/square-pyramidal-numbers.lagda.md new file mode 100644 index 0000000000..dca47f71c6 --- /dev/null +++ b/src/elementary-number-theory/square-pyramidal-numbers.lagda.md @@ -0,0 +1,170 @@ +# The square pyramidal numbers + +```agda +module elementary-number-theory.square-pyramidal-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.pronic-numbers +open import elementary-number-theory.squares-natural-numbers +open import elementary-number-theory.sums-of-natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.identity-types +open import foundation.transport-along-identifications + +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +The $n$th {{#concept "square pyramidal number" Agda=square-pyramidal-number-ℕ OEIS=A000330 WDID=Q18949 WD="square pyramidal number"}} is the [sum](elementary-number-theory.sums-of-natural-numbers.md) of [squares](elementary-number-theory.squares-natural-numbers.md) + +$$ + 1^2 + 2^2 + \cdots + n^2. +$$ + +A common exercise in mathematical induction states that the $n$th square pyramidal number is equal to + +$$ + \frac{n(n+1)(2n+1)}{6}. +$$ + +The sequence of square pyramidal numbers is listed as A000330 in the [OEIS](literature.oeis.md) {{#cite "OEIS"}}. + +## Definitions + +### The square pyramidal numbers + +```agda +square-pyramidal-number-ℕ : ℕ → ℕ +square-pyramidal-number-ℕ n = + sum-Fin-ℕ (succ-ℕ n) (λ i → square-ℕ (nat-Fin (succ-ℕ n) i)) +``` + +## Properties + +### Computing the square pyramidal numbers + +We will prove the identity + +$$ + \sum_{k=0}^n k^2 = \frac{n(n+1)(2n+1)}{6}. +$$ + +In order to prove this, we first show that the sequence $a_n := n(n+1)(2n+1)$ satisfies the recurrence + +$$ + a_{n+1} = a(n) + 6(n+1)^2. +$$ + +By this recurrence it follows that each $a_n$ is [divisible](elementary-number-theory.divisibility-natural-numbers.md) by $6$, and that $a_n$ is $6$ times the $n$th pyramidal number. + +The computation of the square pyramidal numbers is a common exercise in elementary number theory texts, such as exercise 1 of section 1.2 in {{#cite "Andrews94}}. + +```agda +dividend-square-pyramidal-number-ℕ : ℕ → ℕ +dividend-square-pyramidal-number-ℕ n = pronic-number-ℕ n *ℕ odd-number-ℕ n + +dividend-square-pyramidal-number-succ-ℕ : + (n : ℕ) → + dividend-square-pyramidal-number-ℕ (succ-ℕ n) = + dividend-square-pyramidal-number-ℕ n +ℕ 6 *ℕ square-ℕ (succ-ℕ n) +dividend-square-pyramidal-number-succ-ℕ n = + ( associative-mul-ℕ + ( succ-ℕ n) + ( succ-ℕ (succ-ℕ n)) + ( odd-number-ℕ (succ-ℕ n))) ∙ + ( ap + ( succ-ℕ n *ℕ_) + ( ( ap + ( succ-ℕ (succ-ℕ n) *ℕ_) + ( ap succ-ℕ (left-distributive-mul-add-ℕ 2 n 1))) ∙ + ( left-distributive-mul-add-ℕ (succ-ℕ (succ-ℕ n)) (odd-number-ℕ n) 2) ∙ + ( ( ap-add-ℕ + ( right-distributive-mul-add-ℕ n 2 (odd-number-ℕ n)) + ( ( commutative-mul-ℕ (succ-ℕ (succ-ℕ n)) 2) ∙ + ( left-distributive-mul-add-ℕ 2 (succ-ℕ n) 1))) ∙ + ( associative-add-ℕ + ( n *ℕ odd-number-ℕ n) + ( 2 *ℕ odd-number-ℕ n) + ( 2 *ℕ succ-ℕ n +ℕ 2) ∙ + ( ap + ( n *ℕ odd-number-ℕ n +ℕ_) + ( ( left-swap-add-ℕ (2 *ℕ odd-number-ℕ n) (2 *ℕ succ-ℕ n) 2) ∙ + ( ap + ( 2 *ℕ succ-ℕ n +ℕ_) + ( inv (left-distributive-mul-add-ℕ 2 (odd-number-ℕ n) 1))) ∙ + ( ap + ( 2 *ℕ succ-ℕ n +ℕ_) + ( ( ap (2 *ℕ_) (inv (left-distributive-mul-add-ℕ 2 n 1))) ∙ + ( inv (associative-mul-ℕ 2 2 (succ-ℕ n))))) ∙ + ( inv (right-distributive-mul-add-ℕ 2 4 (succ-ℕ n))))))))) ∙ + ( left-distributive-mul-add-ℕ + ( succ-ℕ n) + ( n *ℕ odd-number-ℕ n) + ( 6 *ℕ succ-ℕ n)) ∙ + ( ap-add-ℕ + ( ( left-swap-mul-ℕ (succ-ℕ n) n (odd-number-ℕ n)) ∙ + ( inv (associative-mul-ℕ n (succ-ℕ n) (odd-number-ℕ n)))) + ( left-swap-mul-ℕ (succ-ℕ n) 6 (succ-ℕ n))) + +div-6-dividend-square-pyramidal-number-ℕ : + (n : ℕ) → div-ℕ 6 (dividend-square-pyramidal-number-ℕ n) +div-6-dividend-square-pyramidal-number-ℕ zero-ℕ = div-zero-ℕ 6 +div-6-dividend-square-pyramidal-number-ℕ (succ-ℕ n) = + tr + ( div-ℕ 6) + ( inv (dividend-square-pyramidal-number-succ-ℕ n)) + ( div-add-ℕ 6 + ( dividend-square-pyramidal-number-ℕ n) + ( 6 *ℕ square-ℕ (succ-ℕ n)) + ( div-6-dividend-square-pyramidal-number-ℕ n) + ( div-mul-ℕ' (square-ℕ (succ-ℕ n)) 6 6 (refl-div-ℕ 6))) + +value-square-pyramidal-number-ℕ : ℕ → ℕ +value-square-pyramidal-number-ℕ n = + quotient-div-ℕ 6 + ( dividend-square-pyramidal-number-ℕ n) + ( div-6-dividend-square-pyramidal-number-ℕ n) + +eq-value-square-pyramidal-number-ℕ : + (n : ℕ) → + 6 *ℕ value-square-pyramidal-number-ℕ n = dividend-square-pyramidal-number-ℕ n +eq-value-square-pyramidal-number-ℕ n = + eq-quotient-div-ℕ' 6 + ( dividend-square-pyramidal-number-ℕ n) + ( div-6-dividend-square-pyramidal-number-ℕ n) + +compute-square-pyramidal-number-ℕ' : + (n : ℕ) → + 6 *ℕ square-pyramidal-number-ℕ n = dividend-square-pyramidal-number-ℕ n +compute-square-pyramidal-number-ℕ' zero-ℕ = refl +compute-square-pyramidal-number-ℕ' (succ-ℕ n) = + ( left-distributive-mul-add-ℕ 6 + ( square-pyramidal-number-ℕ n) + ( square-ℕ (succ-ℕ n))) ∙ + ( ap (_+ℕ 6 *ℕ square-ℕ (succ-ℕ n)) (compute-square-pyramidal-number-ℕ' n)) ∙ + ( inv (dividend-square-pyramidal-number-succ-ℕ n)) + +compute-square-pyramidal-number-ℕ : + (n : ℕ) → square-pyramidal-number-ℕ n = value-square-pyramidal-number-ℕ n +compute-square-pyramidal-number-ℕ n = + is-injective-left-mul-ℕ 6 + ( is-nonzero-succ-ℕ 5) + ( ( compute-square-pyramidal-number-ℕ' n) ∙ + ( inv (eq-value-square-pyramidal-number-ℕ n))) +``` + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 47911558c6..16a847ec54 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -536,6 +536,7 @@ module _ ## See also - [Cubes of natural numbers](elementary-number-theory.cubes-natural-numbers.md) +- [Square pyramidal numbers](elementary-number-theory.square-pyramidal-numbers.md) - [Squares of integers](elementary-number-theory.squares-integers.md) ## References diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index 2d17a51a26..2778481c79 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -165,6 +165,15 @@ A000244 = exp-ℕ 3 OEIS: [A000244](https://oeis.org/A000244) +### A000330 -- The square pyramidal numbers + +```agda +open import elementary-number-theory.square-pyramidal-numbers using + ( square-pyramidal-number-ℕ) +``` + +OEIS: [A000330](https://oeis.org/A000330) + ### A000720 -- The prime counting function ```agda From adec8f0c5ae59fd58b330ad80ec349498f3d7edb Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 4 Jan 2025 14:48:56 -0500 Subject: [PATCH 28/72] sum of pronic numbers --- .../pronic-numbers.lagda.md | 110 ++++++++++++++++++ .../square-pyramidal-numbers.lagda.md | 2 +- 2 files changed, 111 insertions(+), 1 deletion(-) diff --git a/src/elementary-number-theory/pronic-numbers.lagda.md b/src/elementary-number-theory/pronic-numbers.lagda.md index 9763698627..2f1a1d0fdc 100644 --- a/src/elementary-number-theory/pronic-numbers.lagda.md +++ b/src/elementary-number-theory/pronic-numbers.lagda.md @@ -8,13 +8,19 @@ module elementary-number-theory.pronic-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers +open import elementary-number-theory.sums-of-natural-numbers +open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types open import foundation.coproduct-types open import foundation.identity-types +open import foundation.transport-along-identifications + +open import univalent-combinatorics.standard-finite-types ``` @@ -28,6 +34,12 @@ $$ n(n+1). $$ +The pronic numbers satisfy the recurrence + +$$ + a_{n+1} = a_n + 2(n+1). +$$ + The sequence `0, 2, 6, 12, 20, 30, 42, ⋯` of pronic numbers is listed as [A002378](https://oeis.org/A002378) in the OEIS {{#cite "OEIS"}}. The $n$th pronic number is [even](elementary-number-theory.parity-natural-numbers.md) for @@ -45,6 +57,16 @@ pronic-number-ℕ n = n *ℕ succ-ℕ n ## Properties +### The pronic number of a successor + +```agda +pronic-number-succ-ℕ : + (n : ℕ) → pronic-number-ℕ (succ-ℕ n) = pronic-number-ℕ n +ℕ 2 *ℕ succ-ℕ n +pronic-number-succ-ℕ n = + ( left-distributive-mul-add-ℕ (succ-ℕ n) n 2) ∙ + ( ap-add-ℕ (commutative-mul-ℕ (succ-ℕ n) n) (commutative-mul-ℕ (succ-ℕ n) 2)) +``` + ### The $n$th pronic number $n(n + 1)$ is even ```agda @@ -79,9 +101,97 @@ compute-pronic-number-succ-ℕ n = right-distributive-mul-add-ℕ n 2 (succ-ℕ n) ``` +### The sum of pronic numbers + +The sum of the pronic numbers from $1$ to $n$ is given by the identity + +$$ + \sum_{k=0}^n k(k+1) = \frac{n(n+1)(n+3)}{3} = 2\binom{n+3}{3} +$$ + +This sequence of numbers starts out as + +$$ + 0,\ 2,  8,\ 20,\ 40,\ 70,\ \ldots +$$ + +and it is listed as A007290 in the [OEIS](literature.oeis.md) {{#cite "OEIS"}}. The computation of the sum of pronic numbers is stated as exercise 4 in section 1.2 of {{#cite "Andrews94}}. + +```agda +sum-of-pronic-numbers-ℕ : ℕ → ℕ +sum-of-pronic-numbers-ℕ n = + sum-Fin-ℕ (succ-ℕ n) (λ i → pronic-number-ℕ (nat-Fin (succ-ℕ n) i)) + +dividend-sum-of-pronic-numbers-ℕ : ℕ → ℕ +dividend-sum-of-pronic-numbers-ℕ n = pronic-number-ℕ n *ℕ succ-ℕ (succ-ℕ n) + +dividend-sum-of-pronic-numbers-succ-ℕ : + (n : ℕ) → + dividend-sum-of-pronic-numbers-ℕ (succ-ℕ n) = + dividend-sum-of-pronic-numbers-ℕ n +ℕ 3 *ℕ pronic-number-ℕ (succ-ℕ n) +dividend-sum-of-pronic-numbers-succ-ℕ n = + ( commutative-mul-ℕ (pronic-number-ℕ (succ-ℕ n)) (n +ℕ 3)) ∙ + ( right-distributive-mul-add-ℕ n 3 (pronic-number-ℕ (succ-ℕ n))) ∙ + ( ap + ( _+ℕ 3 *ℕ pronic-number-ℕ (succ-ℕ n)) + ( inv (associative-mul-ℕ n (succ-ℕ n) (succ-ℕ (succ-ℕ n))))) + +div-3-dividend-sum-of-pronic-numbers-ℕ : + (n : ℕ) → div-ℕ 3 (dividend-sum-of-pronic-numbers-ℕ n) +div-3-dividend-sum-of-pronic-numbers-ℕ zero-ℕ = div-zero-ℕ 3 +div-3-dividend-sum-of-pronic-numbers-ℕ (succ-ℕ n) = + tr + ( div-ℕ 3) + ( inv (dividend-sum-of-pronic-numbers-succ-ℕ n)) + ( div-add-ℕ 3 + ( dividend-sum-of-pronic-numbers-ℕ n) + ( 3 *ℕ pronic-number-ℕ (succ-ℕ n)) + ( div-3-dividend-sum-of-pronic-numbers-ℕ n) + ( div-mul-ℕ' (pronic-number-ℕ (succ-ℕ n)) 3 3 (refl-div-ℕ 3))) + +value-sum-of-pronic-numbers-ℕ : ℕ → ℕ +value-sum-of-pronic-numbers-ℕ n = + quotient-div-ℕ 3 + ( dividend-sum-of-pronic-numbers-ℕ n) + ( div-3-dividend-sum-of-pronic-numbers-ℕ n) + +eq-value-sum-of-pronic-numbers-ℕ : + (n : ℕ) → + 3 *ℕ value-sum-of-pronic-numbers-ℕ n = dividend-sum-of-pronic-numbers-ℕ n +eq-value-sum-of-pronic-numbers-ℕ n = + eq-quotient-div-ℕ' 3 + ( dividend-sum-of-pronic-numbers-ℕ n) + ( div-3-dividend-sum-of-pronic-numbers-ℕ n) + +compute-sum-of-pronic-numbers-ℕ' : + (n : ℕ) → + 3 *ℕ sum-of-pronic-numbers-ℕ n = dividend-sum-of-pronic-numbers-ℕ n +compute-sum-of-pronic-numbers-ℕ' zero-ℕ = refl +compute-sum-of-pronic-numbers-ℕ' (succ-ℕ n) = + ( left-distributive-mul-add-ℕ 3 + ( sum-of-pronic-numbers-ℕ n) + ( pronic-number-ℕ (succ-ℕ n))) ∙ + ( ap + ( _+ℕ 3 *ℕ pronic-number-ℕ (succ-ℕ n)) + ( compute-sum-of-pronic-numbers-ℕ' n)) ∙ + ( inv (dividend-sum-of-pronic-numbers-succ-ℕ n)) + +compute-sum-of-pronic-numbers-ℕ : + (n : ℕ) → sum-of-pronic-numbers-ℕ n = value-sum-of-pronic-numbers-ℕ n +compute-sum-of-pronic-numbers-ℕ n = + is-injective-left-mul-ℕ 3 + ( is-nonzero-succ-ℕ 2) + ( ( compute-sum-of-pronic-numbers-ℕ' n) ∙ + ( inv (eq-value-sum-of-pronic-numbers-ℕ n))) +``` + ## See also - [Nicomachus's Theorem](elementary-number-theory.nicomachuss-theorem.md) - [Square pyramidal numbers](elementary-number-theory.square-pyramidal-numbers.md) - [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) - [Triangular numbers](elementary-number-theory.triangular-numbers.md) + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/square-pyramidal-numbers.lagda.md b/src/elementary-number-theory/square-pyramidal-numbers.lagda.md index dca47f71c6..f7d8cbe506 100644 --- a/src/elementary-number-theory/square-pyramidal-numbers.lagda.md +++ b/src/elementary-number-theory/square-pyramidal-numbers.lagda.md @@ -64,7 +64,7 @@ $$ In order to prove this, we first show that the sequence $a_n := n(n+1)(2n+1)$ satisfies the recurrence $$ - a_{n+1} = a(n) + 6(n+1)^2. + a_{n+1} = a_n + 6(n+1)^2. $$ By this recurrence it follows that each $a_n$ is [divisible](elementary-number-theory.divisibility-natural-numbers.md) by $6$, and that $a_n$ is $6$ times the $n$th pyramidal number. From 4f10aa57da02f0a63c4ba4341c77e1968555a2e7 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 4 Jan 2025 15:46:15 -0500 Subject: [PATCH 29/72] The sum of the first n odd numbers is the nth square --- .../parity-natural-numbers.lagda.md | 2 +- .../squares-natural-numbers.lagda.md | 29 +++++++++++++++++++ 2 files changed, 30 insertions(+), 1 deletion(-) diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index e6d0913690..3ea3e75065 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -409,6 +409,6 @@ is-odd-add-succ-self-ℕ n = Further laws of parity are proven in other files, e.g.: - [Parity of integers](elementary-number-theory.parity-integers.md) -- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md) +- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md). Here we also show that the sum of the first $n$ odd numbers is $n^2$. - The fact that the pronic numbers $n(n+1)$ are even is proven in [Pronic numbers](elementary-number-theory.pronic-numbers.md). diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 16a847ec54..5e37e47a06 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -20,6 +20,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.pronic-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import elementary-number-theory.sums-of-natural-numbers open import foundation.action-on-identifications-functions open import foundation.coproduct-types @@ -33,6 +34,8 @@ open import foundation.universe-levels open import foundation-core.cartesian-product-types open import foundation-core.transport-along-identifications + +open import univalent-combinatorics.standard-finite-types ``` @@ -533,6 +536,32 @@ module _ ( bounded-div-square-ℕ) ``` +### The $n$th square is the sum of the first $n$ odd numbers + +We show that + +$$ + 1 + 3 + \cdots + (2n+1) = n^2. +$$ + +This solves exercise 5 in section 1.2 of {{#cite "Andrews94"}}. + +```agda +sum-of-odd-numbers-ℕ : ℕ → ℕ +sum-of-odd-numbers-ℕ n = + sum-Fin-ℕ n (λ i → odd-number-ℕ (nat-Fin n i)) + +compute-sum-of-odd-numbers-ℕ : + (n : ℕ) → sum-of-odd-numbers-ℕ n = square-ℕ n +compute-sum-of-odd-numbers-ℕ zero-ℕ = refl +compute-sum-of-odd-numbers-ℕ (succ-ℕ zero-ℕ) = refl +compute-sum-of-odd-numbers-ℕ (succ-ℕ (succ-ℕ n)) = + ( ap + ( _+ℕ odd-number-ℕ (succ-ℕ n)) + ( compute-sum-of-odd-numbers-ℕ (succ-ℕ n))) ∙ + ( inv (square-succ-ℕ (succ-ℕ n))) +``` + ## See also - [Cubes of natural numbers](elementary-number-theory.cubes-natural-numbers.md) From 5dd2b237b25ef3363baa2d9df00b141ae33ad817 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 4 Jan 2025 18:34:16 -0500 Subject: [PATCH 30/72] number of divisors --- src/literature/oeis.lagda.md | 7 +++++++ .../counting-decidable-subtypes.lagda.md | 7 ++++++- 2 files changed, 13 insertions(+), 1 deletion(-) diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index 2778481c79..8a6a366206 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -45,6 +45,13 @@ A000004 : ℕ → ℕ A000004 _ = zero-ℕ ``` +### A000005 -- The number of divisors + +```agda +open import elementary-number-theory.number-of-divisors using + ( number-of-divisors-ℕ) +``` + OEIS: [A000004](https://oeis.org/A000004) ### A000007 -- The characteristic function for 0 diff --git a/src/univalent-combinatorics/counting-decidable-subtypes.lagda.md b/src/univalent-combinatorics/counting-decidable-subtypes.lagda.md index 2dc4eea133..5d6b13a671 100644 --- a/src/univalent-combinatorics/counting-decidable-subtypes.lagda.md +++ b/src/univalent-combinatorics/counting-decidable-subtypes.lagda.md @@ -106,11 +106,16 @@ abstract count-decidable-subtype : {l1 l2 : Level} {X : UU l1} (P : decidable-subtype l2 X) → - (count X) → count (type-decidable-subtype P) + count X → count (type-decidable-subtype P) count-decidable-subtype P e = count-decidable-subtype' P ( number-of-elements-count e) ( equiv-count e) + +number-of-elements-decidable-subtype : + {l1 l2 : Level} {X : UU l1} (P : decidable-subtype l2 X) (e : count X) → ℕ +number-of-elements-decidable-subtype P e = + number-of-elements-count (count-decidable-subtype P e) ``` ### The elements in the domain of a decidable embedding can be counted if the elements of the codomain can be counted From f94ad9d6fd3a37188378a64b6118a7eaf94dad49 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 4 Jan 2025 19:52:57 -0500 Subject: [PATCH 31/72] LeVeque's strict bound on the growth of the Fibonacci numbers --- .../fibonacci-sequence.lagda.md | 86 +++++++++++++++++++ ...strict-inequality-natural-numbers.lagda.md | 66 ++++++++++++++ 2 files changed, 152 insertions(+) diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index 89af95e039..d1546a51f7 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -8,16 +8,19 @@ module elementary-number-theory.fibonacci-sequence where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.greatest-common-divisor-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.relatively-prime-natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.transport-along-identifications +open import foundation.unit-type ``` @@ -226,8 +229,91 @@ preserves-div-Fibonacci-ℕ m n H = div-Fibonacci-div-ℕ (Fibonacci-ℕ m) m n H (refl-div-ℕ (Fibonacci-ℕ m)) ``` +### LeVeque's strict bound on the growth of the Fibonacci numbers + +The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This claim appears on pages 7 and 8 in section 1.2 of {{#cite "Leveque12volI"}} as an example of a proof by induction. + +**Proof.** We will show by induction that $4^n F(n) < 7^n$. In the base case the claim reduces to the strict inequality $0 < 1$, which is true. Furthermore, we have that $4 F(1) = 4 < 7$. For the inductive step, assume that $4^n F(n) < 7^n$ and that $4^{n+1} F(n+1) < 7^{n+1}$. Then we have + +$$ + 4^{n+2} F(n+2) = 4^{n+2} F(n+1) + 4^{n+2} F(n) < 4\cdot 7^{n+1} + 4^2\cdot 7^n = 44\cdot 7^n < 7^{n+2}. +$$ + +```agda +leveque-strict-bound-Fibonacci-ℕ : + (n : ℕ) → exp-ℕ 4 n *ℕ Fibonacci-ℕ n <-ℕ exp-ℕ 7 n +leveque-strict-bound-Fibonacci-ℕ zero-ℕ = star +leveque-strict-bound-Fibonacci-ℕ (succ-ℕ zero-ℕ) = star +leveque-strict-bound-Fibonacci-ℕ (succ-ℕ (succ-ℕ n)) = + concatenate-eq-le-ℕ + { exp-ℕ 4 (n +ℕ 2) *ℕ Fibonacci-ℕ (n +ℕ 2)} + { 4 *ℕ (exp-ℕ 4 (n +ℕ 1) *ℕ Fibonacci-ℕ (n +ℕ 1)) +ℕ + 16 *ℕ (exp-ℕ 4 n *ℕ Fibonacci-ℕ n)} + { exp-ℕ 7 (n +ℕ 2)} + ( ( left-distributive-mul-add-ℕ + ( exp-ℕ 4 (n +ℕ 2)) + ( Fibonacci-ℕ (succ-ℕ n)) + ( Fibonacci-ℕ n)) ∙ + ( ap-add-ℕ + ( ( ap + ( _*ℕ Fibonacci-ℕ (succ-ℕ n)) + ( commutative-mul-ℕ (exp-ℕ 4 (n +ℕ 1)) 4)) ∙ + ( associative-mul-ℕ 4 (exp-ℕ 4 (n +ℕ 1)) (Fibonacci-ℕ (succ-ℕ n)))) + ( ( ap + ( _*ℕ Fibonacci-ℕ n) + ( ( associative-mul-ℕ (exp-ℕ 4 n) 4 4) ∙ + ( commutative-mul-ℕ (exp-ℕ 4 n) 16))) ∙ + ( associative-mul-ℕ 16 (exp-ℕ 4 n) (Fibonacci-ℕ n))))) + ( concatenate-le-eq-le-ℕ + ( 4 *ℕ (exp-ℕ 4 (n +ℕ 1) *ℕ Fibonacci-ℕ (succ-ℕ n)) +ℕ + 16 *ℕ (exp-ℕ 4 n *ℕ Fibonacci-ℕ n)) + ( 4 *ℕ exp-ℕ 7 (succ-ℕ n) +ℕ 16 *ℕ exp-ℕ 7 n) + ( 44 *ℕ exp-ℕ 7 n) + ( exp-ℕ 7 (n +ℕ 2)) + ( preserves-strict-order-add-ℕ + ( 4 *ℕ (exp-ℕ 4 (n +ℕ 1) *ℕ Fibonacci-ℕ (succ-ℕ n))) + ( 4 *ℕ exp-ℕ 7 (succ-ℕ n)) + ( 16 *ℕ (exp-ℕ 4 n *ℕ Fibonacci-ℕ n)) + ( 16 *ℕ exp-ℕ 7 n) + ( preserves-strict-order-left-mul-ℕ 4 + ( exp-ℕ 4 (n +ℕ 1) *ℕ Fibonacci-ℕ (succ-ℕ n)) + ( exp-ℕ 7 (succ-ℕ n)) + ( is-nonzero-succ-ℕ 3) + ( leveque-strict-bound-Fibonacci-ℕ (succ-ℕ n))) + ( preserves-strict-order-left-mul-ℕ 16 + ( exp-ℕ 4 n *ℕ Fibonacci-ℕ n) + ( exp-ℕ 7 n) + ( is-nonzero-succ-ℕ 15) + ( leveque-strict-bound-Fibonacci-ℕ n))) + ( ( ap-add-ℕ + { 4 *ℕ exp-ℕ 7 (succ-ℕ n)} + { 16 *ℕ exp-ℕ 7 n} + { 28 *ℕ exp-ℕ 7 n} + { 16 *ℕ exp-ℕ 7 n} + ( ( ap (4 *ℕ_) (commutative-mul-ℕ (exp-ℕ 7 n) 7)) ∙ + ( inv (associative-mul-ℕ 4 7 (exp-ℕ 7 n)))) + ( refl)) ∙ + ( inv (right-distributive-mul-add-ℕ 28 16 (exp-ℕ 7 n)))) + ( concatenate-le-eq-ℕ + { 44 *ℕ exp-ℕ 7 n} + { 49 *ℕ exp-ℕ 7 n} + { exp-ℕ 7 (n +ℕ 2)} + ( preserves-strict-order-right-mul-ℕ + ( exp-ℕ 7 n) + ( 44) + ( 49) + ( is-nonzero-exp-ℕ 7 n (is-nonzero-succ-ℕ 6)) + ( star)) + ( ( commutative-mul-ℕ 49 (exp-ℕ 7 n)) ∙ + ( inv (associative-mul-ℕ (exp-ℕ 7 n) 7 7))))) +``` + ## External links - [Fibonacci sequence](https://en.wikipedia.org/wiki/Fibonacci_sequence) at Wikipedia - [A000045](https://oeis.org/A000045) in the OEIS + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index d2d0f0019f..9e315aa466 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -144,6 +144,11 @@ transitive-le-ℕ : (n m l : ℕ) → (le-ℕ n m) → (le-ℕ m l) → (le-ℕ transitive-le-ℕ zero-ℕ (succ-ℕ m) (succ-ℕ l) p q = star transitive-le-ℕ (succ-ℕ n) (succ-ℕ m) (succ-ℕ l) p q = transitive-le-ℕ n m l p q + +concatenate-le-eq-le-ℕ : + (a b c d : ℕ) → a <-ℕ b → b = c → c <-ℕ d → a <-ℕ d +concatenate-le-eq-le-ℕ a b .b d H refl K = + transitive-le-ℕ a b d H K ``` ### A sharper variant of transitivity @@ -421,3 +426,64 @@ preserves-strict-order-mul-ℕ (succ-ℕ a) (succ-ℕ b) (succ-ℕ c) (succ-ℕ ( K)) ( inv (double-successor-law-mul-ℕ b d)) ``` + +### Multiplication by a nonzero natural number on the left or right is strictly order preserving + +```agda +preserves-strict-order-left-mul-succ-ℕ : + (k x y : ℕ) → x <-ℕ y → succ-ℕ k *ℕ x <-ℕ succ-ℕ k *ℕ y +preserves-strict-order-left-mul-succ-ℕ zero-ℕ x y H = + concatenate-eq-le-eq-ℕ (left-unit-law-mul-ℕ x) H (inv (left-unit-law-mul-ℕ y)) +preserves-strict-order-left-mul-succ-ℕ (succ-ℕ k) x y H = + concatenate-eq-le-eq-ℕ + ( left-successor-law-mul-ℕ (succ-ℕ k) x) + ( preserves-strict-order-add-ℕ + ( succ-ℕ k *ℕ x) + ( succ-ℕ k *ℕ y) + ( x) + ( y) + ( preserves-strict-order-left-mul-succ-ℕ k x y H) + ( H)) + ( inv (left-successor-law-mul-ℕ (succ-ℕ k) y)) + +preserves-strict-order-left-mul-is-successor-ℕ : + (k x y : ℕ) → is-successor-ℕ k → x <-ℕ y → k *ℕ x <-ℕ k *ℕ y +preserves-strict-order-left-mul-is-successor-ℕ .(succ-ℕ l) x y (l , refl) = + preserves-strict-order-left-mul-succ-ℕ l x y + +preserves-strict-order-left-mul-ℕ : + (k x y : ℕ) → is-nonzero-ℕ k → x <-ℕ y → k *ℕ x <-ℕ k *ℕ y +preserves-strict-order-left-mul-ℕ k x y H = + preserves-strict-order-left-mul-is-successor-ℕ k x y + ( is-successor-is-nonzero-ℕ H) + +preserves-strict-order-right-mul-succ-ℕ : + (k x y : ℕ) → x <-ℕ y → x *ℕ succ-ℕ k <-ℕ y *ℕ succ-ℕ k +preserves-strict-order-right-mul-succ-ℕ zero-ℕ x y H = + concatenate-eq-le-eq-ℕ + ( right-unit-law-mul-ℕ x) + ( H) + ( inv (right-unit-law-mul-ℕ y)) +preserves-strict-order-right-mul-succ-ℕ (succ-ℕ k) x y H = + concatenate-eq-le-eq-ℕ + ( right-successor-law-mul-ℕ x (succ-ℕ k)) + ( preserves-strict-order-add-ℕ + ( x *ℕ succ-ℕ k) + ( y *ℕ succ-ℕ k) + ( x) + ( y) + ( preserves-strict-order-right-mul-succ-ℕ k x y H) + ( H)) + ( inv (right-successor-law-mul-ℕ y (succ-ℕ k))) + +preserves-strict-order-right-mul-is-successor-ℕ : + (k x y : ℕ) → is-successor-ℕ k → x <-ℕ y → x *ℕ k <-ℕ y *ℕ k +preserves-strict-order-right-mul-is-successor-ℕ .(succ-ℕ l) x y (l , refl) = + preserves-strict-order-right-mul-succ-ℕ l x y + +preserves-strict-order-right-mul-ℕ : + (k x y : ℕ) → is-nonzero-ℕ k → x <-ℕ y → x *ℕ k <-ℕ y *ℕ k +preserves-strict-order-right-mul-ℕ k x y H = + preserves-strict-order-right-mul-is-successor-ℕ k x y + ( is-successor-is-nonzero-ℕ H) +``` From c6b25ed4f6ede64dfe1f5afeccd79f05869307a5 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 5 Jan 2025 14:15:15 -0500 Subject: [PATCH 32/72] generalizing LeVeque's upper bound for the Fibonacci sequence --- .../bezouts-lemma-natural-numbers.lagda.md | 2 +- .../cofibonacci.lagda.md | 4 +- .../equality-natural-numbers.lagda.md | 4 +- .../exponentiation-natural-numbers.lagda.md | 17 +- .../fibonacci-sequence.lagda.md | 174 ++++++++++++------ .../products-of-natural-numbers.lagda.md | 6 +- .../squares-natural-numbers.lagda.md | 25 ++- ...strict-inequality-natural-numbers.lagda.md | 101 +++++++--- 8 files changed, 238 insertions(+), 95 deletions(-) diff --git a/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md b/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md index 407b072465..5dff84516e 100755 --- a/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md @@ -1368,7 +1368,7 @@ remainder-min-dist-succ-x-is-distance x y = coeff-nonnegative : leq-ℤ one-ℤ ((int-ℕ q) *ℤ (int-ℕ s)) coeff-nonnegative = tr (leq-ℤ one-ℤ) (inv (mul-int-ℕ q s)) (leq-int-ℕ 1 (q *ℕ s) - (leq-succ-le-ℕ 0 (q *ℕ s) (le-is-nonzero-ℕ (q *ℕ s) + (leq-succ-le-ℕ 0 (q *ℕ s) (le-zero-is-nonzero-ℕ (q *ℕ s) (is-nonzero-mul-ℕ q s quotient-min-dist-succ-x-nonzero min-dist-succ-x-coeff-nonzero)))) diff --git a/src/elementary-number-theory/cofibonacci.lagda.md b/src/elementary-number-theory/cofibonacci.lagda.md index e2a6f001ba..16ac9c058c 100644 --- a/src/elementary-number-theory/cofibonacci.lagda.md +++ b/src/elementary-number-theory/cofibonacci.lagda.md @@ -57,9 +57,9 @@ is-decidable-is-multiple-of-cofibonacci : (m x : ℕ) → is-decidable (is-multiple-of-cofibonacci m x) is-decidable-is-multiple-of-cofibonacci m x = is-decidable-function-type - ( is-decidable-is-nonzero-ℕ m) + ( is-decidable-zero-is-nonzero-ℕ m) ( is-decidable-product - ( is-decidable-is-nonzero-ℕ x) + ( is-decidable-zero-is-nonzero-ℕ x) ( is-decidable-div-ℕ m (Fibonacci-ℕ x))) cofibonacci-multiple : (m : ℕ) → Σ ℕ (is-multiple-of-cofibonacci m) diff --git a/src/elementary-number-theory/equality-natural-numbers.lagda.md b/src/elementary-number-theory/equality-natural-numbers.lagda.md index bb335e5e58..06ac27ceeb 100644 --- a/src/elementary-number-theory/equality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/equality-natural-numbers.lagda.md @@ -132,8 +132,8 @@ is-decidable-is-zero-ℕ n = has-decidable-equality-ℕ n zero-ℕ is-decidable-is-zero-ℕ' : (n : ℕ) → is-decidable (is-zero-ℕ' n) is-decidable-is-zero-ℕ' n = has-decidable-equality-ℕ zero-ℕ n -is-decidable-is-nonzero-ℕ : (n : ℕ) → is-decidable (is-nonzero-ℕ n) -is-decidable-is-nonzero-ℕ n = +is-decidable-zero-is-nonzero-ℕ : (n : ℕ) → is-decidable (is-nonzero-ℕ n) +is-decidable-zero-is-nonzero-ℕ n = is-decidable-neg (is-decidable-is-zero-ℕ n) is-decidable-is-one-ℕ : (n : ℕ) → is-decidable (is-one-ℕ n) diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index abbf1734ae..e852edcb4a 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -14,9 +14,11 @@ open import elementary-number-theory.commutative-semiring-of-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.products-of-natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions open import foundation.identity-types +open import foundation.unit-type ``` @@ -48,6 +50,9 @@ power-ℕ = power-Commutative-Semiring ℕ-Commutative-Semiring ### Tarski's high school arithmetic laws for exponentiation ```agda +right-unit-law-exp-ℕ : (n : ℕ) → exp-ℕ n 1 = n +right-unit-law-exp-ℕ n = left-unit-law-mul-ℕ n + annihilation-law-exp-ℕ : (n : ℕ) → 1 ^ℕ n = 1 annihilation-law-exp-ℕ zero-ℕ = refl annihilation-law-exp-ℕ (succ-ℕ n) = @@ -75,14 +80,22 @@ exp-mul-ℕ x (succ-ℕ y) z = ( inv (right-distributive-exp-mul-ℕ (x ^ℕ y) x z))) ``` -### The exponent $m^n$ is always nonzero +### The exponent $m^n$ is nonzero if $m$ is nonzero ```agda is-nonzero-exp-ℕ : (m n : ℕ) → is-nonzero-ℕ m → is-nonzero-ℕ (m ^ℕ n) -is-nonzero-exp-ℕ m zero-ℕ p = is-nonzero-one-ℕ +is-nonzero-exp-ℕ m zero-ℕ p = + is-nonzero-one-ℕ is-nonzero-exp-ℕ m (succ-ℕ n) p = is-nonzero-mul-ℕ (m ^ℕ n) m (is-nonzero-exp-ℕ m n p) p + +le-zero-exp-ℕ : + (m n : ℕ) → 0 <-ℕ m → 0 <-ℕ m ^ℕ n +le-zero-exp-ℕ m zero-ℕ H = + star +le-zero-exp-ℕ m (succ-ℕ n) H = + preserves-strict-order-mul-ℕ 0 (m ^ℕ n) 0 m (le-zero-exp-ℕ m n H) H ``` ### The exponent $m^n$ is equal to the $n$-fold product of $m$ diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index d1546a51f7..762e771a56 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -11,10 +11,12 @@ open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.greatest-common-divisor-natural-numbers +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.relatively-prime-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import elementary-number-theory.squares-natural-numbers open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types @@ -231,81 +233,141 @@ preserves-div-Fibonacci-ℕ m n H = ### LeVeque's strict bound on the growth of the Fibonacci numbers -The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This claim appears on pages 7 and 8 in section 1.2 of {{#cite "Leveque12volI"}} as an example of a proof by induction. +The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This claim appears on pages 7 and 8 in section 1.2 of {{#cite "Leveque12volI"}} as an instructive example of a proof by induction. The upper bound works for any fraction $\frac{b}{a}$ where $a(b+a) Date: Mon, 6 Jan 2025 11:03:51 -0500 Subject: [PATCH 33/72] edits --- references.bib | 10 +++ .../fermat-numbers.lagda.md | 78 ++++++++++++++++++- .../fibonacci-sequence.lagda.md | 30 ++++--- .../infinitude-of-primes.lagda.md | 9 ++- .../mersenne-primes.lagda.md | 21 +++-- .../natural-numbers.lagda.md | 60 +++++++++++--- .../prime-numbers.lagda.md | 75 ++++++++++++------ .../products-of-natural-numbers.lagda.md | 12 +++ .../sieve-of-eratosthenes.lagda.md | 6 +- src/univalent-combinatorics/counting.lagda.md | 2 +- 10 files changed, 242 insertions(+), 61 deletions(-) diff --git a/references.bib b/references.bib index ef8570c226..2a44d6d83e 100644 --- a/references.bib +++ b/references.bib @@ -10,6 +10,16 @@ @online{100theorems url = {https://www.cs.ru.nl/~freek/100/} } +@book{AZ18, + author = {Aigner, Martin and Ziegler, Gnter M.}, + title = {Proofs from THE BOOK}, + year = {2018}, + isbn = {3662572648}, + publisher = {Springer Publishing Company, Incorporated}, + edition = {6th}, + abstract = {This revised and enlarged sixth edition of Proofs from THE BOOKfeatures an entirely new chapter on Van der Waerdens permanent conjecture, as well as additional, highly original and delightful proofs in other chapters.} +} + @article{AKS15, title = {Univalent Categories and the {{Rezk}} Completion}, author = {Ahrens, Benedikt and Kapulkin, Krzysztof and Shulman, Michael}, diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index 8a76e9fc21..be8ad43e8e 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -8,11 +8,16 @@ module elementary-number-theory.fermat-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.products-of-natural-numbers open import elementary-number-theory.strong-induction-natural-numbers +open import foundation.action-on-identifications-functions +open import foundation.identity-types +open import foundation.negated-equality + open import univalent-combinatorics.standard-finite-types ``` @@ -20,7 +25,7 @@ open import univalent-combinatorics.standard-finite-types ## Idea -{{#concept "Fermat numbers" WD="Fermat number" WDID=Q207264 Agda=fermat-number-ℕ}} +{{#concept "Fermat numbers" WD="Fermat number" WDID=Q207264 OEIS=A000215 Agda=fermat-number-ℕ}} are numbers of the form $F_n := 2^{2^n}+1$. The first five Fermat numbers are ```text @@ -28,7 +33,7 @@ are numbers of the form $F_n := 2^{2^n}+1$. The first five Fermat numbers are ``` The sequence of Fermat numbers is listed as A000215 in the -[Online Encyclopedia of Integer Sequences](literature.oeis.md). +[Online Encyclopedia of Integer Sequences](literature.oeis.md) {{#cite "OEIS"}}. Alternatively, the Fermat numbers can be defined with [strong induction](elementary-number-theory.strong-induction-natural-numbers.md) @@ -40,9 +45,10 @@ F (n + 1) := 2 + Π_{i≤n} F_i ``` This recurrence implies that any two Fermat numbers are -[relatively prime](elementary-number-theory.relatively-prime-natural-numbers.md). +[relatively prime](elementary-number-theory.relatively-prime-natural-numbers.md), because it follows from this recurrence that any divisor of two distinct Fermat numbers must divide the number $2$. Since the Fermat numbers are all odd, this implies that any divisor of two distinct Fermat numbers must be the number $1$. + Goldbach used this observation to prove the -[infinitude of primes](elementary-number-theory.infinitude-of-primes.md): Since +[infinitude of primes](elementary-number-theory.infinitude-of-primes.md) {{#cite "AZ18"}} : Since there are infinitely many Fermat numbers, and all of them are relatively prime, there must be infinitely many prime numbers. Fermat numbers also feature in a series of long-standing open problems in mathematics, including: @@ -71,9 +77,73 @@ recursive-fermat-number-ℕ = ( succ-ℕ n) ( λ i → f (nat-Fin (succ-ℕ n) i) (upper-bound-nat-Fin n i))) ( 2)) + +compute-succ-recursive-fermat-number-ℕ : + (n : ℕ) → + recursive-fermat-number-ℕ (succ-ℕ n) = + Π-ℕ (succ-ℕ n) (λ i → recursive-fermat-number-ℕ (nat-Fin (succ-ℕ n) i)) +ℕ 2 +compute-succ-recursive-fermat-number-ℕ = + compute-succ-strong-rec-ℕ 3 + ( λ n f → + add-ℕ + ( Π-ℕ + ( succ-ℕ n) + ( λ i → f (nat-Fin (succ-ℕ n) i) (upper-bound-nat-Fin n i))) + ( 2)) +``` + +## Properties + +### The Fermat number at a successor + +Any Fermat number of the form $F(n+1)=2^{2^{n+1}}+1$ can be computed as + +$$ + F(n+1)-2=(2^{2^n})^2-1=(2^{2^n}+1)(2^{2^n}-1)=F(n)(F(n)-2). +$$ + +```text +fermat-number-succ-ℕ : + (n : ℕ) → fermat-number-ℕ (succ-ℕ n) = +``` + +### The two definitions of the Fermat numbers agree + +**Proof.** The proof is by strong induction on the natural numbers. The base case holds by reflexivity. For the inductive step, assume that + +```agda +compute-recursive-fermat-number-ℕ : + (n : ℕ) → recursive-fermat-number-ℕ n = fermat-number-ℕ n +compute-recursive-fermat-number-ℕ = + strong-ind-ℕ + ( λ n → recursive-fermat-number-ℕ n = fermat-number-ℕ n) + ( refl) + ( λ n H → + ( compute-succ-recursive-fermat-number-ℕ n) ∙ + ( ap + ( _+ℕ 2) + ( preserves-htpy-Π-ℕ + ( succ-ℕ n) + ( λ i → H (nat-Fin (succ-ℕ n) i) (upper-bound-nat-Fin n i)))) ∙ + {!!}) +``` + +### Any two distinct Fermat numbers are relatively prime + +**Proof.** By the recursive definition of the Fermat numbers, we have that $F_{n+1}=(\prod_{n=0}^n F_n) + 2$. This implies that if $d$ divides $F_m$ and $F_n$ for some $n>m$, then $d|2$. However, the Fermat numbers are odd, so $d=1$. + +```agda +is-one-div-fermat-number-ℕ : + (m n d : ℕ) → m ≠ n → + div-ℕ d (fermat-number-ℕ m) → div-ℕ d (fermat-number-ℕ n) → is-one-ℕ d +is-one-div-fermat-number-ℕ m n d H K L = {!!} ``` ## External link - [Fermat number](https://en.wikipedia.org/wiki/Fermat_number) at Wikipedia - [A000215](https://oeis.org/A000215) in the OEIS + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index 762e771a56..d44d21a6a8 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -236,10 +236,16 @@ preserves-div-Fibonacci-ℕ m n H = The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This claim appears on pages 7 and 8 in section 1.2 of {{#cite "Leveque12volI"}} as an instructive example of a proof by induction. The upper bound works for any fraction $\frac{b}{a}$ where $a(b+a) Date: Mon, 6 Jan 2025 14:05:20 -0500 Subject: [PATCH 34/72] mersenne numbers and bounded natural numbers --- references.bib | 3 +- .../bounded-natural-numbers.lagda.md | 59 +++++++++++++ .../inequality-natural-numbers.lagda.md | 4 +- .../mersenne-exponents.lagda.md | 32 +++++++ .../mersenne-numbers.lagda.md | 74 ++++++++++++++++ .../mersenne-primes.lagda.md | 26 +++--- .../number-of-divisors.lagda.md | 87 +++++++++++++++++++ ...strict-inequality-natural-numbers.lagda.md | 4 + .../strictly-bounded-natural-numbers.lagda.md | 59 +++++++++++++ src/literature/oeis.lagda.md | 27 ++++++ .../standard-finite-types.lagda.md | 6 ++ 11 files changed, 367 insertions(+), 14 deletions(-) create mode 100644 src/elementary-number-theory/bounded-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/mersenne-exponents.lagda.md create mode 100644 src/elementary-number-theory/mersenne-numbers.lagda.md create mode 100644 src/elementary-number-theory/number-of-divisors.lagda.md create mode 100644 src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md diff --git a/references.bib b/references.bib index f91bddc4d9..09920fc085 100644 --- a/references.bib +++ b/references.bib @@ -11,7 +11,7 @@ @online{100theorems } @book{AZ18, - author = {Aigner, Martin and Ziegler, Gnter M.}, + author = {Aigner, Martin and Ziegler, G\"{u}nter M.}, title = {Proofs from THE BOOK}, year = {2018}, isbn = {3662572648}, @@ -74,7 +74,6 @@ @book{Andrews94 publisher = {Dover Publications} } - @misc{Awodey22, author = {{Awodey}, Steve}, title = "{On Hofmann-Streicher universes}", diff --git a/src/elementary-number-theory/bounded-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-natural-numbers.lagda.md new file mode 100644 index 0000000000..5e01d46ef4 --- /dev/null +++ b/src/elementary-number-theory/bounded-natural-numbers.lagda.md @@ -0,0 +1,59 @@ +# Bounded natural numbers + +```agda +module elementary-number-theory.bounded-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.natural-numbers + +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.universe-levels + +open import univalent-combinatorics.counting +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +The type of {{#concept "bounded natural numbers" Agda=bounded-ℕ}} with upper bound $n$ is the type + +$$ + \mathbb{N}_{\leq n} := \{k : ℕ \mid k \leq n\}. +$$ + +The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) to the [standard finite type](univalent-combinatorics.standard-finite-types.md) $\mathsf{Fin}(n+1)$. + +## Definition + +### The bounded natural numbers + +```agda +bounded-ℕ : ℕ → UU lzero +bounded-ℕ n = Σ ℕ (λ k → k ≤-ℕ n) +``` + +## Properties + +### The type $\mathbb{N}_{\leq b}$ is equivalent to the standard finite type $\mathsf{Fin}(n+1)$ + +```agda +equiv-count-bounded-ℕ : + (n : ℕ) → Fin (succ-ℕ n) ≃ bounded-ℕ n +equiv-count-bounded-ℕ n = {!!} + +count-bounded-ℕ : + (n : ℕ) → count (bounded-ℕ n) +pr1 (count-bounded-ℕ n) = succ-ℕ n +pr2 (count-bounded-ℕ n) = equiv-count-bounded-ℕ n +``` + +## See also + +- [The strictly bounded natural numbers](elementary-number-theory.strictly-bounded-natural-numbers.md) diff --git a/src/elementary-number-theory/inequality-natural-numbers.lagda.md b/src/elementary-number-theory/inequality-natural-numbers.lagda.md index ab185fae08..e559497525 100644 --- a/src/elementary-number-theory/inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/inequality-natural-numbers.lagda.md @@ -532,5 +532,5 @@ leq-left-factor-eq-ℕ m n .(m *ℕ n) H refl = ## See also -- Strict inequality of the natural numbers is defined in - [`strict-inequality-natural-numbers`](elementary-number-theory.strict-inequality-natural-numbers.md) +- [The bounded natural numbers](elementary-number-theory.bounded-natural-numbers.md) +- [Strict inequality of the natural numbers](elementary-number-theory.strict-inequality-natural-numbers.md) diff --git a/src/elementary-number-theory/mersenne-exponents.lagda.md b/src/elementary-number-theory/mersenne-exponents.lagda.md new file mode 100644 index 0000000000..dfd2114046 --- /dev/null +++ b/src/elementary-number-theory/mersenne-exponents.lagda.md @@ -0,0 +1,32 @@ +# Mersenne exponents + +```agda +module elementary-number-theory.mersenne-exponents where +``` + +
Imports + +```agda +open import elementary-number-theory.mersenne-primes +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.prime-numbers + +open import foundation.cartesian-product-types +open import foundation.universe-levels +``` + +
+ +## Idea + +A {{#concept "Mersenne exponent" Agda=mersenne-exponent-ℕ OEIS=A000043}} is a [prime number](elementary-number-theory.prime-numbers.md) $p$ such that the [Mersenne number](elementary-number-theory.mersenne-numbers.md) $M_p$ is a [Mersenne prime](elementary-number-theory.mersenne-primes.md). + +## Definitions + +### The predicate of being a Mersenne exponent + +```agda +is-mersenne-exponent-ℕ : ℕ → UU lzero +is-mersenne-exponent-ℕ n = + is-prime-ℕ n × is-prime-mersenne-number-ℕ n +``` diff --git a/src/elementary-number-theory/mersenne-numbers.lagda.md b/src/elementary-number-theory/mersenne-numbers.lagda.md new file mode 100644 index 0000000000..ccea42c695 --- /dev/null +++ b/src/elementary-number-theory/mersenne-numbers.lagda.md @@ -0,0 +1,74 @@ +# Mersenne numbers + +```agda +module elementary-number-theory.mersenne-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.exponentiation-natural-numbers +open import elementary-number-theory.infinitude-of-primes +open import elementary-number-theory.natural-numbers + +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.universe-levels +``` + +
+ +## Idea + +The $n$th {{#concept "Mersenne number" Agda=mersenne-number-ℕ OEIS=A000225}} +$M_n$ is defined to be + +$$ + M_n := 2^n-1. +$$ + +[Mersenne primes](elementary-number-theory.mersenne-primes.md) are Mersenne +numbers that are [prime](elementary-number-theory.prime-numbers.md). + +## Definitions + +### The Mersenne numbers + +```agda +mersenne-number-ℕ : + ℕ → ℕ +mersenne-number-ℕ n = + pred-is-nonzero-ℕ (exp-ℕ 2 n) (is-nonzero-exp-ℕ 2 n is-nonzero-two-ℕ) +``` + +### The Mersenne numbers of prime powers + +Some sources only refer to numbers of the form $2^p-1$ with $p$ prime as +Mersenne numbers. The sequence $n ↦ 2^{p(n)}-1$, where $p(n)$ is the $n$th prime +number, is listed as A001348 in the [OEIS](literature.oeis.md) {{#cite "OEIS"}}. +The first few numbers in this sequence are + +```text + 3, 7, 31, 127, 2047, ... +``` + +```agda +mersenne-number-prime-ℕ : + ℕ → ℕ +mersenne-number-prime-ℕ = + mersenne-number-ℕ ∘ prime-ℕ +``` + +### The predicate of being a Mersenne number + +```agda +is-mersenne-number-ℕ : + ℕ → UU lzero +is-mersenne-number-ℕ = + fiber mersenne-number-ℕ +``` + +## See also + +- [Mersenne exponents](elementary-number-theory.mersenne-exponents.md) +- [Mersenne primes](elementary-number-theory.mersenne-primes.md) diff --git a/src/elementary-number-theory/mersenne-primes.lagda.md b/src/elementary-number-theory/mersenne-primes.lagda.md index 322e2f48fc..e3fc94020a 100644 --- a/src/elementary-number-theory/mersenne-primes.lagda.md +++ b/src/elementary-number-theory/mersenne-primes.lagda.md @@ -9,6 +9,7 @@ module elementary-number-theory.mersenne-primes where ```agda open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers +open import elementary-number-theory.mersenne-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.prime-numbers @@ -22,23 +23,28 @@ open import foundation.universe-levels ## Idea -A {{#concept "Mersenne prime" Agda=is-mersenne-prime-ℕ}} is a prime number that is one less than a power of two. +A +{{#concept "Mersenne prime" Agda=is-mersenne-prime-ℕ WDID=Q186875 WD="Mersenne prime"}} +is a [prime number](elementary-number-theory.prime-numbers.md) that is one less +than a +[power of two](elementary-number-theory.exponentiation-natural-numbers.md). ## Definitions -### The Mersenne numbers +### The predicate of being a Mersenne prime ```agda -mersenne-number-ℕ : ℕ → ℕ -mersenne-number-ℕ = ? +is-mersenne-prime-ℕ : + ℕ → UU lzero +is-mersenne-prime-ℕ n = + is-prime-ℕ n × is-mersenne-number-ℕ n ``` -### The predicate of being a Mersenne prime +### The predicate on Mersenne numbers of being prime ```agda -is-mersenne-prime-ℕ : ℕ → UU lzero -is-mersenne-prime-ℕ n = is-prime-ℕ n × Σ ℕ (λ k → dist-ℕ (exp-ℕ 2 k) 1 = n) - -is-mersenne-prime-power-ℕ : ℕ → UU lzero -is-mersenne-prime-power-ℕ k = is-prime-ℕ (dist-ℕ (exp-ℕ 2 k) 1) +is-prime-mersenne-number-ℕ : + ℕ → UU lzero +is-prime-mersenne-number-ℕ k = + is-prime-ℕ (mersenne-number-ℕ k) ``` diff --git a/src/elementary-number-theory/number-of-divisors.lagda.md b/src/elementary-number-theory/number-of-divisors.lagda.md new file mode 100644 index 0000000000..58cc3774e7 --- /dev/null +++ b/src/elementary-number-theory/number-of-divisors.lagda.md @@ -0,0 +1,87 @@ +# The number of divisors function + +```agda +module elementary-number-theory.number-of-divisors where +``` + +
Imports + +```agda +open import elementary-number-theory.bounded-divisibility-natural-numbers +open import elementary-number-theory.modular-arithmetic-standard-finite-types +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers + +open import foundation.decidable-subtypes +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.universe-levels + +open import univalent-combinatorics.counting +open import univalent-combinatorics.counting-decidable-subtypes +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +The {{#concept "number of divisors" Agda=number-of-divisors-ℕ OEIS=A000005 WDID=Q12890701 WD="number of divisors"}} $\tau(n)$ of a [natural number](elementary-number-theory.natural-numbers.md) $n$ is the number of natural numbers $d \leq n$ that [divide](elementary-number-theory.divisibility.natural-numbers.md) the number $n$, i.e., + +$$ + \tau(n) := |\{d \in \mathbb{N} \mid d|n\}|. +$$ + +The number of divisors function is listed as A000005 in the [OEIS](literature.oeis.md) {{#cite "OEIS"}}. + +## Definitions + +### The type of divisors + +```agda +decidable-subtype-of-divisors-ℕ : + (n : ℕ) → decidable-subtype lzero (Fin (succ-ℕ n)) +decidable-subtype-of-divisors-ℕ n i = + bounded-div-ℕ-Decidable-Prop (nat-Fin (succ-ℕ n) i) n + +type-of-divisors-ℕ : ℕ → UU lzero +type-of-divisors-ℕ n = + type-decidable-subtype (decidable-subtype-of-divisors-ℕ n) + +quotient-type-of-divisors-ℕ : + (n : ℕ) → type-of-divisors-ℕ n → type-of-divisors-ℕ n +pr1 (quotient-type-of-divisors-ℕ n (d , H)) = + mod-succ-ℕ n (quotient-bounded-div-ℕ (nat-Fin (succ-ℕ n) d) n H) +pr1 (pr2 (quotient-type-of-divisors-ℕ n (d , H))) = + nat-Fin (succ-ℕ n) d +pr1 (pr2 (pr2 (quotient-type-of-divisors-ℕ n (d , H)))) = + upper-bound-nat-Fin n d +pr2 (pr2 (pr2 (quotient-type-of-divisors-ℕ n (d , H)))) = + {!is-section-nat-Fin n!} +``` + +### The number of divisors function + +Note that the entry `number-of-elements-decidable-subtype` refers to `count-decidable-subtype'`, which is `abstract` and therefore doesn't compute. + +```agda +count-type-of-divisors-ℕ : + (n : ℕ) → count (type-of-divisors-ℕ n) +count-type-of-divisors-ℕ n = + count-decidable-subtype + ( decidable-subtype-of-divisors-ℕ n) + ( count-Fin (succ-ℕ n)) + +number-of-divisors-ℕ : ℕ → ℕ +number-of-divisors-ℕ n = + number-of-elements-decidable-subtype + ( decidable-subtype-of-divisors-ℕ n) + ( count-Fin (succ-ℕ n)) +``` + +### Involution on the type of divisors + +For every natural number `n : ℕ`, the map `d ↦ n/d` mapping a divisor of `n` to its quotient is an involution. This operation has a fixed point precisely when `n` is a [square number](elementary-number-theory.square-natural-numbers.md). + +```agda +``` diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 2e7b273813..ec9f968a2c 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -530,3 +530,7 @@ preserves-strict-order-right-mul-ℕ k x y H = preserves-strict-order-right-mul-is-successor-ℕ k x y ( is-successor-is-nonzero-ℕ H) ``` + +## See also + +- [The strictly bounded natural numbers](elementary-number-theory.strictly-bounded-natural-numbers.md) diff --git a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md new file mode 100644 index 0000000000..d680cd9a65 --- /dev/null +++ b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md @@ -0,0 +1,59 @@ +# Strictly bounded natural numbers + +```agda +module elementary-number-theory.strictly-bounded-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers + +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.universe-levels + +open import univalent-combinatorics.counting +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +The type of {{#concept "strictly bounded natural numbers" Agda=strictly-bounded-ℕ}} with strict upper bound $n$ is the type + +$$ + \mathbb{N}_{< n} := \{k : ℕ \mid k < n\}. +$$ + +The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) to the [standard finite type](univalent-combinatorics.standard-finite-types.md) $\mathsf{Fin}(n+1)$. + +## Definition + +### The strictly bounded natural numbers + +```agda +strictly-bounded-ℕ : ℕ → UU lzero +strictly-bounded-ℕ n = Σ ℕ (λ k → k <-ℕ n) +``` + +## Properties + +### The type $\mathbb{N}_{\leq b}$ is equivalent to the standard finite type $\mathsf{Fin}(n+1)$ + +```agda +equiv-count-strictly-bounded-ℕ : + (n : ℕ) → Fin n ≃ strictly-bounded-ℕ n +equiv-count-strictly-bounded-ℕ n = {!!} + +count-strictly-bounded-ℕ : + (n : ℕ) → count (strictly-bounded-ℕ n) +pr1 (count-strictly-bounded-ℕ n) = n +pr2 (count-strictly-bounded-ℕ n) = equiv-count-strictly-bounded-ℕ n +``` + +## See also + +- [The bounded natural numbers](elementary-number-theory.bounded-natural-numbers.md) diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index 8a6a366206..cca06d8c3e 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -100,6 +100,15 @@ open import elementary-number-theory.infinitude-of-primes using OEIS: [A000040](https://oeis.org/A000040) +### A000043 -- The Mersenne exponents + +```agda +open import elementary-number-theory.mersenne-exponents using + ( is-mersenne-exponent-ℕ) +``` + +OEIS: [A000043](https://oeis.org/A000043) + ### A000045 -- The Fibonacci sequence ```agda @@ -163,6 +172,15 @@ open import elementary-number-theory.triangular-numbers using OEIS: [A000217](https://oeis.org/A000217) +### A000225 -- Mersenne numbers + +```agda +open import elementary-number-theory.mersenne-numbers using + ( mersenne-number-ℕ) +``` + +OEIS: [A000225](https://oeis.org/A000225) + ### A000244 -- Powers of `3` ```agda @@ -217,6 +235,15 @@ open import elementary-number-theory.cofibonacci using OEIS: [A001177](https://oeis.org/A001177) +### A001348 -- The Mersenne numbers (at primes) + +```agda +open import elementary-number-theory.mersenne-numbers using + ( mersenne-number-prime-ℕ) +``` + +OEIS: [A001348](https://oeis.org/A001348) + ### A001477 -- The natural numbers ```agda diff --git a/src/univalent-combinatorics/standard-finite-types.lagda.md b/src/univalent-combinatorics/standard-finite-types.lagda.md index 1a60373a5f..7c0ab90427 100644 --- a/src/univalent-combinatorics/standard-finite-types.lagda.md +++ b/src/univalent-combinatorics/standard-finite-types.lagda.md @@ -491,3 +491,9 @@ is-preunivalent-Fin : is-preunivalent Fin is-preunivalent-Fin = is-preunivalent-retraction-equiv-tr-Set Fin-Set retraction-equiv-tr-Fin ``` + +## See also + +- [The bounded natural numbers](elementary-number-theory.bounded-natural-numbers.md) +- [The strictly bounded natural numbers](elementary-number-theory.strictly-bounded-natural-numbers.md) +- [Modular arithmetic on the standard finite types](elementary-number-theory.modular-arithmetic-standard-finite-types.md) From e0318565e84da1e2475bbd3573d179ffe73db338 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 6 Jan 2025 15:33:25 -0500 Subject: [PATCH 35/72] fixing citations --- .../fermat-numbers.lagda.md | 25 ++++++++++++------- .../fibonacci-sequence.lagda.md | 24 ++++++++++++++---- .../infinitude-of-primes.lagda.md | 5 +++- .../mersenne-numbers.lagda.md | 2 +- .../nicomachuss-theorem.lagda.md | 11 ++++---- .../number-of-divisors.lagda.md | 18 ++++++++++--- .../pronic-numbers.lagda.md | 18 +++++++------ .../square-pyramidal-numbers.lagda.md | 22 +++++++++++----- .../squares-natural-numbers.lagda.md | 4 +-- 9 files changed, 89 insertions(+), 40 deletions(-) diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index be8ad43e8e..3e169f638b 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -33,7 +33,7 @@ are numbers of the form $F_n := 2^{2^n}+1$. The first five Fermat numbers are ``` The sequence of Fermat numbers is listed as A000215 in the -[Online Encyclopedia of Integer Sequences](literature.oeis.md) {{#cite "OEIS"}}. +[Online Encyclopedia of Integer Sequences](literature.oeis.md) {{#cite OEIS}}. Alternatively, the Fermat numbers can be defined with [strong induction](elementary-number-theory.strong-induction-natural-numbers.md) @@ -45,13 +45,17 @@ F (n + 1) := 2 + Π_{i≤n} F_i ``` This recurrence implies that any two Fermat numbers are -[relatively prime](elementary-number-theory.relatively-prime-natural-numbers.md), because it follows from this recurrence that any divisor of two distinct Fermat numbers must divide the number $2$. Since the Fermat numbers are all odd, this implies that any divisor of two distinct Fermat numbers must be the number $1$. +[relatively prime](elementary-number-theory.relatively-prime-natural-numbers.md), +because it follows from this recurrence that any divisor of two distinct Fermat +numbers must divide the number $2$. Since the Fermat numbers are all odd, this +implies that any divisor of two distinct Fermat numbers must be the number $1$. Goldbach used this observation to prove the -[infinitude of primes](elementary-number-theory.infinitude-of-primes.md) {{#cite "AZ18"}} : Since -there are infinitely many Fermat numbers, and all of them are relatively prime, -there must be infinitely many prime numbers. Fermat numbers also feature in a -series of long-standing open problems in mathematics, including: +[infinitude of primes](elementary-number-theory.infinitude-of-primes.md) +{{#cite AZ18}} : Since there are infinitely many Fermat numbers, and all of them +are relatively prime, there must be infinitely many prime numbers. Fermat +numbers also feature in a series of long-standing open problems in mathematics, +including: - Are there infinitely many prime Fermat numbers? - Is $F_n$ composite for all $n\geq 5$? @@ -104,12 +108,13 @@ $$ ```text fermat-number-succ-ℕ : - (n : ℕ) → fermat-number-ℕ (succ-ℕ n) = + (n : ℕ) → fermat-number-ℕ (succ-ℕ n) = ``` ### The two definitions of the Fermat numbers agree -**Proof.** The proof is by strong induction on the natural numbers. The base case holds by reflexivity. For the inductive step, assume that +**Proof.** The proof is by strong induction on the natural numbers. The base +case holds by reflexivity. For the inductive step, assume that ```agda compute-recursive-fermat-number-ℕ : @@ -130,7 +135,9 @@ compute-recursive-fermat-number-ℕ = ### Any two distinct Fermat numbers are relatively prime -**Proof.** By the recursive definition of the Fermat numbers, we have that $F_{n+1}=(\prod_{n=0}^n F_n) + 2$. This implies that if $d$ divides $F_m$ and $F_n$ for some $n>m$, then $d|2$. However, the Fermat numbers are odd, so $d=1$. +**Proof.** By the recursive definition of the Fermat numbers, we have that +$F_{n+1}=(\prod_{n=0}^n F_n) + 2$. This implies that if $d$ divides $F_m$ and +$F_n$ for some $n>m$, then $d|2$. However, the Fermat numbers are odd, so $d=1$. ```agda is-one-div-fermat-number-ℕ : diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index d44d21a6a8..c0e10b8df6 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -233,21 +233,35 @@ preserves-div-Fibonacci-ℕ m n H = ### LeVeque's strict bound on the growth of the Fibonacci numbers -The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This claim appears on pages 7 and 8 in section 1.2 of {{#cite "Leveque12volI"}} as an instructive example of a proof by induction. The upper bound works for any fraction $\frac{b}{a}$ where $a(b+a) Date: Mon, 6 Jan 2025 15:37:13 -0500 Subject: [PATCH 36/72] resolve merge conflicts --- src/literature/oeis.lagda.md | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index b763939b7d..3a159097a8 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -314,16 +314,16 @@ open import elementary-number-theory.ackermann-function using ( simplified-ackermann-ℕ) ``` -<<<<<<< HEAD OEIS: [A046859](https://oeis.org/A046859) -======= -### [A058129](https://oeis.org/A058129) The number of monoids of order `n` up to isomorphism + +### A058129 -- The number of monoids of order `n` up to isomorphism ```agda open import finite-group-theory.finite-monoids using ( number-of-monoids-of-order) ``` ->>>>>>> 871a0295de0f8e48d52561158d55e9c7fdd69ab9 + +OEIS: [A058129](https://oeis.org/A058129) ## References From be13fbffb0d57a30a047f4fb12d8d4d5c52c601d Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 6 Jan 2025 21:36:38 -0500 Subject: [PATCH 37/72] involution on the type of divisors --- .../bounded-natural-numbers.lagda.md | 64 +++++- .../number-of-divisors.lagda.md | 183 ++++++++++++++++-- .../squares-natural-numbers.lagda.md | 88 +++++---- .../strictly-bounded-natural-numbers.lagda.md | 72 ++++++- src/literature/oeis.lagda.md | 4 +- src/univalent-combinatorics/counting.lagda.md | 10 +- 6 files changed, 354 insertions(+), 67 deletions(-) diff --git a/src/elementary-number-theory/bounded-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-natural-numbers.lagda.md index 5e01d46ef4..c95c04bc67 100644 --- a/src/elementary-number-theory/bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bounded-natural-numbers.lagda.md @@ -7,11 +7,19 @@ module elementary-number-theory.bounded-natural-numbers where
Imports ```agda +open import elementary-number-theory.congruence-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.retractions +open import foundation.sections +open import foundation.sets +open import foundation.subtypes open import foundation.universe-levels open import univalent-combinatorics.counting @@ -37,16 +45,70 @@ The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) ```agda bounded-ℕ : ℕ → UU lzero bounded-ℕ n = Σ ℕ (λ k → k ≤-ℕ n) + +nat-bounded-ℕ : (n : ℕ) → bounded-ℕ n → ℕ +nat-bounded-ℕ n = pr1 + +upper-bound-nat-bounded-ℕ : (n : ℕ) (x : bounded-ℕ n) → nat-bounded-ℕ n x ≤-ℕ n +upper-bound-nat-bounded-ℕ n = pr2 ``` ## Properties +### The type of bounded natural numbers is a set + +```agda +is-set-bounded-ℕ : + (n : ℕ) → is-set (bounded-ℕ n) +is-set-bounded-ℕ n = + is-set-type-subtype + ( λ k → leq-ℕ-Prop k n) + ( is-set-ℕ) +``` + ### The type $\mathbb{N}_{\leq b}$ is equivalent to the standard finite type $\mathsf{Fin}(n+1)$ ```agda +bounded-nat-Fin : + (n : ℕ) → Fin (succ-ℕ n) → bounded-ℕ n +pr1 (bounded-nat-Fin n x) = nat-Fin (succ-ℕ n) x +pr2 (bounded-nat-Fin n x) = upper-bound-nat-Fin n x + +fin-bounded-ℕ : + (n : ℕ) → bounded-ℕ n → Fin (succ-ℕ n) +fin-bounded-ℕ n x = + mod-succ-ℕ n (nat-bounded-ℕ n x) + +is-section-fin-bounded-ℕ : + (n : ℕ) → is-section (bounded-nat-Fin n) (fin-bounded-ℕ n) +is-section-fin-bounded-ℕ n i = + eq-type-subtype + ( λ x → leq-ℕ-Prop x n) + ( eq-cong-le-ℕ + ( succ-ℕ n) + ( nat-Fin (succ-ℕ n) (fin-bounded-ℕ n i)) + ( nat-bounded-ℕ n i) + ( strict-upper-bound-nat-Fin (succ-ℕ n) (fin-bounded-ℕ n i)) + ( le-succ-leq-ℕ (nat-bounded-ℕ n i) n (upper-bound-nat-bounded-ℕ n i)) + ( cong-nat-mod-succ-ℕ n (nat-bounded-ℕ n i))) + +is-retraction-fin-bounded-ℕ : + (n : ℕ) → is-retraction (bounded-nat-Fin n) (fin-bounded-ℕ n) +is-retraction-fin-bounded-ℕ n = + is-section-nat-Fin n + +is-equiv-bounded-nat-Fin : + (n : ℕ) → is-equiv (bounded-nat-Fin n) +is-equiv-bounded-nat-Fin n = + is-equiv-is-invertible + ( fin-bounded-ℕ n) + ( is-section-fin-bounded-ℕ n) + ( is-retraction-fin-bounded-ℕ n) + equiv-count-bounded-ℕ : (n : ℕ) → Fin (succ-ℕ n) ≃ bounded-ℕ n -equiv-count-bounded-ℕ n = {!!} +pr1 (equiv-count-bounded-ℕ n) = bounded-nat-Fin n +pr2 (equiv-count-bounded-ℕ n) = is-equiv-bounded-nat-Fin n count-bounded-ℕ : (n : ℕ) → count (bounded-ℕ n) diff --git a/src/elementary-number-theory/number-of-divisors.lagda.md b/src/elementary-number-theory/number-of-divisors.lagda.md index 0c1adb4618..8ac0801357 100644 --- a/src/elementary-number-theory/number-of-divisors.lagda.md +++ b/src/elementary-number-theory/number-of-divisors.lagda.md @@ -8,13 +8,24 @@ module elementary-number-theory.number-of-divisors where ```agda open import elementary-number-theory.bounded-divisibility-natural-numbers +open import elementary-number-theory.bounded-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.squares-natural-numbers +open import foundation.action-on-identifications-functions open import foundation.decidable-subtypes open import foundation.dependent-pair-types +open import foundation.fixed-points-endofunctions +open import foundation.function-types open import foundation.identity-types +open import foundation.involutions +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes open import foundation.universe-levels open import univalent-combinatorics.counting @@ -46,24 +57,91 @@ The number of divisors function is listed as A000005 in the ```agda decidable-subtype-of-divisors-ℕ : - (n : ℕ) → decidable-subtype lzero (Fin (succ-ℕ n)) + (n : ℕ) → decidable-subtype lzero (bounded-ℕ n) decidable-subtype-of-divisors-ℕ n i = - bounded-div-ℕ-Decidable-Prop (nat-Fin (succ-ℕ n) i) n + bounded-div-ℕ-Decidable-Prop (nat-bounded-ℕ n i) n -type-of-divisors-ℕ : ℕ → UU lzero -type-of-divisors-ℕ n = +divisor-ℕ : + ℕ → UU lzero +divisor-ℕ n = type-decidable-subtype (decidable-subtype-of-divisors-ℕ n) -quotient-type-of-divisors-ℕ : - (n : ℕ) → type-of-divisors-ℕ n → type-of-divisors-ℕ n -pr1 (quotient-type-of-divisors-ℕ n (d , H)) = - mod-succ-ℕ n (quotient-bounded-div-ℕ (nat-Fin (succ-ℕ n) d) n H) -pr1 (pr2 (quotient-type-of-divisors-ℕ n (d , H))) = - nat-Fin (succ-ℕ n) d -pr1 (pr2 (pr2 (quotient-type-of-divisors-ℕ n (d , H)))) = - upper-bound-nat-Fin n d -pr2 (pr2 (pr2 (quotient-type-of-divisors-ℕ n (d , H)))) = - {!is-section-nat-Fin n!} +nat-divisor-ℕ : + (n : ℕ) → divisor-ℕ n → ℕ +nat-divisor-ℕ n = + nat-bounded-ℕ n ∘ pr1 + +bounded-div-divisor-ℕ : + (n : ℕ) (d : divisor-ℕ n) → bounded-div-ℕ (nat-divisor-ℕ n d) n +bounded-div-divisor-ℕ n = + pr2 + +div-divisor-ℕ : + (n : ℕ) (d : divisor-ℕ n) → div-ℕ (nat-divisor-ℕ n d) n +div-divisor-ℕ n d = + div-bounded-div-ℕ (nat-divisor-ℕ n d) n (bounded-div-divisor-ℕ n d) + +eq-divisor-ℕ : + (n : ℕ) (x y : divisor-ℕ n) → + nat-divisor-ℕ n x = nat-divisor-ℕ n y → x = y +eq-divisor-ℕ n x y p = + eq-type-subtype + ( λ i → bounded-div-ℕ-Prop (nat-bounded-ℕ n i) n) + ( eq-type-subtype + ( λ x → leq-ℕ-Prop x n) + ( p)) + +is-set-divisor-ℕ : + (n : ℕ) → is-set (divisor-ℕ n) +is-set-divisor-ℕ n = + is-set-type-subtype + ( λ i → bounded-div-ℕ-Prop (nat-bounded-ℕ n i) n) + ( is-set-bounded-ℕ n) + +divisor-ℕ-Set : + (n : ℕ) → Set lzero +pr1 (divisor-ℕ-Set n) = divisor-ℕ n +pr2 (divisor-ℕ-Set n) = is-set-divisor-ℕ n + +upper-bound-divisor-ℕ : + (n : ℕ) (d : divisor-ℕ n) → nat-divisor-ℕ n d ≤-ℕ n +upper-bound-divisor-ℕ n = + upper-bound-nat-bounded-ℕ n ∘ pr1 + +nat-quotient-divisor-ℕ : + (n : ℕ) → divisor-ℕ n → ℕ +nat-quotient-divisor-ℕ n (d , H) = + quotient-bounded-div-ℕ (nat-bounded-ℕ n d) n H + +upper-bound-quotient-divisor-ℕ : + (n : ℕ) (d : divisor-ℕ n) → nat-quotient-divisor-ℕ n d ≤-ℕ n +upper-bound-quotient-divisor-ℕ n (d , H) = + upper-bound-quotient-bounded-div-ℕ (nat-bounded-ℕ n d) n H + +eq-quotient-divisor-ℕ : + (n : ℕ) (d : divisor-ℕ n) → + nat-quotient-divisor-ℕ n d *ℕ nat-divisor-ℕ n d = n +eq-quotient-divisor-ℕ n d = + eq-quotient-bounded-div-ℕ (nat-divisor-ℕ n d) n (bounded-div-divisor-ℕ n d) + +eq-quotient-divisor-ℕ' : + (n : ℕ) (d : divisor-ℕ n) → + nat-divisor-ℕ n d *ℕ nat-quotient-divisor-ℕ n d = n +eq-quotient-divisor-ℕ' n d = + eq-quotient-bounded-div-ℕ' (nat-divisor-ℕ n d) n (bounded-div-divisor-ℕ n d) + +quotient-divisor-ℕ : + (n : ℕ) → divisor-ℕ n → divisor-ℕ n +pr1 (pr1 (quotient-divisor-ℕ n d)) = + nat-quotient-divisor-ℕ n d +pr2 (pr1 (quotient-divisor-ℕ n d)) = + upper-bound-quotient-divisor-ℕ n d +pr1 (pr2 (quotient-divisor-ℕ n d)) = + nat-divisor-ℕ n d +pr1 (pr2 (pr2 (quotient-divisor-ℕ n d))) = + upper-bound-divisor-ℕ n d +pr2 (pr2 (pr2 (quotient-divisor-ℕ n d))) = + eq-quotient-divisor-ℕ' n d ``` ### The number of divisors function @@ -72,18 +150,18 @@ Note that the entry `number-of-elements-decidable-subtype` refers to `count-decidable-subtype'`, which is `abstract` and therefore doesn't compute. ```agda -count-type-of-divisors-ℕ : - (n : ℕ) → count (type-of-divisors-ℕ n) -count-type-of-divisors-ℕ n = +count-divisor-ℕ : + (n : ℕ) → count (divisor-ℕ n) +count-divisor-ℕ n = count-decidable-subtype ( decidable-subtype-of-divisors-ℕ n) - ( count-Fin (succ-ℕ n)) + ( count-bounded-ℕ n) number-of-divisors-ℕ : ℕ → ℕ number-of-divisors-ℕ n = number-of-elements-decidable-subtype ( decidable-subtype-of-divisors-ℕ n) - ( count-Fin (succ-ℕ n)) + ( count-bounded-ℕ n) ``` ### Involution on the type of divisors @@ -93,5 +171,72 @@ its quotient is an involution. This operation has a fixed point precisely when `n` is a [square number](elementary-number-theory.square-natural-numbers.md). ```agda +is-involution-quotient-divisor-ℕ : + (n : ℕ) → is-involution (quotient-divisor-ℕ n) +is-involution-quotient-divisor-ℕ n d = + eq-divisor-ℕ n (quotient-divisor-ℕ n (quotient-divisor-ℕ n d)) d refl +``` + +### The involution on the type of divisors of `n` has a fixed point if and only if `n` is a square number + +```agda +is-square-has-fixed-point-quotient-divisor-ℕ : + (n : ℕ) → fixed-point (quotient-divisor-ℕ n) → is-square-ℕ n +pr1 (is-square-has-fixed-point-quotient-divisor-ℕ n (d , p)) = + nat-divisor-ℕ n d +pr2 (is-square-has-fixed-point-quotient-divisor-ℕ n (d , p)) = + ( inv (eq-quotient-divisor-ℕ n d)) ∙ + ( ap (_*ℕ nat-divisor-ℕ n d) (ap (nat-divisor-ℕ n) p)) + +square-root-divisor-is-square-ℕ : + (n : ℕ) (H : is-square-ℕ n) → divisor-ℕ n +square-root-divisor-is-square-ℕ ._ (n , refl) = + ( ( n , is-inflationary-square-ℕ n) , + ( is-inflationary-bounded-div-square-ℕ n)) + +is-fixed-point-square-root-divisor-is-square-ℕ : + (n : ℕ) (H : is-square-ℕ n) → + quotient-divisor-ℕ n (square-root-divisor-is-square-ℕ n H) = + square-root-divisor-is-square-ℕ n H +is-fixed-point-square-root-divisor-is-square-ℕ ._ H@(n , refl) = + eq-divisor-ℕ + ( square-ℕ n) + ( quotient-divisor-ℕ _ (square-root-divisor-is-square-ℕ _ H)) + ( square-root-divisor-is-square-ℕ _ H) + ( refl) + +has-fixed-point-quotient-divisor-is-square-ℕ : + (n : ℕ) → is-square-ℕ n → fixed-point (quotient-divisor-ℕ n) +pr1 (has-fixed-point-quotient-divisor-is-square-ℕ n H) = + square-root-divisor-is-square-ℕ n H +pr2 (has-fixed-point-quotient-divisor-is-square-ℕ n H) = + is-fixed-point-square-root-divisor-is-square-ℕ n H +``` +### The type of fixed points of the involution on the type of divisors is a proposition + +This essentially claims that only the square root of a square can be the fixed point of the involution on the type of divisors. + +```agda +all-elements-equal-fixed-point-quotient-divisor-ℕ : + (n : ℕ) → all-elements-equal (fixed-point (quotient-divisor-ℕ n)) +all-elements-equal-fixed-point-quotient-divisor-ℕ n (x , p) (y , q) = + eq-type-subtype + ( λ d → Id-Prop (divisor-ℕ-Set n) (quotient-divisor-ℕ n d) d) + ( eq-divisor-ℕ n x y + ( is-injective-square-ℕ + ( ( ap (nat-divisor-ℕ n x *ℕ_) (ap (nat-divisor-ℕ n) (inv p))) ∙ + ( eq-quotient-divisor-ℕ' n x) ∙ + ( inv (eq-quotient-divisor-ℕ' n y)) ∙ + ( ap (nat-divisor-ℕ n y *ℕ_) (ap (nat-divisor-ℕ n) q))))) + +is-prop-fixed-point-quotient-divisor-ℕ : + (n : ℕ) → is-prop (fixed-point (quotient-divisor-ℕ n)) +is-prop-fixed-point-quotient-divisor-ℕ n = + is-prop-all-elements-equal + ( all-elements-equal-fixed-point-quotient-divisor-ℕ n) ``` + +## Remarks + +In the properties above we have shown that the quotient operation on the type of divisors of any natural number $n$ is an involution with at most one fixed point, and it has a fixed point if and only if $n$ is square. This implies that the number of divisors is odd if and only if $n$ is a square. However, in the agda-unimath library we don't have the appropriate infrastructure yet for counting the elements of types with involutions. The conclusion that the number of divisors of $n$ is odd if and only if $n$ is a square will be formalized in the future. diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index b605b399ae..7a9d7077d9 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -28,6 +28,7 @@ open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.identity-types +open import foundation.injective-maps open import foundation.negation open import foundation.unit-type open import foundation.universe-levels @@ -201,6 +202,16 @@ le-zero-le-zero-square-ℕ n = reflects-strict-order-square-ℕ 0 n ``` +### The square function is injective + +```agda +is-injective-square-ℕ : is-injective square-ℕ +is-injective-square-ℕ {x} {y} p = + antisymmetric-leq-ℕ x y + ( reflects-order-square-ℕ x y (leq-eq-ℕ (square-ℕ x) (square-ℕ y) p)) + ( reflects-order-square-ℕ y x (leq-eq-ℕ (square-ℕ y) (square-ℕ x) (inv p))) +``` + ### Squares distribute over multiplication ```agda @@ -210,44 +221,18 @@ distributive-square-mul-ℕ m n = interchange-law-mul-mul-ℕ m n m n ``` -### Any number is less than or equal to its own square - -**Proof.** The proof is by induction. Since `0` is the least natural number, be -base case is trivial. Now consider a natural number `n` such that `n ≤ n²`. Then -we have - -```text - (n + 1 ≤ (n + 1)²) ↔ n + 1 ≤ (n + 2) * n + 1 - ↔ n ≤ n² + n + n. -``` - -The last inequality follows by the following chain of inequalities - -```text - n ≤ n² -- by the induction hypothesis - ≤ n² + n -- since a ≤ a + b for any a,b - ≤ n² + n + n -- since a ≤ a + b for any a,b -``` +### The square function is inflationary ```agda -lower-bound-square-ℕ : +is-inflationary-square-ℕ : (n : ℕ) → n ≤-ℕ square-ℕ n -lower-bound-square-ℕ zero-ℕ = star -lower-bound-square-ℕ (succ-ℕ n) = - concatenate-leq-eq-ℕ - ( succ-ℕ n) - ( transitive-leq-ℕ - ( n) - ( square-ℕ n) - ( square-ℕ n +ℕ n +ℕ n) - ( transitive-leq-ℕ - ( square-ℕ n) - ( square-ℕ n +ℕ n) - ( square-ℕ n +ℕ n +ℕ n) - ( leq-add-ℕ (square-ℕ n +ℕ n) n) - ( leq-add-ℕ (square-ℕ n) n)) - ( lower-bound-square-ℕ n)) - ( inv (square-succ-ℕ' n)) +is-inflationary-square-ℕ zero-ℕ = + star +is-inflationary-square-ℕ (succ-ℕ n) = + concatenate-eq-leq-ℕ + ( square-ℕ (succ-ℕ n)) + ( inv (right-unit-law-mul-ℕ (succ-ℕ n))) + ( preserves-order-right-mul-ℕ (succ-ℕ n) 1 (succ-ℕ n) star) ``` ### If a number `n` has a square root, then the square root is at most `n` @@ -256,16 +241,16 @@ lower-bound-square-ℕ (succ-ℕ n) = upper-bound-square-root-ℕ : (n : ℕ) (H : is-square-ℕ n) → square-root-ℕ n H ≤-ℕ n upper-bound-square-root-ℕ .(square-ℕ x) (x , refl) = - lower-bound-square-ℕ x + is-inflationary-square-ℕ x ``` ### Any number greater than 1 is strictly less than its square ```agda -strict-lower-bound-square-ℕ : +strict-is-inflationary-square-ℕ : (n : ℕ) → 1 <-ℕ n → n <-ℕ square-ℕ n -strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ zero-ℕ)) H = star -strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = +strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ zero-ℕ)) H = star +strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = concatenate-le-eq-ℕ ( n +ℕ 3) ( (n +ℕ 4) *ℕ (n +ℕ 2) +ℕ 1) @@ -274,7 +259,7 @@ strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = ( n +ℕ 2) ( square-ℕ (n +ℕ 2)) ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) - ( strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ n)) star) + ( strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ n)) star) ( transitive-le-ℕ ( square-ℕ (n +ℕ 2)) ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) @@ -290,7 +275,7 @@ strict-lower-bound-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = strict-upper-bound-square-root-ℕ : (n : ℕ) → 1 <-ℕ n → (H : is-square-ℕ n) → square-root-ℕ n H <-ℕ n strict-upper-bound-square-root-ℕ ._ B (succ-ℕ (succ-ℕ x) , refl) = - strict-lower-bound-square-ℕ (x +ℕ 2) star + strict-is-inflationary-square-ℕ (x +ℕ 2) star ``` ### If `n² ≤ n`, then `n ≤ 1` @@ -305,7 +290,7 @@ leq-one-leq-square-ℕ (succ-ℕ (succ-ℕ n)) H = ( contradiction-le-ℕ ( n +ℕ 2) ( square-ℕ (n +ℕ 2)) - ( strict-lower-bound-square-ℕ (n +ℕ 2) star) + ( strict-is-inflationary-square-ℕ (n +ℕ 2) star) ( H)) ``` @@ -328,7 +313,7 @@ not-le-square-ℕ n H = ( square-ℕ n) ( n) ( H) - ( lower-bound-square-ℕ n) + ( is-inflationary-square-ℕ n) ``` ### Being a square natural number is decidable @@ -352,6 +337,23 @@ is-decidable-is-square-ℕ n = ( is-decidable-has-greater-root-ℕ n) ``` +### Any number divides its own square + +In other words, the squaring function is inflationary with respect to divisibility. + +```agda +is-inflationary-bounded-div-square-ℕ : + (n : ℕ) → bounded-div-ℕ n (square-ℕ n) +pr1 (is-inflationary-bounded-div-square-ℕ n) = n +pr1 (pr2 (is-inflationary-bounded-div-square-ℕ n)) = is-inflationary-square-ℕ n +pr2 (pr2 (is-inflationary-bounded-div-square-ℕ n)) = refl + +is-inflationary-div-square-ℕ : + (n : ℕ) → div-ℕ n (square-ℕ n) +is-inflationary-div-square-ℕ n = + div-bounded-div-ℕ n (square-ℕ n) (is-inflationary-bounded-div-square-ℕ n) +``` + ### Equivalent characterizations of a number being even in terms of its square Consider a natural number `n`. The following are equivalent: diff --git a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md index d680cd9a65..9e7dc01d97 100644 --- a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md @@ -7,11 +7,18 @@ module elementary-number-theory.strictly-bounded-natural-numbers where
Imports ```agda +open import elementary-number-theory.congruence-natural-numbers +open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.retractions +open import foundation.sections +open import foundation.sets +open import foundation.subtypes open import foundation.universe-levels open import univalent-combinatorics.counting @@ -37,16 +44,79 @@ The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) ```agda strictly-bounded-ℕ : ℕ → UU lzero strictly-bounded-ℕ n = Σ ℕ (λ k → k <-ℕ n) + +nat-strictly-bounded-ℕ : + (n : ℕ) → strictly-bounded-ℕ n → ℕ +nat-strictly-bounded-ℕ n = + pr1 + +strict-upper-bound-nat-strictly-bounded-ℕ : + (n : ℕ) (x : strictly-bounded-ℕ n) → nat-strictly-bounded-ℕ n x <-ℕ n +strict-upper-bound-nat-strictly-bounded-ℕ n = + pr2 ``` ## Properties +### The type of bounded natural numbers is a set + +```agda +is-set-strictly-bounded-ℕ : + (n : ℕ) → is-set (strictly-bounded-ℕ n) +is-set-strictly-bounded-ℕ n = + is-set-type-subtype + ( λ k → le-ℕ-Prop k n) + ( is-set-ℕ) +``` + ### The type $\mathbb{N}_{\leq b}$ is equivalent to the standard finite type $\mathsf{Fin}(n+1)$ ```agda +strictly-bounded-nat-Fin : + (n : ℕ) → Fin n → strictly-bounded-ℕ n +pr1 (strictly-bounded-nat-Fin n i) = nat-Fin n i +pr2 (strictly-bounded-nat-Fin n i) = strict-upper-bound-nat-Fin n i + +fin-strictly-bounded-ℕ : + (n : ℕ) → strictly-bounded-ℕ n → Fin n +fin-strictly-bounded-ℕ (succ-ℕ n) x = + mod-succ-ℕ n (nat-strictly-bounded-ℕ (succ-ℕ n) x) + +is-section-fin-strictly-bounded-ℕ : + (n : ℕ) → is-section (strictly-bounded-nat-Fin n) (fin-strictly-bounded-ℕ n) +is-section-fin-strictly-bounded-ℕ (succ-ℕ n) i = + eq-type-subtype + ( λ x → le-ℕ-Prop x (succ-ℕ n)) + ( eq-cong-le-ℕ + ( succ-ℕ n) + ( nat-Fin (succ-ℕ n) (fin-strictly-bounded-ℕ (succ-ℕ n) i)) + ( nat-strictly-bounded-ℕ (succ-ℕ n) i) + ( strict-upper-bound-nat-Fin + ( succ-ℕ n) + ( fin-strictly-bounded-ℕ (succ-ℕ n) i)) + ( strict-upper-bound-nat-strictly-bounded-ℕ (succ-ℕ n) i) + ( cong-nat-mod-succ-ℕ n (nat-strictly-bounded-ℕ (succ-ℕ n) i))) + +is-retraction-fin-strictly-bounded-ℕ : + (n : ℕ) → + is-retraction (strictly-bounded-nat-Fin n) (fin-strictly-bounded-ℕ n) +is-retraction-fin-strictly-bounded-ℕ (succ-ℕ n) i = + is-section-nat-Fin n i + +is-equiv-strictly-bounded-nat-Fin : + (n : ℕ) → is-equiv (strictly-bounded-nat-Fin n) +is-equiv-strictly-bounded-nat-Fin n = + is-equiv-is-invertible + ( fin-strictly-bounded-ℕ n) + ( is-section-fin-strictly-bounded-ℕ n) + ( is-retraction-fin-strictly-bounded-ℕ n) + equiv-count-strictly-bounded-ℕ : (n : ℕ) → Fin n ≃ strictly-bounded-ℕ n -equiv-count-strictly-bounded-ℕ n = {!!} +pr1 (equiv-count-strictly-bounded-ℕ n) = + strictly-bounded-nat-Fin n +pr2 (equiv-count-strictly-bounded-ℕ n) = + is-equiv-strictly-bounded-nat-Fin n count-strictly-bounded-ℕ : (n : ℕ) → count (strictly-bounded-ℕ n) diff --git a/src/literature/oeis.lagda.md b/src/literature/oeis.lagda.md index 3a159097a8..bc9af5e560 100644 --- a/src/literature/oeis.lagda.md +++ b/src/literature/oeis.lagda.md @@ -45,6 +45,8 @@ A000004 : ℕ → ℕ A000004 _ = zero-ℕ ``` +OEIS: [A000004](https://oeis.org/A000004) + ### A000005 -- The number of divisors ```agda @@ -52,7 +54,7 @@ open import elementary-number-theory.number-of-divisors using ( number-of-divisors-ℕ) ``` -OEIS: [A000004](https://oeis.org/A000004) +OEIS: [A000005](https://oeis.org/A000005) ### A000007 -- The characteristic function for 0 diff --git a/src/univalent-combinatorics/counting.lagda.md b/src/univalent-combinatorics/counting.lagda.md index a7abb0d0c1..9755b78b0a 100644 --- a/src/univalent-combinatorics/counting.lagda.md +++ b/src/univalent-combinatorics/counting.lagda.md @@ -34,10 +34,12 @@ open import univalent-combinatorics.standard-finite-types ## Idea -The elements of a type `X` can be counted by establishing an equivalence +A {{#concept "count" Agda=count WDID=Q247154 WD=counting}} of the elements of a type `X` consists of a [natural number](elementary-number-theory.natural-numbers.md) `n` and an [equivalence](foundation-core.equivalences.md) `Fin n ≃ X`. -## Definition +## Definitions + +### Counts of a type ```agda count : {l : Level} → UU l → UU l @@ -217,3 +219,7 @@ count-type-trunc-Prop (pair (succ-ℕ k) e) = ( is-prop-type-trunc-Prop) ( unit-trunc-Prop (map-equiv e (zero-Fin k)))) ``` + +## See also + +- [Finite types](univalent-combinatorics.finite-types.md) From ed68ab3d6a62ca827d5d9fce0b5ff019ba637d40 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 6 Jan 2025 23:36:40 -0500 Subject: [PATCH 38/72] The formula for the distance of squares --- .../fermat-numbers.lagda.md | 2 +- .../multiplication-natural-numbers.lagda.md | 10 ++++ .../squares-natural-numbers.lagda.md | 58 +++++++++++++++++++ 3 files changed, 69 insertions(+), 1 deletion(-) diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index 3e169f638b..d753d37a77 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -98,7 +98,7 @@ compute-succ-recursive-fermat-number-ℕ = ## Properties -### The Fermat number at a successor +### The Fermat number of a successor Any Fermat number of the form $F(n+1)=2^{2^{n+1}}+1$ can be computed as diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md index 0cc4824158..fbee2ce1fd 100644 --- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md @@ -183,6 +183,16 @@ abstract ( ( left-distributive-mul-add-ℕ z x y) ∙ ( ( ap (_+ℕ (z *ℕ y)) (commutative-mul-ℕ z x)) ∙ ( ap ((x *ℕ z) +ℕ_) (commutative-mul-ℕ z y)))) + +abstract + double-distributive-mul-add-ℕ : + (w x y z : ℕ) → (w +ℕ x) *ℕ (y +ℕ z) = w *ℕ y +ℕ w *ℕ z +ℕ x *ℕ y +ℕ x *ℕ z + double-distributive-mul-add-ℕ w x y z = + ( right-distributive-mul-add-ℕ w x (y +ℕ z)) ∙ + ( ap-add-ℕ + ( left-distributive-mul-add-ℕ w y z) + ( left-distributive-mul-add-ℕ x y z)) ∙ + ( inv (associative-add-ℕ (w *ℕ y +ℕ w *ℕ z) (x *ℕ y) (x *ℕ z))) ``` ### Multiplication is associative diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 7a9d7077d9..2e8fa31dda 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -559,6 +559,64 @@ module _ ( bounded-div-square-ℕ) ``` +### Squares of sums of natural numbers + +```agda +square-add-ℕ : + (m n : ℕ) → square-ℕ (m +ℕ n) = square-ℕ m +ℕ 2 *ℕ (m *ℕ n) +ℕ square-ℕ n +square-add-ℕ m n = + ( double-distributive-mul-add-ℕ m n m n) ∙ + ( ap + ( _+ℕ square-ℕ n) + ( ( associative-add-ℕ (square-ℕ m) (m *ℕ n) (n *ℕ m)) ∙ + ( ap + ( square-ℕ m +ℕ_) + ( ( ap (m *ℕ n +ℕ_) (commutative-mul-ℕ n m)) ∙ + ( inv (left-two-law-mul-ℕ (m *ℕ n))))))) +``` + +### The formula for the distance between squares + +The formula for the distance between squares is more commonly known as the formula for the difference of squares. However, since we prefer using the distance operation on the natural numbers over the partial difference operation, we will state and prove the analogous formula using the distance function. + +The formula for the difference of squares of integers is formalized in its usual form. + +```agda +distance-of-squares-ℕ' : + (m n k : ℕ) → m +ℕ k = n → + dist-ℕ (square-ℕ m) (square-ℕ n) = dist-ℕ m n *ℕ (m +ℕ n) +distance-of-squares-ℕ' m ._ k refl = + ( ap + ( dist-ℕ (square-ℕ m)) + ( ( square-add-ℕ m k) ∙ + ( associative-add-ℕ (square-ℕ m) (2 *ℕ (m *ℕ k)) (square-ℕ k)))) ∙ + ( dist-add-ℕ (square-ℕ m) (2 *ℕ (m *ℕ k) +ℕ square-ℕ k)) ∙ + ( ap + ( _+ℕ square-ℕ k) + ( inv (associative-mul-ℕ 2 m k) ∙ ap (_*ℕ k) (left-two-law-mul-ℕ m))) ∙ + ( inv (right-distributive-mul-add-ℕ (m +ℕ m) k k)) ∙ + ( ap (_*ℕ k) (associative-add-ℕ m m k)) ∙ + ( commutative-mul-ℕ (m +ℕ (m +ℕ k)) k) ∙ + ( ap (_*ℕ (m +ℕ (m +ℕ k))) (inv (dist-add-ℕ m k))) + +abstract + distance-of-squares-ℕ : + (m n : ℕ) → dist-ℕ (square-ℕ m) (square-ℕ n) = dist-ℕ m n *ℕ (m +ℕ n) + distance-of-squares-ℕ m n + with + linear-leq-ℕ m n + ... | inl H = + distance-of-squares-ℕ' m n + ( pr1 (subtraction-leq-ℕ m n H)) + ( commutative-add-ℕ m _ ∙ pr2 (subtraction-leq-ℕ m n H)) + ... | inr H = + ( symmetric-dist-ℕ (square-ℕ m) (square-ℕ n)) ∙ + ( distance-of-squares-ℕ' n m + ( pr1 (subtraction-leq-ℕ n m H)) + ( commutative-add-ℕ n _ ∙ pr2 (subtraction-leq-ℕ n m H))) ∙ + ( ap-mul-ℕ (symmetric-dist-ℕ n m) (commutative-add-ℕ n m)) +``` + ### The $n$th square is the sum of the first $n$ odd numbers We show that From e0a8dee9a54d28f40c2e07351754b449d4b26288 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 7 Jan 2025 15:58:18 -0500 Subject: [PATCH 39/72] refactoring exponentiation --- ...of-elements-commutative-semirings.lagda.md | 69 +++++++++++- .../exponentiation-natural-numbers.lagda.md | 102 +++++++++++------- .../fermat-numbers.lagda.md | 11 +- .../squares-natural-numbers.lagda.md | 11 ++ .../powers-of-elements-monoids.lagda.md | 12 ++- .../powers-of-elements-semirings.lagda.md | 19 ++++ 6 files changed, 183 insertions(+), 41 deletions(-) diff --git a/src/commutative-algebra/powers-of-elements-commutative-semirings.lagda.md b/src/commutative-algebra/powers-of-elements-commutative-semirings.lagda.md index 0bda772d1a..0ae3eadf64 100644 --- a/src/commutative-algebra/powers-of-elements-commutative-semirings.lagda.md +++ b/src/commutative-algebra/powers-of-elements-commutative-semirings.lagda.md @@ -8,8 +8,10 @@ module commutative-algebra.powers-of-elements-commutative-semirings where ```agda open import commutative-algebra.commutative-semirings +open import commutative-algebra.homomorphisms-commutative-semirings open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import foundation.identity-types @@ -36,6 +38,21 @@ power-Commutative-Semiring A = power-Semiring (semiring-Commutative-Semiring A) ## Properties +### `1ⁿ = 1` + +```agda +module _ + {l : Level} (A : Commutative-Semiring l) + where + + power-one-Commutative-Semiring : + (n : ℕ) → + power-Commutative-Semiring A n (one-Commutative-Semiring A) = + one-Commutative-Semiring A + power-one-Commutative-Semiring = + power-one-Semiring (semiring-Commutative-Semiring A) +``` + ### `xⁿ⁺¹ = xⁿx` ```agda @@ -51,6 +68,21 @@ module _ power-succ-Semiring (semiring-Commutative-Semiring A) ``` +### `xⁿ⁺¹ = xxⁿ` + +```agda +module _ + {l : Level} (A : Commutative-Semiring l) + where + + power-succ-Commutative-Semiring' : + (n : ℕ) (x : type-Commutative-Semiring A) → + power-Commutative-Semiring A (succ-ℕ n) x = + mul-Commutative-Semiring A x (power-Commutative-Semiring A n x) + power-succ-Commutative-Semiring' = + power-succ-Semiring' (semiring-Commutative-Semiring A) +``` + ### Powers by sums of natural numbers are products of powers ```agda @@ -68,7 +100,7 @@ module _ distributive-power-add-Semiring (semiring-Commutative-Semiring A) ``` -### If `x` commutes with `y`, then powers distribute over the product of `x` and `y` +### Powers distribute over products ```agda module _ @@ -87,3 +119,38 @@ module _ ( n) ( commutative-mul-Commutative-Semiring A x y) ``` + +### Powers by products of natural numbers are iterated powers + +```agda +module _ + {l : Level} (A : Commutative-Semiring l) + where + + power-mul-Commutative-Semiring : + (m n : ℕ) {x : type-Commutative-Semiring A} → + power-Commutative-Semiring A (m *ℕ n) x = + power-Commutative-Semiring A n (power-Commutative-Semiring A m x) + power-mul-Commutative-Semiring = + power-mul-Semiring (semiring-Commutative-Semiring A) +``` + +### Homomorphisms of semirings preserve powers + +```agda +module _ + {l1 l2 : Level} + (A : Commutative-Semiring l1) (B : Commutative-Semiring l2) + (f : hom-Commutative-Semiring A B) + where + + preserves-powers-hom-Commutative-Semiring : + (n : ℕ) (x : type-Commutative-Semiring A) → + map-hom-Commutative-Semiring A B f (power-Commutative-Semiring A n x) = + power-Commutative-Semiring B n (map-hom-Commutative-Semiring A B f x) + preserves-powers-hom-Commutative-Semiring = + preserves-powers-hom-Semiring + ( semiring-Commutative-Semiring A) + ( semiring-Commutative-Semiring B) + ( f) +``` diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index e852edcb4a..a569184553 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -29,55 +29,75 @@ The {{#concept "exponential" Agda=exp-ℕ WD="exponentiation" WDID=Q33456}} $m^n of two [natural numbers](elementary-number-theory.natural-numbers.md) is the number obtained by multiplying $m$ with itself $n$ times. +The natural numbers satisfy Tarski's high school arithmetic laws for exponentiation. + ## Definition +### Exponentiation of natural numbers + ```agda +power-ℕ : ℕ → ℕ → ℕ +power-ℕ = power-Commutative-Semiring ℕ-Commutative-Semiring + exp-ℕ : ℕ → ℕ → ℕ -exp-ℕ m 0 = 1 -exp-ℕ m (succ-ℕ n) = (exp-ℕ m n) *ℕ m +exp-ℕ m n = power-ℕ n m infixr 45 _^ℕ_ _^ℕ_ = exp-ℕ ``` +## Properties + +### `1ⁿ = 1` + ```agda -power-ℕ : ℕ → ℕ → ℕ -power-ℕ = power-Commutative-Semiring ℕ-Commutative-Semiring +annihilation-law-exp-ℕ : (n : ℕ) → 1 ^ℕ n = 1 +annihilation-law-exp-ℕ = power-one-Commutative-Semiring ℕ-Commutative-Semiring ``` -## Properties +### `xⁿ⁺¹ = xⁿx` -### Tarski's high school arithmetic laws for exponentiation +```agda +exp-succ-ℕ : + (n : ℕ) (x : ℕ) → x ^ℕ succ-ℕ n = x ^ℕ n *ℕ x +exp-succ-ℕ = + power-succ-Commutative-Semiring ℕ-Commutative-Semiring +``` + +### `xⁿ⁺¹ = xxⁿ` ```agda -right-unit-law-exp-ℕ : (n : ℕ) → exp-ℕ n 1 = n -right-unit-law-exp-ℕ n = left-unit-law-mul-ℕ n +exp-succ-ℕ' : + (n : ℕ) (x : ℕ) → x ^ℕ succ-ℕ n = x *ℕ x ^ℕ n +exp-succ-ℕ' = + power-succ-Commutative-Semiring' ℕ-Commutative-Semiring +``` -annihilation-law-exp-ℕ : (n : ℕ) → 1 ^ℕ n = 1 -annihilation-law-exp-ℕ zero-ℕ = refl -annihilation-law-exp-ℕ (succ-ℕ n) = - right-unit-law-mul-ℕ (1 ^ℕ n) ∙ annihilation-law-exp-ℕ n +### Powers by sums of natural numbers are products of powers +```agda left-distributive-exp-add-ℕ : - (x y z : ℕ) → x ^ℕ (y +ℕ z) = (x ^ℕ y) *ℕ (x ^ℕ z) -left-distributive-exp-add-ℕ x y zero-ℕ = inv (right-unit-law-mul-ℕ (x ^ℕ y)) -left-distributive-exp-add-ℕ x y (succ-ℕ z) = - ( ap (_*ℕ x) (left-distributive-exp-add-ℕ x y z)) ∙ - ( associative-mul-ℕ (x ^ℕ y) (x ^ℕ z) x) + (m n : ℕ) {x : ℕ} → x ^ℕ (m +ℕ n) = x ^ℕ m *ℕ x ^ℕ n +left-distributive-exp-add-ℕ = + distributive-power-add-Commutative-Semiring ℕ-Commutative-Semiring +``` +### Powers distribute over products + +```agda right-distributive-exp-mul-ℕ : - (x y z : ℕ) → (x *ℕ y) ^ℕ z = (x ^ℕ z) *ℕ (y ^ℕ z) -right-distributive-exp-mul-ℕ x y zero-ℕ = refl -right-distributive-exp-mul-ℕ x y (succ-ℕ z) = - ( ap (_*ℕ (x *ℕ y)) (right-distributive-exp-mul-ℕ x y z)) ∙ - ( interchange-law-mul-mul-ℕ (x ^ℕ z) (y ^ℕ z) x y) - -exp-mul-ℕ : (x y z : ℕ) → x ^ℕ (y *ℕ z) = (x ^ℕ y) ^ℕ z -exp-mul-ℕ x zero-ℕ z = inv (annihilation-law-exp-ℕ z) -exp-mul-ℕ x (succ-ℕ y) z = - ( left-distributive-exp-add-ℕ x (y *ℕ z) z) ∙ - ( ( ap (_*ℕ (x ^ℕ z)) (exp-mul-ℕ x y z)) ∙ - ( inv (right-distributive-exp-mul-ℕ (x ^ℕ y) x z))) + (n : ℕ) (x y : ℕ) → (x *ℕ y) ^ℕ n = x ^ℕ n *ℕ y ^ℕ n +right-distributive-exp-mul-ℕ = + distributive-power-mul-Commutative-Semiring ℕ-Commutative-Semiring +``` + +### Powers by products of natural numbers are iterated powers + +```agda +exp-mul-ℕ : + (m n : ℕ) {x : ℕ} → x ^ℕ (m *ℕ n) = (x ^ℕ m) ^ℕ n +exp-mul-ℕ = + power-mul-Commutative-Semiring ℕ-Commutative-Semiring ``` ### The exponent $m^n$ is nonzero if $m$ is nonzero @@ -85,17 +105,26 @@ exp-mul-ℕ x (succ-ℕ y) z = ```agda is-nonzero-exp-ℕ : (m n : ℕ) → is-nonzero-ℕ m → is-nonzero-ℕ (m ^ℕ n) -is-nonzero-exp-ℕ m zero-ℕ p = +is-nonzero-exp-ℕ m zero-ℕ H = is-nonzero-one-ℕ -is-nonzero-exp-ℕ m (succ-ℕ n) p = - is-nonzero-mul-ℕ (m ^ℕ n) m (is-nonzero-exp-ℕ m n p) p +is-nonzero-exp-ℕ m (succ-ℕ zero-ℕ) H = H +is-nonzero-exp-ℕ m (succ-ℕ (succ-ℕ n)) H = + is-nonzero-mul-ℕ (m ^ℕ succ-ℕ n) m (is-nonzero-exp-ℕ m (succ-ℕ n) H) H le-zero-exp-ℕ : (m n : ℕ) → 0 <-ℕ m → 0 <-ℕ m ^ℕ n le-zero-exp-ℕ m zero-ℕ H = star -le-zero-exp-ℕ m (succ-ℕ n) H = - preserves-strict-order-mul-ℕ 0 (m ^ℕ n) 0 m (le-zero-exp-ℕ m n H) H +le-zero-exp-ℕ m (succ-ℕ zero-ℕ) H = + H +le-zero-exp-ℕ m (succ-ℕ (succ-ℕ n)) H = + preserves-strict-order-mul-ℕ + ( 0) + ( m ^ℕ succ-ℕ n) + ( 0) + ( m) + ( le-zero-exp-ℕ m (succ-ℕ n) H) + ( H) ``` ### The exponent $m^n$ is equal to the $n$-fold product of $m$ @@ -104,6 +133,7 @@ le-zero-exp-ℕ m (succ-ℕ n) H = compute-constant-product-ℕ : (m n : ℕ) → Π-ℕ n (λ _ → m) = exp-ℕ m n compute-constant-product-ℕ m zero-ℕ = refl -compute-constant-product-ℕ m (succ-ℕ n) = - ap (_*ℕ m) (compute-constant-product-ℕ m n) +compute-constant-product-ℕ m (succ-ℕ zero-ℕ) = left-unit-law-add-ℕ m +compute-constant-product-ℕ m (succ-ℕ (succ-ℕ n)) = + ap (_*ℕ m) (compute-constant-product-ℕ m (succ-ℕ n)) ``` diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index d753d37a77..e94f631fc8 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -8,6 +8,7 @@ module elementary-number-theory.fermat-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.natural-numbers @@ -106,9 +107,13 @@ $$ F(n+1)-2=(2^{2^n})^2-1=(2^{2^n}+1)(2^{2^n}-1)=F(n)(F(n)-2). $$ -```text -fermat-number-succ-ℕ : - (n : ℕ) → fermat-number-ℕ (succ-ℕ n) = +```agda +dist-fermat-number-succ-two-ℕ : + (n : ℕ) → + dist-ℕ 2 (fermat-number-ℕ (succ-ℕ n)) = + Π-ℕ (succ-ℕ n) (λ k → fermat-number-ℕ (nat-Fin (succ-ℕ n) k)) +dist-fermat-number-succ-two-ℕ n = + {!!} ``` ### The two definitions of the Fermat numbers agree diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 2e8fa31dda..d3448fa6d2 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -14,6 +14,7 @@ open import elementary-number-theory.decidable-types open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -78,6 +79,16 @@ square-root-ℕ _ (root , _) = root ## Properties +### Squares are second powers + +The equality `square-ℕ n = n ^ℕ 2` holds by definition. + +```agda +compute-exp-2-ℕ : + (n : ℕ) → square-ℕ n = n ^ℕ 2 +compute-exp-2-ℕ n = refl +``` + ### Squares of successors For any `n` we have the equations diff --git a/src/group-theory/powers-of-elements-monoids.lagda.md b/src/group-theory/powers-of-elements-monoids.lagda.md index 66c4e4f41b..e9d0e489f9 100644 --- a/src/group-theory/powers-of-elements-monoids.lagda.md +++ b/src/group-theory/powers-of-elements-monoids.lagda.md @@ -25,10 +25,20 @@ open import group-theory.monoids ## Idea -The **power operation** on a [monoid](group-theory.monoids.md) is the map +The {{#concept "power operation" Disambiguation="monoid" Agda=power-Monoid}} on a [monoid](group-theory.monoids.md) is the map `n x ↦ xⁿ`, which is defined by [iteratively](foundation.iterating-functions.md) multiplying `x` with itself `n` times. +We define the power operation such that the following equalities hold by definition: + +```text + x⁰ ≐ 1 + x¹ ≐ x + xⁿ⁺² ≐ xⁿ⁺¹ · x. +``` + +This setup requires one extra step for the most basic properties, but it avoids having a superficial factor of the multiplicative unit in its definition for `n ≥ 1`. + ## Definitions ### Powers of elements of monoids diff --git a/src/ring-theory/powers-of-elements-semirings.lagda.md b/src/ring-theory/powers-of-elements-semirings.lagda.md index ca2f29b47c..a9b8996b46 100644 --- a/src/ring-theory/powers-of-elements-semirings.lagda.md +++ b/src/ring-theory/powers-of-elements-semirings.lagda.md @@ -16,6 +16,7 @@ open import foundation.universe-levels open import group-theory.powers-of-elements-monoids +open import ring-theory.homomorphisms-semirings open import ring-theory.semirings ``` @@ -145,3 +146,21 @@ module _ power-Semiring R (m *ℕ n) x = power-Semiring R n (power-Semiring R m x) power-mul-Semiring = power-mul-Monoid (multiplicative-monoid-Semiring R) ``` + +### Homomorphisms of semirings preserve powers + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (S : Semiring l2) (f : hom-Semiring R S) + where + + preserves-powers-hom-Semiring : + (n : ℕ) (x : type-Semiring R) → + map-hom-Semiring R S f (power-Semiring R n x) = + power-Semiring S n (map-hom-Semiring R S f x) + preserves-powers-hom-Semiring = + preserves-powers-hom-Monoid + ( multiplicative-monoid-Semiring R) + ( multiplicative-monoid-Semiring S) + ( hom-multiplicative-monoid-hom-Semiring R S f) +``` From 1e2dc898abde8e7fe8d8df9501518addba6f7623 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 8 Jan 2025 12:22:18 -0500 Subject: [PATCH 40/72] strictly ordered sets and inflationary maps --- .../directed-families-posets.lagda.md | 4 +- src/elementary-number-theory.lagda.md | 2 +- ...lagda.md => 2-adic-decomposition.lagda.md} | 170 +++++++++++++++++- .../distance-natural-numbers.lagda.md | 2 +- .../exponentiation-natural-numbers.lagda.md | 20 +++ .../fermat-numbers.lagda.md | 41 ++++- .../parity-natural-numbers.lagda.md | 87 ++++++++- .../squares-natural-numbers.lagda.md | 2 +- .../type-arithmetic-natural-numbers.lagda.md | 2 +- ...rdering-principle-natural-numbers.lagda.md | 5 + src/foundation/binary-relations.lagda.md | 10 +- src/foundation/iterating-functions.lagda.md | 9 +- src/order-theory/chains-posets.lagda.md | 4 +- .../deflationary-maps-posets.lagda.md | 145 +++++++++++++++ .../deflationary-maps-preorders.lagda.md | 146 +++++++++++++++ src/order-theory/galois-connections.lagda.md | 4 +- .../homomorphisms-large-suplattices.lagda.md | 4 +- .../inflationary-maps-posets.lagda.md | 145 +++++++++++++++ .../inflationary-maps-preorders.lagda.md | 146 +++++++++++++++ src/order-theory/opposite-preorders.lagda.md | 2 +- .../order-preserving-maps-posets.lagda.md | 6 +- .../order-preserving-maps-preorders.lagda.md | 58 +++--- .../strictly-ordered-sets.lagda.md | 115 ++++++++++++ .../strictly-ordered-types.lagda.md | 92 ++++++++++ .../powers-of-elements-rings.lagda.md | 12 ++ .../powers-of-elements-semirings.lagda.md | 14 ++ 26 files changed, 1181 insertions(+), 66 deletions(-) rename src/elementary-number-theory/{powers-of-two.lagda.md => 2-adic-decomposition.lagda.md} (59%) create mode 100644 src/order-theory/deflationary-maps-posets.lagda.md create mode 100644 src/order-theory/deflationary-maps-preorders.lagda.md create mode 100644 src/order-theory/inflationary-maps-posets.lagda.md create mode 100644 src/order-theory/inflationary-maps-preorders.lagda.md create mode 100644 src/order-theory/strictly-ordered-sets.lagda.md create mode 100644 src/order-theory/strictly-ordered-types.lagda.md diff --git a/src/domain-theory/directed-families-posets.lagda.md b/src/domain-theory/directed-families-posets.lagda.md index 2684d639aa..0b192c1b84 100644 --- a/src/domain-theory/directed-families-posets.lagda.md +++ b/src/domain-theory/directed-families-posets.lagda.md @@ -138,11 +138,11 @@ module _ ( exists-structure-Prop type-directed-family-hom-Poset _) ( λ z y → intro-exists z - ( preserves-order-map-hom-Poset P Q f + ( preserves-order-hom-Poset P Q f ( family-directed-family-Poset P x u) ( family-directed-family-Poset P x z) ( pr1 y) , - preserves-order-map-hom-Poset P Q f + preserves-order-hom-Poset P Q f ( family-directed-family-Poset P x v) ( family-directed-family-Poset P x z) ( pr2 y))) diff --git a/src/elementary-number-theory.lagda.md b/src/elementary-number-theory.lagda.md index 0e899d079e..c65d8b05e2 100644 --- a/src/elementary-number-theory.lagda.md +++ b/src/elementary-number-theory.lagda.md @@ -131,7 +131,7 @@ open import elementary-number-theory.positive-integer-fractions public open import elementary-number-theory.positive-integers public open import elementary-number-theory.positive-rational-numbers public open import elementary-number-theory.powers-integers public -open import elementary-number-theory.powers-of-two public +open import elementary-number-theory.2-adic-decomposition public open import elementary-number-theory.prime-numbers public open import elementary-number-theory.products-of-natural-numbers public open import elementary-number-theory.proper-divisors-natural-numbers public diff --git a/src/elementary-number-theory/powers-of-two.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md similarity index 59% rename from src/elementary-number-theory/powers-of-two.lagda.md rename to src/elementary-number-theory/2-adic-decomposition.lagda.md index 4fed8c7c76..cc6d24cae8 100644 --- a/src/elementary-number-theory/powers-of-two.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -1,13 +1,14 @@ -# Powers of two +# 2-Adic decomposition ```agda -module elementary-number-theory.powers-of-two where +module elementary-number-theory.2-adic-decomposition where ```
Imports ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.based-strong-induction-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.inequality-natural-numbers @@ -20,6 +21,7 @@ open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types open import foundation.equality-cartesian-product-types open import foundation.function-types +open import foundation.functoriality-dependent-pair-types open import foundation.split-surjective-maps open import foundation.transport-along-identifications open import foundation.universe-levels @@ -36,10 +38,160 @@ open import foundation-core.injective-maps ## Idea -Any natural number `x` can be written as `(2^u(2v-1))-1` for some pair of -natural numbers `(u , v)` +The {{#concept "2-adic decomposition" Agda=2-adic-decomposition-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) $n$ is a factorization of $n$ into a [power](elementary-number-theory.exponentiation-natural-numbers.md) of $2$ and an [odd](elementary-number-theory.parity-natural-numbers.md) natural number. + +The $2$-adic decomposition of the natural numbers can be used to construct an [equivalence](foundation-core.equivalences.md) $\mathbb{N}\times\mathbb{N} \simeq \mathbb{N}$ by mapping + +$$ + (m , n) \mapsto 2^m(2n + 1) - 1. +$$ + +## Definitions + +### The $2$-adic composition function + +```agda +2-adic-composition-ℕ : ℕ → ℕ → ℕ +2-adic-composition-ℕ k l = 2 ^ℕ k *ℕ odd-number-ℕ l +``` + +### The type of $2$-adic decompositions of a natural number + +```agda +2-adic-decomposition-ℕ : ℕ → UU lzero +2-adic-decomposition-ℕ n = + Σ ℕ (λ k → Σ ℕ (λ l → 2-adic-composition-ℕ k l = n)) + +module _ + (n : ℕ) (d : 2-adic-decomposition-ℕ n) + where + + exponent-2-adic-decomposition-ℕ : ℕ + exponent-2-adic-decomposition-ℕ = pr1 d + + index-odd-factor-2-adic-decomposition-ℕ : ℕ + index-odd-factor-2-adic-decomposition-ℕ = pr1 (pr2 d) + + eq-2-adic-decomposition-ℕ : + 2-adic-composition-ℕ + exponent-2-adic-decomposition-ℕ + index-odd-factor-2-adic-decomposition-ℕ = + n + eq-2-adic-decomposition-ℕ = pr2 (pr2 d) +``` + +## Properties + +### The values of the $2$-adic composition function are nonzero + +```agda +is-nonzero-2-adic-composition-ℕ : + (k l : ℕ) → is-nonzero-ℕ (2-adic-composition-ℕ k l) +is-nonzero-2-adic-composition-ℕ k l = + is-nonzero-mul-ℕ + ( 2 ^ℕ k) + ( odd-number-ℕ l) + ( is-nonzero-exp-ℕ 2 k is-nonzero-two-ℕ) + ( is-nonzero-odd-number-ℕ l) +``` + +### Any number that has a $2$-adic decomposition is nonzero ```agda +is-nonzero-2-adic-decomposition-ℕ : + (n : ℕ) → 2-adic-decomposition-ℕ n → is-nonzero-ℕ n +is-nonzero-2-adic-decomposition-ℕ ._ (k , l , refl) = + is-nonzero-2-adic-composition-ℕ k l +``` + +### Every odd number has a $2$-adic decomposition + +```agda +2-adic-decomposition-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → 2-adic-decomposition-ℕ n +pr1 (2-adic-decomposition-is-odd-ℕ n H) = + 0 +pr1 (pr2 (2-adic-decomposition-is-odd-ℕ n H)) = + pr1 (has-odd-expansion-is-odd-ℕ n H) +pr2 (pr2 (2-adic-decomposition-is-odd-ℕ n H)) = + left-unit-law-mul-ℕ _ ∙ pr2 (has-odd-expansion-is-odd-ℕ n H) +``` + +### Every nonzero natural number has a $2$-adic decomposition + +The proof is by [based strong induction](elementary-number-theory.based-strong-induction-natural-numbers.md). +We have already seen that every odd natural number has a $2$-adic expansion. +If $1 \leq x$ is an even number, then write $x = 2 \cdot y$. Given a $2$-adic decomposition $(k , l)$ of the number $y$, it follows that $(k + 1 , l)$ is a $2$-adic decomposition of $x$, because + +$$ + 2^{k+1}(2l + 1) = 2(2^k(2l + 1)) = 2*y = x. +$$ + +```agda +module _ + (x : ℕ) (H : 1 ≤-ℕ x) + (f : + based-□-≤-ℕ 1 (λ y → is-even-ℕ y + is-odd-ℕ y → 2-adic-decomposition-ℕ y) x) + (K@(q , p) : is-even-ℕ (succ-ℕ x)) + where + + 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ : + 2-adic-decomposition-ℕ q + 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ = + f ( q) + ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) + ( upper-bound-quotient-is-even-succ-ℕ x K) + ( is-decidable-is-even-ℕ q) + + exponent-even-case-2-adic-decomposition-is-even-or-odd-ℕ : + ℕ + exponent-even-case-2-adic-decomposition-is-even-or-odd-ℕ = + succ-ℕ + ( exponent-2-adic-decomposition-ℕ + ( q) + ( 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ)) + + even-case-2-adic-decomposition-is-even-or-odd-ℕ : + 2-adic-decomposition-ℕ (succ-ℕ x) + even-case-2-adic-decomposition-is-even-or-odd-ℕ = + ( succ-ℕ (pr1 d) , pr1 (pr2 d) , {!!}) + where + + d = + f ( q) + ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) + {!!} + {!!} + +2-adic-decomposition-is-even-or-odd-ℕ : + (n : ℕ) → 1 ≤-ℕ n → is-even-ℕ n + is-odd-ℕ n → + 2-adic-decomposition-ℕ n +2-adic-decomposition-is-even-or-odd-ℕ = + based-strong-ind-ℕ 1 + ( λ x → (is-even-ℕ x + is-odd-ℕ x) → 2-adic-decomposition-ℕ x) + ( rec-coproduct + ( λ H → ex-falso (is-odd-one-ℕ H)) + ( 2-adic-decomposition-is-odd-ℕ 1)) + ( λ x H f → + rec-coproduct + ( even-case-2-adic-decomposition-is-even-or-odd-ℕ x H f + {- + λ K@(q , p) → + let + + L : 2-adic-decomposition-ℕ q +d L = + f ( q) + ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) + {! upper-bound-quotient-div-ℕ 2 (succ-ℕ x) K!} + {!!} + in + {!f q (leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) ? ?!} + -}) + ( 2-adic-decomposition-is-odd-ℕ (succ-ℕ x))) +``` + +```text pair-expansion : ℕ → UU lzero pair-expansion n = Σ (ℕ × ℕ) @@ -124,7 +276,7 @@ has-pair-expansion n = ### If `(u , v)` and `(u' , v')` are the pairs corresponding the same number `x`, then `u = u'` and `v = v'` -```agda +```text is-pair-expansion-unique : (u u' v v' : ℕ) → ((exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = @@ -183,7 +335,7 @@ A pairing function is a bijection between `ℕ × ℕ` and `ℕ`. ## Definition -```agda +```text pairing-map : ℕ × ℕ → ℕ pairing-map (u , v) = pr1 (is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v)) @@ -191,7 +343,7 @@ pairing-map (u , v) = ### Pairing function is split surjective -```agda +```text is-split-surjective-pairing-map : is-split-surjective pairing-map is-split-surjective-pairing-map n = (u , v) , is-injective-succ-ℕ (q ∙ s) @@ -208,7 +360,7 @@ is-split-surjective-pairing-map n = ### Pairing function is injective -```agda +```text is-injective-pairing-map : is-injective pairing-map is-injective-pairing-map {u , v} {u' , v'} p = ( eq-pair' (is-pair-expansion-unique u u' v v' q)) @@ -223,7 +375,7 @@ is-injective-pairing-map {u , v} {u' , v'} p = ### Pairing function is equivalence -```agda +```text is-equiv-pairing-map : is-equiv pairing-map is-equiv-pairing-map = is-equiv-is-split-surjective-is-injective diff --git a/src/elementary-number-theory/distance-natural-numbers.lagda.md b/src/elementary-number-theory/distance-natural-numbers.lagda.md index 312f4d7e47..a0121c276e 100644 --- a/src/elementary-number-theory/distance-natural-numbers.lagda.md +++ b/src/elementary-number-theory/distance-natural-numbers.lagda.md @@ -150,7 +150,7 @@ triangle-inequality-dist-ℕ (succ-ℕ m) (succ-ℕ n) (succ-ℕ k) = ```agda is-additive-right-inverse-dist-ℕ : - (x y : ℕ) → x ≤-ℕ y → x +ℕ (dist-ℕ x y) = y + (x y : ℕ) → x ≤-ℕ y → x +ℕ dist-ℕ x y = y is-additive-right-inverse-dist-ℕ zero-ℕ zero-ℕ H = refl is-additive-right-inverse-dist-ℕ zero-ℕ (succ-ℕ y) star = left-unit-law-add-ℕ (succ-ℕ y) diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index a569184553..3dacba5c3f 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -10,13 +10,16 @@ module elementary-number-theory.exponentiation-natural-numbers where open import commutative-algebra.powers-of-elements-commutative-semirings open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.commutative-semiring-of-natural-numbers +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.products-of-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions +open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.unit-type ``` @@ -125,6 +128,13 @@ le-zero-exp-ℕ m (succ-ℕ (succ-ℕ n)) H = ( m) ( le-zero-exp-ℕ m (succ-ℕ n) H) ( H) + +leq-one-exp-ℕ : + (m n : ℕ) → 1 ≤-ℕ m → 1 ≤-ℕ m ^ℕ n +leq-one-exp-ℕ m n H = + leq-one-is-nonzero-ℕ + ( m ^ℕ n) + ( is-nonzero-exp-ℕ m n (is-nonzero-leq-one-ℕ m H)) ``` ### The exponent $m^n$ is equal to the $n$-fold product of $m$ @@ -137,3 +147,13 @@ compute-constant-product-ℕ m (succ-ℕ zero-ℕ) = left-unit-law-add-ℕ m compute-constant-product-ℕ m (succ-ℕ (succ-ℕ n)) = ap (_*ℕ m) (compute-constant-product-ℕ m (succ-ℕ n)) ``` + +### The base of the exponent divides its successor exponents + +```text +bounded-div-exp-succ-ℕ : + (m n : ℕ) → bounded-div-ℕ m (m ^ℕ succ-ℕ n) +pr1 (bounded-div-exp-succ-ℕ m n) = m ^ℕ n +pr1 (pr2 (bounded-div-exp-succ-ℕ m n)) = {!!} +pr2 (pr2 (bounded-div-exp-succ-ℕ m n)) = {!!} +``` diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index e94f631fc8..58e31b2127 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -11,13 +11,17 @@ open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.products-of-natural-numbers +open import elementary-number-theory.squares-natural-numbers open import elementary-number-theory.strong-induction-natural-numbers open import foundation.action-on-identifications-functions open import foundation.identity-types open import foundation.negated-equality +open import foundation.unit-type open import univalent-combinatorics.standard-finite-types ``` @@ -99,6 +103,14 @@ compute-succ-recursive-fermat-number-ℕ = ## Properties +### Every Fermat number is greater than $2$ + +```agda +leq-two-fermat-number-ℕ : + (n : ℕ) → 2 ≤-ℕ fermat-number-ℕ n +leq-two-fermat-number-ℕ n = leq-one-exp-ℕ 2 (2 ^ℕ n) star +``` + ### The Fermat number of a successor Any Fermat number of the form $F(n+1)=2^{2^{n+1}}+1$ can be computed as @@ -108,12 +120,29 @@ $$ $$ ```agda -dist-fermat-number-succ-two-ℕ : +dist-fermat-number-two-ℕ : + (n : ℕ) → + dist-ℕ (fermat-number-ℕ n) 2 = + Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)) +dist-fermat-number-two-ℕ zero-ℕ = refl +dist-fermat-number-two-ℕ (succ-ℕ n) = + ap (λ x → dist-ℕ x 1) (ap (2 ^ℕ_) (exp-succ-ℕ n 2) ∙ exp-mul-ℕ (2 ^ℕ n) 2) ∙ + distance-of-squares-ℕ (2 ^ℕ 2 ^ℕ n) 1 ∙ + ap (_*ℕ fermat-number-ℕ n) (dist-fermat-number-two-ℕ n) + +compute-fermat-number-ℕ : (n : ℕ) → - dist-ℕ 2 (fermat-number-ℕ (succ-ℕ n)) = - Π-ℕ (succ-ℕ n) (λ k → fermat-number-ℕ (nat-Fin (succ-ℕ n) k)) -dist-fermat-number-succ-two-ℕ n = - {!!} + fermat-number-ℕ n = Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)) +ℕ 2 +compute-fermat-number-ℕ n = + ( inv + ( rewrite-right-dist-add-ℕ 2 + ( Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k))) + ( fermat-number-ℕ n) + ( leq-two-fermat-number-ℕ n) + ( inv + ( ( symmetric-dist-ℕ 2 (fermat-number-ℕ n)) ∙ + ( dist-fermat-number-two-ℕ n))))) ∙ + ( commutative-add-ℕ 2 (Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)))) ``` ### The two definitions of the Fermat numbers agree @@ -135,7 +164,7 @@ compute-recursive-fermat-number-ℕ = ( preserves-htpy-Π-ℕ ( succ-ℕ n) ( λ i → H (nat-Fin (succ-ℕ n) i) (upper-bound-nat-Fin n i)))) ∙ - {!!}) + ( inv (compute-fermat-number-ℕ (succ-ℕ n)))) ``` ### Any two distinct Fermat numbers are relatively prime diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 3ea3e75065..991964b988 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -8,11 +8,13 @@ module elementary-number-theory.parity-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.congruence-natural-numbers open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.euclidean-division-natural-numbers +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -29,6 +31,7 @@ open import foundation.function-types open import foundation.identity-types open import foundation.negation open import foundation.transport-along-identifications +open import foundation.unit-type open import foundation.universe-levels ``` @@ -52,29 +55,57 @@ and odd natural numbers are those that aren't. ### Even natural numbers ```agda -is-even-ℕ : ℕ → UU lzero -is-even-ℕ n = div-ℕ 2 n +is-even-ℕ : + ℕ → UU lzero +is-even-ℕ n = + div-ℕ 2 n + +quotient-is-even-ℕ : + (n : ℕ) → is-even-ℕ n → ℕ +quotient-is-even-ℕ = + quotient-div-ℕ 2 + +upper-bound-quotient-is-even-ℕ : + (n : ℕ) (H : is-even-ℕ n) → quotient-is-even-ℕ n H ≤-ℕ n +upper-bound-quotient-is-even-ℕ = + upper-bound-quotient-div-ℕ 2 + +eq-quotient-is-even-ℕ : + (n : ℕ) (H : is-even-ℕ n) → quotient-is-even-ℕ n H *ℕ 2 = n +eq-quotient-is-even-ℕ = + eq-quotient-div-ℕ 2 + +eq-quotient-is-even-ℕ' : + (n : ℕ) (H : is-even-ℕ n) → 2 *ℕ quotient-is-even-ℕ n H = n +eq-quotient-is-even-ℕ' = + eq-quotient-div-ℕ' 2 ``` ### The sequence of even numbers ```agda -even-number-ℕ : ℕ → ℕ -even-number-ℕ n = 2 *ℕ n +even-number-ℕ : + ℕ → ℕ +even-number-ℕ n = + 2 *ℕ n ``` ### Odd natural numbers ```agda -is-odd-ℕ : ℕ → UU lzero -is-odd-ℕ n = ¬ (is-even-ℕ n) +is-odd-ℕ : + ℕ → UU lzero +is-odd-ℕ n = + ¬ (is-even-ℕ n) ``` ### The sequence of odd numbers ```agda -odd-number-ℕ : ℕ → ℕ -odd-number-ℕ n = succ-ℕ (2 *ℕ n) +odd-number-ℕ : + ℕ → ℕ +odd-number-ℕ n = + succ-ℕ (2 *ℕ n) ``` ### The predicate of having an odd expansion @@ -298,6 +329,34 @@ is-odd-is-1-mod-2-ℕ (succ-ℕ n) H K = ( div-right-summand-ℕ 2 n 1 (tr (div-ℕ 2) (right-unit-law-dist-ℕ n) H) K) ``` +### If a successor number `n + 1` is even, then its quotient after division by `2` is at most `n` + +**Proof.** Suppose that `q * 2 = n + 1` for some natural number `q`. Then `q` is a successor, say `q = q' + 1`. It follows that + +```text + q + 1 ≤ q + q' + 1 = q + q = q * 2 = n + 1. +``` + +This implies that `q ≤ n`. + +```agda +upper-bound-quotient-bounded-div-two-succ-ℕ : + (n : ℕ) (H : bounded-div-ℕ 2 (succ-ℕ n)) → + quotient-bounded-div-ℕ 2 (succ-ℕ n) H ≤-ℕ n +upper-bound-quotient-bounded-div-two-succ-ℕ n (succ-ℕ q , H , p) = + concatenate-leq-eq-ℕ + ( succ-ℕ (succ-ℕ q)) + ( preserves-order-right-add-ℕ (succ-ℕ q) 1 (succ-ℕ q) star) + ( inv (right-two-law-mul-ℕ (succ-ℕ q)) ∙ p) + +upper-bound-quotient-is-even-succ-ℕ : + (n : ℕ) (H : is-even-ℕ (succ-ℕ n)) → + quotient-is-even-ℕ (succ-ℕ n) H ≤-ℕ n +upper-bound-quotient-is-even-succ-ℕ n H = + upper-bound-quotient-bounded-div-two-succ-ℕ n + ( bounded-div-div-ℕ 2 (succ-ℕ n) H) +``` + ### If any two out of `x`, `y`, and `x + y` are even, then so is the third ```agda @@ -404,6 +463,18 @@ is-odd-add-succ-self-ℕ n = is-odd-succ-is-even-ℕ (n +ℕ n) (is-even-add-self-ℕ n) ``` +### Odd numbers are nonzero + +```agda +is-nonzero-odd-number-ℕ : + (n : ℕ) → is-nonzero-ℕ (odd-number-ℕ n) +is-nonzero-odd-number-ℕ n = is-nonzero-succ-ℕ (2 *ℕ n) + +is-nonzero-is-odd-ℕ : + (n : ℕ) → is-odd-ℕ n → is-nonzero-ℕ n +is-nonzero-is-odd-ℕ .zero-ℕ H refl = H is-even-zero-ℕ +``` + ## See also Further laws of parity are proven in other files, e.g.: diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index d3448fa6d2..3afee61c53 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -60,7 +60,7 @@ of `n` with itself. ```agda square-ℕ : ℕ → ℕ -square-ℕ n = mul-ℕ n n +square-ℕ n = n *ℕ n ``` ### The predicate of being a square of a natural number diff --git a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md index 9d3fa6833f..244a1fe92d 100644 --- a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md +++ b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md @@ -13,7 +13,7 @@ open import elementary-number-theory.integers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers -open import elementary-number-theory.powers-of-two +open import elementary-number-theory.2-adic-decomposition open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types diff --git a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md index 5555b7b04b..fac7df36c0 100644 --- a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md @@ -34,6 +34,11 @@ The {{#concept "well-ordering principle of the natural numbers" WDID=Q2488476 WD of [decidable types](foundation.decidable-types.md) over ℕ equipped with a natural number `n` and an element `p : P n`, we can find a least natural number `n₀` with an element `p₀ : P n₀`. +The well-ordering principle has some useful consequences: +- For any family `P` of decidable types equipped with a natural number `n ≥ b` and an element `p : P n` we can find a least natural number `n₀ ≥ b` with an element `p₀ : P n₀`. +- For any family `P` of decidable types equipped with an element `n ≤ b` and an element `p : P n` we can find a largest natural number `n₁ ≤ b` with an element `p₁ : P n₁`. +- For any strictly inflationary map `f : ℕ → ℕ` such that `f 0 ≤ b` we can find a largest natural number `n₁` such that `f n₁ ≤ b`. + ## Theorem ```agda diff --git a/src/foundation/binary-relations.lagda.md b/src/foundation/binary-relations.lagda.md index e4fe8d4600..b9ad577958 100644 --- a/src/foundation/binary-relations.lagda.md +++ b/src/foundation/binary-relations.lagda.md @@ -21,6 +21,7 @@ open import foundation-core.equivalences open import foundation-core.identity-types open import foundation-core.negation open import foundation-core.propositions +open import foundation-core.sets open import foundation-core.torsorial-type-families ``` @@ -191,7 +192,14 @@ module _ where is-irreflexive : UU (l1 ⊔ l2) - is-irreflexive = (x : A) → ¬ (R x x) + is-irreflexive = (x : A) → ¬ R x x + +module _ + {l1 l2 : Level} {A : UU l1} (R : Relation-Prop l2 A) + where + + is-irreflexive-Relation-Prop : UU (l1 ⊔ l2) + is-irreflexive-Relation-Prop = is-irreflexive (type-Relation-Prop R) ``` ### The predicate of being an asymmetric relation diff --git a/src/foundation/iterating-functions.lagda.md b/src/foundation/iterating-functions.lagda.md index 339d4eed0c..dcc947bd16 100644 --- a/src/foundation/iterating-functions.lagda.md +++ b/src/foundation/iterating-functions.lagda.md @@ -148,10 +148,11 @@ module _ (k l : ℕ) (f : X → X) (x : X) → iterate (exp-ℕ l k) f x = iterate k (iterate l) f x iterate-exp-ℕ zero-ℕ l f x = refl - iterate-exp-ℕ (succ-ℕ k) l f x = - ( iterate-mul-ℕ (exp-ℕ l k) l f x) ∙ - ( ( iterate-exp-ℕ k l (iterate l f) x) ∙ - ( inv (htpy-eq (reassociate-iterate-succ-ℕ k (iterate l) f) x))) + iterate-exp-ℕ (succ-ℕ zero-ℕ) l f x = refl + iterate-exp-ℕ (succ-ℕ (succ-ℕ k)) l f x = + ( iterate-mul-ℕ (exp-ℕ l (succ-ℕ k)) l f x) ∙ + ( ( iterate-exp-ℕ (succ-ℕ k) l (iterate l f) x) ∙ + ( inv (htpy-eq (reassociate-iterate-succ-ℕ (succ-ℕ k) (iterate l) f) x))) module _ {l : Level} (X : Set l) diff --git a/src/order-theory/chains-posets.lagda.md b/src/order-theory/chains-posets.lagda.md index 0136242af1..404b506ee9 100644 --- a/src/order-theory/chains-posets.lagda.md +++ b/src/order-theory/chains-posets.lagda.md @@ -130,7 +130,7 @@ module _ inl-disjunction ( concatenate-eq-leq-eq-Poset' X ( pr2 p) - ( preserves-order-map-hom-Poset (poset-Total-Order T) X f + ( preserves-order-hom-Poset (poset-Total-Order T) X f ( pr1 p) ( pr1 q) ( H)) @@ -139,7 +139,7 @@ module _ inr-disjunction ( concatenate-eq-leq-eq-Poset' X ( pr2 q) - ( preserves-order-map-hom-Poset (poset-Total-Order T) X f + ( preserves-order-hom-Poset (poset-Total-Order T) X f ( pr1 q) ( pr1 p) ( H)) diff --git a/src/order-theory/deflationary-maps-posets.lagda.md b/src/order-theory/deflationary-maps-posets.lagda.md new file mode 100644 index 0000000000..e765e77437 --- /dev/null +++ b/src/order-theory/deflationary-maps-posets.lagda.md @@ -0,0 +1,145 @@ +# Deflationary maps on a poset + +```agda +module order-theory.deflationary-maps-posets where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.deflationary-maps-preorders +open import order-theory.order-preserving-maps-posets +open import order-theory.posets +``` + +
+ +## Idea + +A map $f : P → P$ on a [poset](order-theory.posets.md) $P$ is said to be an +{{#concept "deflationary map" Disambiguation="poset" Agda=deflationary-map-Poset}} if the inequality + +$$ + f(x) \leq x +$$ + +holds for any element $x : P$. In other words, a map on a poset is deflationary precisely when the map on its underlying [preorder](order-theory.preorders.md) is [deflationary](order-theory.deflationary-maps-preorders.md). If $f$ is also [order preserving](order-theory.order-preserving-maps-posets.md) we say that $f$ is an {{#concept "deflationary morphism" Disambiguation="poset" Agda=deflationary-hom-Poset}}. + +## Definitions + +### The predicate of being an deflationary map + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : type-Poset P → type-Poset P) + where + + is-deflationary-prop-map-Poset : + Prop (l1 ⊔ l2) + is-deflationary-prop-map-Poset = + is-deflationary-prop-map-Preorder (preorder-Poset P) f + + is-deflationary-map-Poset : + UU (l1 ⊔ l2) + is-deflationary-map-Poset = + is-deflationary-map-Preorder (preorder-Poset P) f + + is-prop-is-deflationary-map-Poset : + is-prop is-deflationary-map-Poset + is-prop-is-deflationary-map-Poset = + is-prop-is-deflationary-map-Preorder (preorder-Poset P) f +``` + +### The type of deflationary maps on a poset + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) + where + + deflationary-map-Poset : + UU (l1 ⊔ l2) + deflationary-map-Poset = + deflationary-map-Preorder (preorder-Poset P) + +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : deflationary-map-Poset P) + where + + map-deflationary-map-Poset : + type-Poset P → type-Poset P + map-deflationary-map-Poset = + map-deflationary-map-Preorder (preorder-Poset P) f + + is-deflationary-deflationary-map-Poset : + is-deflationary-map-Poset P map-deflationary-map-Poset + is-deflationary-deflationary-map-Poset = + is-deflationary-deflationary-map-Preorder (preorder-Poset P) f +``` + +### The predicate on order preserving maps of being deflationary + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : hom-Poset P P) + where + + is-deflationary-prop-hom-Poset : Prop (l1 ⊔ l2) + is-deflationary-prop-hom-Poset = + is-deflationary-prop-hom-Preorder (preorder-Poset P) f + + is-deflationary-hom-Poset : UU (l1 ⊔ l2) + is-deflationary-hom-Poset = + is-deflationary-hom-Preorder (preorder-Poset P) f + + is-prop-is-deflationary-hom-Poset : + is-prop is-deflationary-hom-Poset + is-prop-is-deflationary-hom-Poset = + is-prop-is-deflationary-hom-Preorder (preorder-Poset P) f +``` + +### The type of deflationary morphisms on a poset + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) + where + + deflationary-hom-Poset : UU (l1 ⊔ l2) + deflationary-hom-Poset = + deflationary-hom-Preorder (preorder-Poset P) + +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : deflationary-hom-Poset P) + where + + hom-deflationary-hom-Poset : + hom-Poset P P + hom-deflationary-hom-Poset = + hom-deflationary-hom-Preorder (preorder-Poset P) f + + map-deflationary-hom-Poset : + type-Poset P → type-Poset P + map-deflationary-hom-Poset = + map-deflationary-hom-Preorder (preorder-Poset P) f + + preserves-order-deflationary-hom-Poset : + preserves-order-Poset P P map-deflationary-hom-Poset + preserves-order-deflationary-hom-Poset = + preserves-order-deflationary-hom-Preorder (preorder-Poset P) f + + is-deflationary-deflationary-hom-Poset : + is-deflationary-map-Poset P map-deflationary-hom-Poset + is-deflationary-deflationary-hom-Poset = + is-deflationary-deflationary-hom-Preorder (preorder-Poset P) f + + deflationary-map-deflationary-hom-Poset : + deflationary-map-Poset P + deflationary-map-deflationary-hom-Poset = + deflationary-map-deflationary-hom-Preorder (preorder-Poset P) f +``` diff --git a/src/order-theory/deflationary-maps-preorders.lagda.md b/src/order-theory/deflationary-maps-preorders.lagda.md new file mode 100644 index 0000000000..0368370cb8 --- /dev/null +++ b/src/order-theory/deflationary-maps-preorders.lagda.md @@ -0,0 +1,146 @@ +# Deflationary maps on a preorder + +```agda +module order-theory.deflationary-maps-preorders where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.order-preserving-maps-preorders +open import order-theory.preorders +``` + +
+ +## Idea + +A map $f : P → P$ on a [preorder](order-theory.preorders.md) $P$ is said to be an +{{#concept "deflationary map" Disambiguation="preorder" Agda=deflationary-map-Preorder}} if the inequality + +$$ + f(x) \leq x +$$ + +holds for any element $x : P$. If $f$ is also [order preserving](order-theory.order-preserving-maps-preorders.md) we say that $f$ is an {{#concept "deflationary morphism" Disambiguation="preorder" Agda=deflationary-hom-Preorder}}. + +## Definitions + +### The predicate of being an deflationary map + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : type-Preorder P → type-Preorder P) + where + + is-deflationary-prop-map-Preorder : + Prop (l1 ⊔ l2) + is-deflationary-prop-map-Preorder = + Π-Prop (type-Preorder P) (λ x → leq-prop-Preorder P (f x) x) + + is-deflationary-map-Preorder : + UU (l1 ⊔ l2) + is-deflationary-map-Preorder = + type-Prop is-deflationary-prop-map-Preorder + + is-prop-is-deflationary-map-Preorder : + is-prop is-deflationary-map-Preorder + is-prop-is-deflationary-map-Preorder = + is-prop-type-Prop is-deflationary-prop-map-Preorder +``` + +### The type of deflationary maps on a preorder + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) + where + + deflationary-map-Preorder : + UU (l1 ⊔ l2) + deflationary-map-Preorder = + type-subtype (is-deflationary-prop-map-Preorder P) + +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : deflationary-map-Preorder P) + where + + map-deflationary-map-Preorder : + type-Preorder P → type-Preorder P + map-deflationary-map-Preorder = + pr1 f + + is-deflationary-deflationary-map-Preorder : + is-deflationary-map-Preorder P map-deflationary-map-Preorder + is-deflationary-deflationary-map-Preorder = + pr2 f +``` + +### The predicate on order preserving maps of being deflationary + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : hom-Preorder P P) + where + + is-deflationary-prop-hom-Preorder : Prop (l1 ⊔ l2) + is-deflationary-prop-hom-Preorder = + is-deflationary-prop-map-Preorder P (map-hom-Preorder P P f) + + is-deflationary-hom-Preorder : UU (l1 ⊔ l2) + is-deflationary-hom-Preorder = + is-deflationary-map-Preorder P (map-hom-Preorder P P f) + + is-prop-is-deflationary-hom-Preorder : + is-prop is-deflationary-hom-Preorder + is-prop-is-deflationary-hom-Preorder = + is-prop-is-deflationary-map-Preorder P (map-hom-Preorder P P f) +``` + +### The type of deflationary morphisms on a preorder + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) + where + + deflationary-hom-Preorder : UU (l1 ⊔ l2) + deflationary-hom-Preorder = + type-subtype (is-deflationary-prop-hom-Preorder P) + +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : deflationary-hom-Preorder P) + where + + hom-deflationary-hom-Preorder : + hom-Preorder P P + hom-deflationary-hom-Preorder = + pr1 f + + map-deflationary-hom-Preorder : + type-Preorder P → type-Preorder P + map-deflationary-hom-Preorder = + map-hom-Preorder P P hom-deflationary-hom-Preorder + + preserves-order-deflationary-hom-Preorder : + preserves-order-Preorder P P map-deflationary-hom-Preorder + preserves-order-deflationary-hom-Preorder = + preserves-order-hom-Preorder P P hom-deflationary-hom-Preorder + + is-deflationary-deflationary-hom-Preorder : + is-deflationary-map-Preorder P map-deflationary-hom-Preorder + is-deflationary-deflationary-hom-Preorder = + pr2 f + + deflationary-map-deflationary-hom-Preorder : + deflationary-map-Preorder P + pr1 deflationary-map-deflationary-hom-Preorder = + map-deflationary-hom-Preorder + pr2 deflationary-map-deflationary-hom-Preorder = + is-deflationary-deflationary-hom-Preorder +``` diff --git a/src/order-theory/galois-connections.lagda.md b/src/order-theory/galois-connections.lagda.md index 5ae6e96259..e0ac6170c7 100644 --- a/src/order-theory/galois-connections.lagda.md +++ b/src/order-theory/galois-connections.lagda.md @@ -84,7 +84,7 @@ module _ preserves-order-upper-adjoint-Galois-Connection : preserves-order-Poset Q P map-upper-adjoint-Galois-Connection preserves-order-upper-adjoint-Galois-Connection = - preserves-order-map-hom-Poset Q P upper-adjoint-Galois-Connection + preserves-order-hom-Poset Q P upper-adjoint-Galois-Connection is-upper-adjoint-upper-adjoint-Galois-Connection : is-upper-adjoint-Galois-Connection upper-adjoint-Galois-Connection @@ -102,7 +102,7 @@ module _ preserves-order-lower-adjoint-Galois-Connection : preserves-order-Poset P Q map-lower-adjoint-Galois-Connection preserves-order-lower-adjoint-Galois-Connection = - preserves-order-map-hom-Poset P Q lower-adjoint-Galois-Connection + preserves-order-hom-Poset P Q lower-adjoint-Galois-Connection adjoint-relation-Galois-Connection : (x : type-Poset P) (y : type-Poset Q) → diff --git a/src/order-theory/homomorphisms-large-suplattices.lagda.md b/src/order-theory/homomorphisms-large-suplattices.lagda.md index ac054229b8..181c422651 100644 --- a/src/order-theory/homomorphisms-large-suplattices.lagda.md +++ b/src/order-theory/homomorphisms-large-suplattices.lagda.md @@ -81,14 +81,14 @@ module _ ( large-poset-Large-Suplattice L) ( hom-large-poset-hom-Large-Suplattice f) - preserves-order-map-hom-Large-Suplattice : + preserves-order-hom-Large-Suplattice : {l1 l2 : Level} (x : type-Large-Suplattice K l1) (y : type-Large-Suplattice K l2) → leq-Large-Suplattice K x y → leq-Large-Suplattice L ( map-hom-Large-Suplattice x) ( map-hom-Large-Suplattice y) - preserves-order-map-hom-Large-Suplattice = + preserves-order-hom-Large-Suplattice = preserves-order-hom-Large-Poset ( large-poset-Large-Suplattice K) ( large-poset-Large-Suplattice L) diff --git a/src/order-theory/inflationary-maps-posets.lagda.md b/src/order-theory/inflationary-maps-posets.lagda.md new file mode 100644 index 0000000000..d6e25e7dbf --- /dev/null +++ b/src/order-theory/inflationary-maps-posets.lagda.md @@ -0,0 +1,145 @@ +# Inflationary maps on a poset + +```agda +module order-theory.inflationary-maps-posets where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.inflationary-maps-preorders +open import order-theory.order-preserving-maps-posets +open import order-theory.posets +``` + +
+ +## Idea + +A map $f : P → P$ on a [poset](order-theory.posets.md) $P$ is said to be an +{{#concept "inflationary map" Disambiguation="poset" Agda=inflationary-map-Poset}} if the inequality + +$$ + x \leq f(x) +$$ + +holds for any element $x : P$. In other words, a map on a poset is inflationary precisely when the map on its underlying [preorder](order-theory.preorders.md) is [inflationary](order-theory.inflationary-maps-preorders.md). If $f$ is also [order preserving](order-theory.order-preserving-maps-posets.md) we say that $f$ is an {{#concept "inflationary morphism" Disambiguation="poset" Agda=inflationary-hom-Poset}}. + +## Definitions + +### The predicate of being an inflationary map + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : type-Poset P → type-Poset P) + where + + is-inflationary-prop-map-Poset : + Prop (l1 ⊔ l2) + is-inflationary-prop-map-Poset = + is-inflationary-prop-map-Preorder (preorder-Poset P) f + + is-inflationary-map-Poset : + UU (l1 ⊔ l2) + is-inflationary-map-Poset = + is-inflationary-map-Preorder (preorder-Poset P) f + + is-prop-is-inflationary-map-Poset : + is-prop is-inflationary-map-Poset + is-prop-is-inflationary-map-Poset = + is-prop-is-inflationary-map-Preorder (preorder-Poset P) f +``` + +### The type of inflationary maps on a poset + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) + where + + inflationary-map-Poset : + UU (l1 ⊔ l2) + inflationary-map-Poset = + inflationary-map-Preorder (preorder-Poset P) + +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : inflationary-map-Poset P) + where + + map-inflationary-map-Poset : + type-Poset P → type-Poset P + map-inflationary-map-Poset = + map-inflationary-map-Preorder (preorder-Poset P) f + + is-inflationary-inflationary-map-Poset : + is-inflationary-map-Poset P map-inflationary-map-Poset + is-inflationary-inflationary-map-Poset = + is-inflationary-inflationary-map-Preorder (preorder-Poset P) f +``` + +### The predicate on order preserving maps of being inflationary + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : hom-Poset P P) + where + + is-inflationary-prop-hom-Poset : Prop (l1 ⊔ l2) + is-inflationary-prop-hom-Poset = + is-inflationary-prop-hom-Preorder (preorder-Poset P) f + + is-inflationary-hom-Poset : UU (l1 ⊔ l2) + is-inflationary-hom-Poset = + is-inflationary-hom-Preorder (preorder-Poset P) f + + is-prop-is-inflationary-hom-Poset : + is-prop is-inflationary-hom-Poset + is-prop-is-inflationary-hom-Poset = + is-prop-is-inflationary-hom-Preorder (preorder-Poset P) f +``` + +### The type of inflationary morphisms on a poset + +```agda +module _ + {l1 l2 : Level} (P : Poset l1 l2) + where + + inflationary-hom-Poset : UU (l1 ⊔ l2) + inflationary-hom-Poset = + inflationary-hom-Preorder (preorder-Poset P) + +module _ + {l1 l2 : Level} (P : Poset l1 l2) (f : inflationary-hom-Poset P) + where + + hom-inflationary-hom-Poset : + hom-Poset P P + hom-inflationary-hom-Poset = + hom-inflationary-hom-Preorder (preorder-Poset P) f + + map-inflationary-hom-Poset : + type-Poset P → type-Poset P + map-inflationary-hom-Poset = + map-inflationary-hom-Preorder (preorder-Poset P) f + + preserves-order-inflationary-hom-Poset : + preserves-order-Poset P P map-inflationary-hom-Poset + preserves-order-inflationary-hom-Poset = + preserves-order-inflationary-hom-Preorder (preorder-Poset P) f + + is-inflationary-inflationary-hom-Poset : + is-inflationary-map-Poset P map-inflationary-hom-Poset + is-inflationary-inflationary-hom-Poset = + is-inflationary-inflationary-hom-Preorder (preorder-Poset P) f + + inflationary-map-inflationary-hom-Poset : + inflationary-map-Poset P + inflationary-map-inflationary-hom-Poset = + inflationary-map-inflationary-hom-Preorder (preorder-Poset P) f +``` diff --git a/src/order-theory/inflationary-maps-preorders.lagda.md b/src/order-theory/inflationary-maps-preorders.lagda.md new file mode 100644 index 0000000000..a057e74a03 --- /dev/null +++ b/src/order-theory/inflationary-maps-preorders.lagda.md @@ -0,0 +1,146 @@ +# Inflationary maps on a preorder + +```agda +module order-theory.inflationary-maps-preorders where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.order-preserving-maps-preorders +open import order-theory.preorders +``` + +
+ +## Idea + +A map $f : P → P$ on a [preorder](order-theory.preorders.md) $P$ is said to be an +{{#concept "inflationary map" Disambiguation="preorder" Agda=inflationary-map-Preorder}} if the inequality + +$$ + x \leq f(x) +$$ + +holds for any element $x : P$. If $f$ is also [order preserving](order-theory.order-preserving-maps-preorders.md) we say that $f$ is an {{#concept "inflationary morphism" Disambiguation="preorder" Agda=inflationary-hom-Preorder}}. + +## Definitions + +### The predicate of being an inflationary map + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : type-Preorder P → type-Preorder P) + where + + is-inflationary-prop-map-Preorder : + Prop (l1 ⊔ l2) + is-inflationary-prop-map-Preorder = + Π-Prop (type-Preorder P) (λ x → leq-prop-Preorder P x (f x)) + + is-inflationary-map-Preorder : + UU (l1 ⊔ l2) + is-inflationary-map-Preorder = + type-Prop is-inflationary-prop-map-Preorder + + is-prop-is-inflationary-map-Preorder : + is-prop is-inflationary-map-Preorder + is-prop-is-inflationary-map-Preorder = + is-prop-type-Prop is-inflationary-prop-map-Preorder +``` + +### The type of inflationary maps on a preorder + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) + where + + inflationary-map-Preorder : + UU (l1 ⊔ l2) + inflationary-map-Preorder = + type-subtype (is-inflationary-prop-map-Preorder P) + +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : inflationary-map-Preorder P) + where + + map-inflationary-map-Preorder : + type-Preorder P → type-Preorder P + map-inflationary-map-Preorder = + pr1 f + + is-inflationary-inflationary-map-Preorder : + is-inflationary-map-Preorder P map-inflationary-map-Preorder + is-inflationary-inflationary-map-Preorder = + pr2 f +``` + +### The predicate on order preserving maps of being inflationary + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : hom-Preorder P P) + where + + is-inflationary-prop-hom-Preorder : Prop (l1 ⊔ l2) + is-inflationary-prop-hom-Preorder = + is-inflationary-prop-map-Preorder P (map-hom-Preorder P P f) + + is-inflationary-hom-Preorder : UU (l1 ⊔ l2) + is-inflationary-hom-Preorder = + is-inflationary-map-Preorder P (map-hom-Preorder P P f) + + is-prop-is-inflationary-hom-Preorder : + is-prop is-inflationary-hom-Preorder + is-prop-is-inflationary-hom-Preorder = + is-prop-is-inflationary-map-Preorder P (map-hom-Preorder P P f) +``` + +### The type of inflationary morphisms on a preorder + +```agda +module _ + {l1 l2 : Level} (P : Preorder l1 l2) + where + + inflationary-hom-Preorder : UU (l1 ⊔ l2) + inflationary-hom-Preorder = + type-subtype (is-inflationary-prop-hom-Preorder P) + +module _ + {l1 l2 : Level} (P : Preorder l1 l2) (f : inflationary-hom-Preorder P) + where + + hom-inflationary-hom-Preorder : + hom-Preorder P P + hom-inflationary-hom-Preorder = + pr1 f + + map-inflationary-hom-Preorder : + type-Preorder P → type-Preorder P + map-inflationary-hom-Preorder = + map-hom-Preorder P P hom-inflationary-hom-Preorder + + preserves-order-inflationary-hom-Preorder : + preserves-order-Preorder P P map-inflationary-hom-Preorder + preserves-order-inflationary-hom-Preorder = + preserves-order-hom-Preorder P P hom-inflationary-hom-Preorder + + is-inflationary-inflationary-hom-Preorder : + is-inflationary-map-Preorder P map-inflationary-hom-Preorder + is-inflationary-inflationary-hom-Preorder = + pr2 f + + inflationary-map-inflationary-hom-Preorder : + inflationary-map-Preorder P + pr1 inflationary-map-inflationary-hom-Preorder = + map-inflationary-hom-Preorder + pr2 inflationary-map-inflationary-hom-Preorder = + is-inflationary-inflationary-hom-Preorder +``` diff --git a/src/order-theory/opposite-preorders.lagda.md b/src/order-theory/opposite-preorders.lagda.md index 39c2c42ec0..2e6db98be2 100644 --- a/src/order-theory/opposite-preorders.lagda.md +++ b/src/order-theory/opposite-preorders.lagda.md @@ -80,7 +80,7 @@ module _ hom-Preorder (opposite-Preorder P) (opposite-Preorder Q) opposite-hom-Preorder f = ( map-hom-Preorder P Q f) , - ( λ x y p → preserves-order-map-hom-Preorder P Q f y x p) + ( λ x y p → preserves-order-hom-Preorder P Q f y x p) ``` ## Properties diff --git a/src/order-theory/order-preserving-maps-posets.lagda.md b/src/order-theory/order-preserving-maps-posets.lagda.md index 077313f75b..163aee3f3f 100644 --- a/src/order-theory/order-preserving-maps-posets.lagda.md +++ b/src/order-theory/order-preserving-maps-posets.lagda.md @@ -64,10 +64,10 @@ module _ map-hom-Poset : hom-Poset → type-Poset P → type-Poset Q map-hom-Poset = map-hom-Preorder (preorder-Poset P) (preorder-Poset Q) - preserves-order-map-hom-Poset : + preserves-order-hom-Poset : (f : hom-Poset) → preserves-order-Poset (map-hom-Poset f) - preserves-order-map-hom-Poset = - preserves-order-map-hom-Preorder (preorder-Poset P) (preorder-Poset Q) + preserves-order-hom-Poset = + preserves-order-hom-Preorder (preorder-Poset P) (preorder-Poset Q) ``` ### Homotopies of order preserving maps diff --git a/src/order-theory/order-preserving-maps-preorders.lagda.md b/src/order-theory/order-preserving-maps-preorders.lagda.md index d3bf16d8ee..a45df09656 100644 --- a/src/order-theory/order-preserving-maps-preorders.lagda.md +++ b/src/order-theory/order-preserving-maps-preorders.lagda.md @@ -28,21 +28,24 @@ open import order-theory.preorders ## Idea -A map `f : P → Q` between the underlying types of two preorders is said to be -**order preserving** if `x ≤ y` in `P` implies `f x ≤ f y` in `Q`. +A map `f : P → Q` between the underlying types of two +[preorders](order-theory.preorders.md) is said to be an +{{#concept "order preserving map" Disambiguation="preorder" Agda=hom-Preorder}} +if for any two elements `x ≤ y` in `P` we have `f x ≤ f y` in `Q`. ## Definition -### Order preserving maps +### The predicate of being an order preserving map ```agda module _ {l1 l2 l3 l4 : Level} (P : Preorder l1 l2) (Q : Preorder l3 l4) + (f : type-Preorder P → type-Preorder Q) where preserves-order-prop-Preorder : - (type-Preorder P → type-Preorder Q) → Prop (l1 ⊔ l2 ⊔ l4) - preserves-order-prop-Preorder f = + Prop (l1 ⊔ l2 ⊔ l4) + preserves-order-prop-Preorder = Π-Prop ( type-Preorder P) ( λ x → @@ -54,26 +57,37 @@ module _ ( leq-prop-Preorder Q (f x) (f y)))) preserves-order-Preorder : - (type-Preorder P → type-Preorder Q) → UU (l1 ⊔ l2 ⊔ l4) - preserves-order-Preorder f = - type-Prop (preserves-order-prop-Preorder f) + UU (l1 ⊔ l2 ⊔ l4) + preserves-order-Preorder = + type-Prop preserves-order-prop-Preorder is-prop-preserves-order-Preorder : - (f : type-Preorder P → type-Preorder Q) → - is-prop (preserves-order-Preorder f) - is-prop-preserves-order-Preorder f = - is-prop-type-Prop (preserves-order-prop-Preorder f) + is-prop preserves-order-Preorder + is-prop-preserves-order-Preorder = + is-prop-type-Prop preserves-order-prop-Preorder +``` + +### The type of order preserving maps + +```agda +module _ + {l1 l2 l3 l4 : Level} (P : Preorder l1 l2) (Q : Preorder l3 l4) + where - hom-Preorder : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + hom-Preorder : + UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) hom-Preorder = - Σ (type-Preorder P → type-Preorder Q) preserves-order-Preorder + Σ (type-Preorder P → type-Preorder Q) (preserves-order-Preorder P Q) - map-hom-Preorder : hom-Preorder → type-Preorder P → type-Preorder Q - map-hom-Preorder = pr1 + map-hom-Preorder : + hom-Preorder → type-Preorder P → type-Preorder Q + map-hom-Preorder = + pr1 - preserves-order-map-hom-Preorder : - (f : hom-Preorder) → preserves-order-Preorder (map-hom-Preorder f) - preserves-order-map-hom-Preorder = pr2 + preserves-order-hom-Preorder : + (f : hom-Preorder) → preserves-order-Preorder P Q (map-hom-Preorder f) + preserves-order-hom-Preorder = + pr2 ``` ### Homotopies of order preserving maps @@ -101,7 +115,7 @@ module _ ( is-prop-preserves-order-Preorder P Q) ( map-hom-Preorder P Q f) ( refl-htpy) - ( preserves-order-map-hom-Preorder P Q f) + ( preserves-order-hom-Preorder P Q f) is-equiv-htpy-eq-hom-Preorder : (f g : hom-Preorder P Q) → is-equiv (htpy-eq-hom-Preorder f g) @@ -150,10 +164,10 @@ module _ preserves-order-Preorder P R ( map-hom-Preorder Q R g ∘ map-hom-Preorder P Q f) preserves-order-comp-Preorder g f x y H = - preserves-order-map-hom-Preorder Q R g + preserves-order-hom-Preorder Q R g ( map-hom-Preorder P Q f x) ( map-hom-Preorder P Q f y) - ( preserves-order-map-hom-Preorder P Q f x y H) + ( preserves-order-hom-Preorder P Q f x y H) comp-hom-Preorder : (g : hom-Preorder Q R) (f : hom-Preorder P Q) → diff --git a/src/order-theory/strictly-ordered-sets.lagda.md b/src/order-theory/strictly-ordered-sets.lagda.md new file mode 100644 index 0000000000..f53b3caf1d --- /dev/null +++ b/src/order-theory/strictly-ordered-sets.lagda.md @@ -0,0 +1,115 @@ +# Strictly ordered sets + +```agda +module order-theory.strictly-ordered-sets where +``` + +
Imports + +```agda +open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.negation +open import foundation.propositions +open import foundation.sets +open import foundation.universe-levels + +open import order-theory.strictly-ordered-types +``` + +
+ +## Idea + +A {{#concept "strictly ordered set" Agda=Strictly-Ordered-Set}} is a [strictly ordered type](order-theory.strictly-ordered-types.md) whose underlying type is a [set](foundation-core.sets.md). More specifically, a strictly ordered set consists of a set $A$, a [binary relation](foundation.binary-relations.md) $<$ on $A$ valued in the [propositions](foundation-core.propositions.md), such that the relation $<$ is {{#concept "irreflexive"}} and transitive: + +- For any $x:A$ we have $x \nle x$. +- For any $x,y,z:A$ we have $yImports + +```agda +open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.negation +open import foundation.propositions +open import foundation.universe-levels +``` + +
+ +## Idea + +A {{#concept "strictly ordered type" Agda=Strictly-Ordered-Type}} consists of a type $A$, a [binary relation](foundation.binary-relations.md) $<$ on $A$ valued in the [propositions](foundation-core.propositions.md), such that the relation $<$ is {{#concept "irreflexive"}} and transitive: + +- For any $x:A$ we have $x \nle x$. +- For any $x,y,z:A$ we have $y Date: Wed, 8 Jan 2025 13:45:09 -0500 Subject: [PATCH 41/72] inflationary maps of strictly ordered types --- ...onary-maps-strictly-ordered-types.lagda.md | 162 ++++++++++++++++++ .../strict-order-preserving-maps.lagda.md | 153 +++++++++++++++++ 2 files changed, 315 insertions(+) create mode 100644 src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md create mode 100644 src/order-theory/strict-order-preserving-maps.lagda.md diff --git a/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md b/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md new file mode 100644 index 0000000000..069d80d2c3 --- /dev/null +++ b/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md @@ -0,0 +1,162 @@ +# Inflationary maps on a strictly ordered type + +```agda +module order-theory.inflationary-maps-strictly-ordered-types where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.strict-order-preserving-maps +open import order-theory.strictly-ordered-types +``` + +
+ +## Idea + +A map $f : P → P$ on a [strictly ordered type](order-theory.strictly-ordered-types.md) $P$ is said to be an +{{#concept "inflationary map" Disambiguation="strictly ordered type" Agda=inflationary-map-Strictly-Ordered-Type}} if the inequality + +$$ + x < f(x) +$$ + +holds for any element $x : P$. If $f$ is also [order preserving](order-theory.order-preserving-maps-strictly-ordered-types.md) we say that $f$ is an {{#concept "inflationary morphism" Disambiguation="strictly ordered type" Agda=inflationary-hom-Strictly-Ordered-Type}}. + +## Definitions + +### The predicate of being an inflationary map + +```agda +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + (f : type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P) + where + + is-inflationary-prop-map-Strictly-Ordered-Type : + Prop (l1 ⊔ l2) + is-inflationary-prop-map-Strictly-Ordered-Type = + Π-Prop + ( type-Strictly-Ordered-Type P) + ( λ x → le-prop-Strictly-Ordered-Type P x (f x)) + + is-inflationary-map-Strictly-Ordered-Type : + UU (l1 ⊔ l2) + is-inflationary-map-Strictly-Ordered-Type = + type-Prop is-inflationary-prop-map-Strictly-Ordered-Type + + is-prop-is-inflationary-map-Strictly-Ordered-Type : + is-prop is-inflationary-map-Strictly-Ordered-Type + is-prop-is-inflationary-map-Strictly-Ordered-Type = + is-prop-type-Prop is-inflationary-prop-map-Strictly-Ordered-Type +``` + +### The type of inflationary maps on a strictly ordered type + +```agda +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + where + + inflationary-map-Strictly-Ordered-Type : + UU (l1 ⊔ l2) + inflationary-map-Strictly-Ordered-Type = + type-subtype (is-inflationary-prop-map-Strictly-Ordered-Type P) + +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + (f : inflationary-map-Strictly-Ordered-Type P) + where + + map-inflationary-map-Strictly-Ordered-Type : + type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P + map-inflationary-map-Strictly-Ordered-Type = + pr1 f + + is-inflationary-inflationary-map-Strictly-Ordered-Type : + is-inflationary-map-Strictly-Ordered-Type P + ( map-inflationary-map-Strictly-Ordered-Type) + is-inflationary-inflationary-map-Strictly-Ordered-Type = + pr2 f +``` + +### The predicate on order preserving maps of being inflationary + +```agda +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + (f : hom-Strictly-Ordered-Type P P) + where + + is-inflationary-prop-hom-Strictly-Ordered-Type : + Prop (l1 ⊔ l2) + is-inflationary-prop-hom-Strictly-Ordered-Type = + is-inflationary-prop-map-Strictly-Ordered-Type P + ( map-hom-Strictly-Ordered-Type P P f) + + is-inflationary-hom-Strictly-Ordered-Type : + UU (l1 ⊔ l2) + is-inflationary-hom-Strictly-Ordered-Type = + is-inflationary-map-Strictly-Ordered-Type P + ( map-hom-Strictly-Ordered-Type P P f) + + is-prop-is-inflationary-hom-Strictly-Ordered-Type : + is-prop is-inflationary-hom-Strictly-Ordered-Type + is-prop-is-inflationary-hom-Strictly-Ordered-Type = + is-prop-is-inflationary-map-Strictly-Ordered-Type P + ( map-hom-Strictly-Ordered-Type P P f) +``` + +### The type of inflationary morphisms on a strictly ordered type + +```agda +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + where + + inflationary-hom-Strictly-Ordered-Type : + UU (l1 ⊔ l2) + inflationary-hom-Strictly-Ordered-Type = + type-subtype (is-inflationary-prop-hom-Strictly-Ordered-Type P) + +module _ + {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) + (f : inflationary-hom-Strictly-Ordered-Type P) + where + + hom-inflationary-hom-Strictly-Ordered-Type : + hom-Strictly-Ordered-Type P P + hom-inflationary-hom-Strictly-Ordered-Type = + pr1 f + + map-inflationary-hom-Strictly-Ordered-Type : + type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P + map-inflationary-hom-Strictly-Ordered-Type = + map-hom-Strictly-Ordered-Type P P hom-inflationary-hom-Strictly-Ordered-Type + + preserves-order-inflationary-hom-Strictly-Ordered-Type : + preserves-strict-order-map-Strictly-Ordered-Type P P + ( map-inflationary-hom-Strictly-Ordered-Type) + preserves-order-inflationary-hom-Strictly-Ordered-Type = + preserves-strict-order-hom-Strictly-Ordered-Type P P + ( hom-inflationary-hom-Strictly-Ordered-Type) + + is-inflationary-inflationary-hom-Strictly-Ordered-Type : + is-inflationary-map-Strictly-Ordered-Type P + ( map-inflationary-hom-Strictly-Ordered-Type) + is-inflationary-inflationary-hom-Strictly-Ordered-Type = + pr2 f + + inflationary-map-inflationary-hom-Strictly-Ordered-Type : + inflationary-map-Strictly-Ordered-Type P + pr1 inflationary-map-inflationary-hom-Strictly-Ordered-Type = + map-inflationary-hom-Strictly-Ordered-Type + pr2 inflationary-map-inflationary-hom-Strictly-Ordered-Type = + is-inflationary-inflationary-hom-Strictly-Ordered-Type +``` diff --git a/src/order-theory/strict-order-preserving-maps.lagda.md b/src/order-theory/strict-order-preserving-maps.lagda.md new file mode 100644 index 0000000000..0b8e200905 --- /dev/null +++ b/src/order-theory/strict-order-preserving-maps.lagda.md @@ -0,0 +1,153 @@ +# Strict order preserving maps + +```agda +module order-theory.strict-order-preserving-maps where +``` + +
Imports + +```agda +open import foundation.binary-relations +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.strictly-ordered-sets +open import order-theory.strictly-ordered-types +``` + +
+ +## Idea + +Consider two [strictly ordered types](order-theory.strictly-ordered-types.md) $P$ and $Q$, with orderings $<_P$ and $<_Q$ respectively. A {{#concept "strict order preserving map" Agda=hom-Strictly-Ordered-Type}} consists of map $f : P → Q$ between their underlying types such that for any two elements $x<_P y$ in $P$ we have $f(x)<_Q f(y)$ in $Q$. + +## Definitions + +### The predicate on maps between strictly ordered types of preserving the strict ordering + +```agda +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Type l1 l2) (Q : Strictly-Ordered-Type l3 l4) + (f : type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type Q) + where + + preserves-strict-order-prop-map-Strictly-Ordered-Type : + Prop (l1 ⊔ l2 ⊔ l4) + preserves-strict-order-prop-map-Strictly-Ordered-Type = + Π-Prop + ( type-Strictly-Ordered-Type P) + ( λ x → + Π-Prop + ( type-Strictly-Ordered-Type P) + ( λ y → + hom-Prop + ( le-prop-Strictly-Ordered-Type P x y) + ( le-prop-Strictly-Ordered-Type Q (f x) (f y)))) + + preserves-strict-order-map-Strictly-Ordered-Type : + UU (l1 ⊔ l2 ⊔ l4) + preserves-strict-order-map-Strictly-Ordered-Type = + type-Prop preserves-strict-order-prop-map-Strictly-Ordered-Type + + is-prop-preserves-strict-order-map-Strictly-Ordered-Type : + is-prop preserves-strict-order-map-Strictly-Ordered-Type + is-prop-preserves-strict-order-map-Strictly-Ordered-Type = + is-prop-type-Prop preserves-strict-order-prop-map-Strictly-Ordered-Type +``` + +### The type of strict order preserving maps between strictly ordered types + +```agda +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Type l1 l2) (Q : Strictly-Ordered-Type l3 l4) + where + + hom-Strictly-Ordered-Type : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + hom-Strictly-Ordered-Type = + type-subtype (preserves-strict-order-prop-map-Strictly-Ordered-Type P Q) + +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Type l1 l2) (Q : Strictly-Ordered-Type l3 l4) + (f : hom-Strictly-Ordered-Type P Q) + where + + map-hom-Strictly-Ordered-Type : + type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type Q + map-hom-Strictly-Ordered-Type = + pr1 f + + preserves-strict-order-hom-Strictly-Ordered-Type : + preserves-strict-order-map-Strictly-Ordered-Type P Q + ( map-hom-Strictly-Ordered-Type) + preserves-strict-order-hom-Strictly-Ordered-Type = + pr2 f +``` + +### The predicate on maps between strictly ordered sets of preserving the strict ordering + +```agda +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Set l1 l2) (Q : Strictly-Ordered-Set l3 l4) + (f : type-Strictly-Ordered-Set P → type-Strictly-Ordered-Set Q) + where + + preserves-strict-order-prop-map-Strictly-Ordered-Set : + Prop (l1 ⊔ l2 ⊔ l4) + preserves-strict-order-prop-map-Strictly-Ordered-Set = + preserves-strict-order-prop-map-Strictly-Ordered-Type + ( strictly-ordered-type-Strictly-Ordered-Set P) + ( strictly-ordered-type-Strictly-Ordered-Set Q) + ( f) + + preserves-strict-order-map-Strictly-Ordered-Set : + UU (l1 ⊔ l2 ⊔ l4) + preserves-strict-order-map-Strictly-Ordered-Set = + preserves-strict-order-map-Strictly-Ordered-Type + ( strictly-ordered-type-Strictly-Ordered-Set P) + ( strictly-ordered-type-Strictly-Ordered-Set Q) + ( f) + + is-prop-preserves-strict-order-map-Strictly-Ordered-Set : + is-prop preserves-strict-order-map-Strictly-Ordered-Set + is-prop-preserves-strict-order-map-Strictly-Ordered-Set = + is-prop-preserves-strict-order-map-Strictly-Ordered-Type + ( strictly-ordered-type-Strictly-Ordered-Set P) + ( strictly-ordered-type-Strictly-Ordered-Set Q) + ( f) +``` + +### The type of strict order preserving maps between strictly ordered sets + +```agda +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Set l1 l2) (Q : Strictly-Ordered-Set l3 l4) + where + + hom-Strictly-Ordered-Set : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + hom-Strictly-Ordered-Set = + type-subtype (preserves-strict-order-prop-map-Strictly-Ordered-Set P Q) + +module _ + {l1 l2 l3 l4 : Level} + (P : Strictly-Ordered-Set l1 l2) (Q : Strictly-Ordered-Set l3 l4) + (f : hom-Strictly-Ordered-Set P Q) + where + + map-hom-Strictly-Ordered-Set : + type-Strictly-Ordered-Set P → type-Strictly-Ordered-Set Q + map-hom-Strictly-Ordered-Set = + pr1 f + + preserves-strict-order-hom-Strictly-Ordered-Set : + preserves-strict-order-map-Strictly-Ordered-Set P Q + ( map-hom-Strictly-Ordered-Set) + preserves-strict-order-hom-Strictly-Ordered-Set = + pr2 f +``` From 4b86dd11055ba497ad428266c9f1933e99f84c0c Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Thu, 9 Jan 2025 17:48:49 -0500 Subject: [PATCH 42/72] lower and upper bounds of structured natural numbers --- ...-strong-induction-natural-numbers.lagda.md | 9 +- .../decidable-types.lagda.md | 73 +++--- .../inequality-natural-numbers.lagda.md | 43 ++-- .../lower-bounds-natural-numbers.lagda.md | 233 +++++++++++++++++- .../strong-induction-natural-numbers.lagda.md | 16 +- .../upper-bounds-natural-numbers.lagda.md | 232 ++++++++++++++++- ...rdering-principle-natural-numbers.lagda.md | 206 +++++++++++----- 7 files changed, 667 insertions(+), 145 deletions(-) diff --git a/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md b/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md index d2478c6ce3..5de7713d58 100644 --- a/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md +++ b/src/elementary-number-theory/based-strong-induction-natural-numbers.lagda.md @@ -101,7 +101,7 @@ succ-based-strong-ind-ℕ : ((x : ℕ) → leq-ℕ k x → based-□-≤-ℕ k P x → P (succ-ℕ x)) → (n : ℕ) → leq-ℕ k n → based-□-≤-ℕ k P n → based-□-≤-ℕ k P (succ-ℕ n) succ-based-strong-ind-ℕ k P pS n N f m M H = - cases-succ-based-strong-ind-ℕ k P pS n N f m M (cases-leq-succ-ℕ H) + cases-succ-based-strong-ind-ℕ k P pS n N f m M (decide-leq-succ-ℕ m n H) cases-htpy-succ-based-strong-ind-ℕ : {l : Level} (k : ℕ) (P : ℕ → UU l) @@ -122,7 +122,7 @@ htpy-succ-based-strong-ind-ℕ : (m : ℕ) (M : k ≤-ℕ m) (H : leq-ℕ m (succ-ℕ n)) (K : leq-ℕ m n) → succ-based-strong-ind-ℕ k P pS n N f m M H = f m M K htpy-succ-based-strong-ind-ℕ k P pS n N f m M H = - cases-htpy-succ-based-strong-ind-ℕ k P pS n N f m M (cases-leq-succ-ℕ H) + cases-htpy-succ-based-strong-ind-ℕ k P pS n N f m M (decide-leq-succ-ℕ m n H) cases-eq-succ-based-strong-ind-ℕ : {l : Level} (k : ℕ) (P : ℕ → UU l) @@ -144,7 +144,8 @@ eq-succ-based-strong-ind-ℕ : (H : leq-ℕ (succ-ℕ n) (succ-ℕ n)) → succ-based-strong-ind-ℕ k P pS n N f (succ-ℕ n) M H = pS n N f eq-succ-based-strong-ind-ℕ k P pS n N f M H = - cases-eq-succ-based-strong-ind-ℕ k P pS n N f M (cases-leq-succ-ℕ H) + cases-eq-succ-based-strong-ind-ℕ k P pS n N f M + ( decide-leq-succ-ℕ (succ-ℕ n) n H) ``` ### The inductive step in the proof of based strong induction @@ -353,7 +354,7 @@ eq-inductive-step-compute-succ-based-strong-ind-ℕ k P p0 pS n N m M = cases-eq-inductive-step-compute-succ-based-strong-ind-ℕ k P p0 pS i I' I m M ( J) - ( cases-leq-succ-ℕ J) + ( decide-leq-succ-ℕ m i J) ( α)) ( n) ( N) diff --git a/src/elementary-number-theory/decidable-types.lagda.md b/src/elementary-number-theory/decidable-types.lagda.md index 047ada9a91..c9f26e8218 100644 --- a/src/elementary-number-theory/decidable-types.lagda.md +++ b/src/elementary-number-theory/decidable-types.lagda.md @@ -35,33 +35,36 @@ decidable. ```agda is-decidable-Σ-ℕ : - {l : Level} (m : ℕ) (P : ℕ → UU l) (d : is-decidable-fam P) → - is-decidable (Σ ℕ (λ x → (leq-ℕ m x) × (P x))) → is-decidable (Σ ℕ P) -is-decidable-Σ-ℕ m P d (inl (pair x (pair l p))) = inl (pair x p) + {l : Level} (m : ℕ) (P : ℕ → UU l) → is-decidable-fam P → + is-decidable (Σ ℕ (λ x → m ≤-ℕ x × P x)) → is-decidable (Σ ℕ P) +is-decidable-Σ-ℕ m P d (inl (x , l , p)) = + inl (x , p) is-decidable-Σ-ℕ zero-ℕ P d (inr f) = - inr (λ p → f (pair (pr1 p) (pair star (pr2 p)))) -is-decidable-Σ-ℕ (succ-ℕ m) P d (inr f) with d zero-ℕ -... | inl p = inl (pair zero-ℕ p) + inr (λ p → f (pr1 p , star , pr2 p)) +is-decidable-Σ-ℕ (succ-ℕ m) P d (inr f) + with + d 0 +... | inl p = inl (0 , p) ... | inr g with is-decidable-Σ-ℕ m ( P ∘ succ-ℕ) ( λ x → d (succ-ℕ x)) - ( inr (λ p → f (pair (succ-ℕ (pr1 p)) (pr2 p)))) -... | inl p = inl (pair (succ-ℕ (pr1 p)) (pr2 p)) + ( inr (λ p → f (succ-ℕ (pr1 p) , pr2 p))) +... | inl p = inl (succ-ℕ (pr1 p) , pr2 p) ... | inr h = inr α where α : Σ ℕ P → empty - α (pair zero-ℕ p) = g p - α (pair (succ-ℕ x) p) = h (pair x p) + α (zero-ℕ , p) = g p + α (succ-ℕ x , p) = h (x , p) ``` ### Bounded sums of decidable families over ℕ are decidable ```agda is-decidable-bounded-Σ-ℕ : - {l1 l2 : Level} (m : ℕ) (P : ℕ → UU l1) (Q : ℕ → UU l2) - (dP : is-decidable-fam P) (dQ : is-decidable-fam Q) - (H : is-upper-bound-ℕ P m) → is-decidable (Σ ℕ (λ x → (P x) × (Q x))) + {l1 l2 : Level} (m : ℕ) (P : ℕ → UU l1) (Q : ℕ → UU l2) → + is-decidable-fam P → is-decidable-fam Q → + is-upper-bound-ℕ P m → is-decidable (Σ ℕ (λ x → (P x) × (Q x))) is-decidable-bounded-Σ-ℕ m P Q dP dQ H = is-decidable-Σ-ℕ ( succ-ℕ m) @@ -76,11 +79,11 @@ is-decidable-bounded-Σ-ℕ m P Q dP dQ H = ( pr1 (pr2 p)))) is-decidable-bounded-Σ-ℕ' : - {l : Level} (m : ℕ) (P : ℕ → UU l) (d : is-decidable-fam P) → - is-decidable (Σ ℕ (λ x → (leq-ℕ x m) × (P x))) + {l : Level} (m : ℕ) (P : ℕ → UU l) → + is-decidable-fam P → is-decidable (Σ ℕ (λ x → x ≤-ℕ m × P x)) is-decidable-bounded-Σ-ℕ' m P d = is-decidable-bounded-Σ-ℕ m - ( λ x → leq-ℕ x m) + ( λ x → x ≤-ℕ m) ( P) ( λ x → is-decidable-leq-ℕ x m) ( d) @@ -91,20 +94,19 @@ is-decidable-bounded-Σ-ℕ' m P d = ```agda is-decidable-strictly-bounded-Σ-ℕ : - {l1 l2 : Level} (m : ℕ) (P : ℕ → UU l1) (Q : ℕ → UU l2) - (dP : is-decidable-fam P) (dQ : is-decidable-fam Q) - (H : is-strict-upper-bound-ℕ P m) → + {l1 l2 : Level} (m : ℕ) (P : ℕ → UU l1) (Q : ℕ → UU l2) → + is-decidable-fam P → is-decidable-fam Q → is-strict-upper-bound-ℕ P m → is-decidable (Σ ℕ (λ x → (P x) × (Q x))) is-decidable-strictly-bounded-Σ-ℕ m P Q dP dQ H = is-decidable-bounded-Σ-ℕ m P Q dP dQ ( is-upper-bound-is-strict-upper-bound-ℕ P m H) is-decidable-strictly-bounded-Σ-ℕ' : - {l : Level} (m : ℕ) (P : ℕ → UU l) (d : is-decidable-fam P) → - is-decidable (Σ ℕ (λ x → (le-ℕ x m) × (P x))) + {l : Level} (m : ℕ) (P : ℕ → UU l) → + is-decidable-fam P → is-decidable (Σ ℕ (λ x → x <-ℕ m × P x)) is-decidable-strictly-bounded-Σ-ℕ' m P d = is-decidable-strictly-bounded-Σ-ℕ m - ( λ x → le-ℕ x m) + ( λ x → x <-ℕ m) ( P) ( λ x → is-decidable-le-ℕ x m) ( d) @@ -115,8 +117,9 @@ is-decidable-strictly-bounded-Σ-ℕ' m P d = ```agda is-decidable-Π-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (m : ℕ) → - is-decidable ((x : ℕ) → (leq-ℕ m x) → P x) → is-decidable ((x : ℕ) → P x) + {l : Level} (P : ℕ → UU l) → + is-decidable-fam P → (m : ℕ) → is-decidable ((x : ℕ) → m ≤-ℕ x → P x) → + is-decidable ((x : ℕ) → P x) is-decidable-Π-ℕ P d zero-ℕ (inr nH) = inr (λ f → nH (λ x y → f x)) is-decidable-Π-ℕ P d zero-ℕ (inl H) = inl (λ x → H x (leq-zero-ℕ x)) is-decidable-Π-ℕ P d (succ-ℕ m) (inr nH) = inr (λ f → nH (λ x y → f x)) @@ -136,8 +139,9 @@ is-decidable-Π-ℕ P d (succ-ℕ m) (inl H) with d zero-ℕ ```agda is-decidable-bounded-Π-ℕ : - {l1 l2 : Level} (P : ℕ → UU l1) (Q : ℕ → UU l2) (dP : is-decidable-fam P) → - (dQ : is-decidable-fam Q) (m : ℕ) (H : is-upper-bound-ℕ P m) → + {l1 l2 : Level} (P : ℕ → UU l1) (Q : ℕ → UU l2) → + is-decidable-fam P → is-decidable-fam Q → + (m : ℕ) → is-upper-bound-ℕ P m → is-decidable ((x : ℕ) → P x → Q x) is-decidable-bounded-Π-ℕ P Q dP dQ m H = is-decidable-Π-ℕ @@ -147,11 +151,11 @@ is-decidable-bounded-Π-ℕ P Q dP dQ m H = ( inl (λ x l p → ex-falso (contradiction-leq-ℕ x m (H x p) l))) is-decidable-bounded-Π-ℕ' : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (m : ℕ) → - is-decidable ((x : ℕ) → (leq-ℕ x m) → P x) + {l : Level} (P : ℕ → UU l) → + is-decidable-fam P → (m : ℕ) → is-decidable ((x : ℕ) → x ≤-ℕ m → P x) is-decidable-bounded-Π-ℕ' P d m = is-decidable-bounded-Π-ℕ - ( λ x → leq-ℕ x m) + ( λ x → x ≤-ℕ m) ( P) ( λ x → is-decidable-leq-ℕ x m) ( d) @@ -163,18 +167,19 @@ is-decidable-bounded-Π-ℕ' P d m = ```agda is-decidable-strictly-bounded-Π-ℕ : - {l1 l2 : Level} (P : ℕ → UU l1) (Q : ℕ → UU l2) (dP : is-decidable-fam P) → - (dQ : is-decidable-fam Q) (m : ℕ) (H : is-strict-upper-bound-ℕ P m) → + {l1 l2 : Level} (P : ℕ → UU l1) (Q : ℕ → UU l2) → + is-decidable-fam P → is-decidable-fam Q → + (m : ℕ) → is-strict-upper-bound-ℕ P m → is-decidable ((x : ℕ) → P x → Q x) is-decidable-strictly-bounded-Π-ℕ P Q dP dQ m H = is-decidable-bounded-Π-ℕ P Q dP dQ m (λ x p → leq-le-ℕ x m (H x p)) is-decidable-strictly-bounded-Π-ℕ' : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (m : ℕ) → - is-decidable ((x : ℕ) → le-ℕ x m → P x) + {l : Level} (P : ℕ → UU l) → is-decidable-fam P → + (m : ℕ) → is-decidable ((x : ℕ) → x <-ℕ m → P x) is-decidable-strictly-bounded-Π-ℕ' P d m = is-decidable-strictly-bounded-Π-ℕ - ( λ x → le-ℕ x m) + ( λ x → x <-ℕ m) ( P) ( λ x → is-decidable-le-ℕ x m) ( d) diff --git a/src/elementary-number-theory/inequality-natural-numbers.lagda.md b/src/elementary-number-theory/inequality-natural-numbers.lagda.md index e559497525..49ed531c42 100644 --- a/src/elementary-number-theory/inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/inequality-natural-numbers.lagda.md @@ -239,6 +239,14 @@ succ-leq-ℕ zero-ℕ = star succ-leq-ℕ (succ-ℕ n) = succ-leq-ℕ n ``` +### If `m` is less than `n`, then it is less than `n+1` + +```agda +leq-succ-leq-ℕ : + (m n : ℕ) → m ≤-ℕ n → m ≤-ℕ succ-ℕ n +leq-succ-leq-ℕ m n p = transitive-leq-ℕ m n (succ-ℕ n) (succ-leq-ℕ n) p +``` + ### Any natural number less than or equal to `n+1` is either less than or equal to `n` or it is `n+1` ```agda @@ -250,14 +258,23 @@ decide-leq-succ-ℕ (succ-ℕ m) zero-ℕ l = inr (ap succ-ℕ (is-zero-leq-zero-ℕ m l)) decide-leq-succ-ℕ (succ-ℕ m) (succ-ℕ n) l = map-coproduct id (ap succ-ℕ) (decide-leq-succ-ℕ m n l) + +decide-leq-refl-succ-ℕ : + {n : ℕ} → decide-leq-succ-ℕ (succ-ℕ n) n (refl-leq-ℕ n) = inr refl +decide-leq-refl-succ-ℕ {zero-ℕ} = refl +decide-leq-refl-succ-ℕ {succ-ℕ n} = + ap (map-coproduct id (ap succ-ℕ)) decide-leq-refl-succ-ℕ ``` -### If `m` is less than `n`, then it is less than `n+1` +### The inequality `m ≤ n` holds if and only if either `m = n` or the inequality `m + 1 ≤ n` holds ```agda -leq-succ-leq-ℕ : - (m n : ℕ) → m ≤-ℕ n → m ≤-ℕ succ-ℕ n -leq-succ-leq-ℕ m n p = transitive-leq-ℕ m n (succ-ℕ n) (succ-leq-ℕ n) p +decide-leq-ℕ : + (m n : ℕ) → m ≤-ℕ n → (m = n) + (succ-ℕ m ≤-ℕ n) +decide-leq-ℕ m zero-ℕ H = inl (is-zero-leq-zero-ℕ m H) +decide-leq-ℕ zero-ℕ (succ-ℕ n) H = inr (leq-zero-ℕ n) +decide-leq-ℕ (succ-ℕ m) (succ-ℕ n) H = + map-coproduct (ap succ-ℕ) id (decide-leq-ℕ m n H) ``` ### The successor of `n` is not less than or equal to `n` @@ -269,24 +286,6 @@ neg-succ-leq-ℕ zero-ℕ = id neg-succ-leq-ℕ (succ-ℕ n) = neg-succ-leq-ℕ n ``` -### If `m ≤ n + 1` then either `m ≤ n` or `m = n + 1` - -```agda -cases-leq-succ-ℕ : - {m n : ℕ} → m ≤-ℕ succ-ℕ n → (m ≤-ℕ n) + (m = succ-ℕ n) -cases-leq-succ-ℕ {zero-ℕ} {n} star = inl star -cases-leq-succ-ℕ {succ-ℕ m} {zero-ℕ} p = - inr (ap succ-ℕ (antisymmetric-leq-ℕ m zero-ℕ p star)) -cases-leq-succ-ℕ {succ-ℕ m} {succ-ℕ n} p = - map-coproduct id (ap succ-ℕ) (cases-leq-succ-ℕ p) - -cases-leq-succ-reflexive-leq-ℕ : - {n : ℕ} → cases-leq-succ-ℕ {succ-ℕ n} {n} (refl-leq-ℕ n) = inr refl -cases-leq-succ-reflexive-leq-ℕ {zero-ℕ} = refl -cases-leq-succ-reflexive-leq-ℕ {succ-ℕ n} = - ap (map-coproduct id (ap succ-ℕ)) cases-leq-succ-reflexive-leq-ℕ -``` - ### `m ≤ n` if and only if `n + 1 ≰ m` ```agda diff --git a/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md index 1a84c70750..1603e4d0f5 100644 --- a/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md @@ -9,9 +9,20 @@ module elementary-number-theory.lower-bounds-natural-numbers where ```agda open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.double-negation +open import foundation.empty-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-extensionality open import foundation.propositions +open import foundation.transport-along-identifications +open import foundation.type-arithmetic-empty-type open import foundation.universe-levels ``` @@ -19,16 +30,64 @@ open import foundation.universe-levels ## Idea -A lower bound for a type family `P` over the natural numbers is a natural number -`n` such that `P x → n ≤ x` for all `x : ℕ`. +Consider a type family $P$ over the [natural numbers](elementary-number-theory.natural-numbers.md). A {{#concept "structured natural number"}} is simply a natural number $n$ equipped with an element $P(n)$. In this file we consider various upper bounds for structured natural numbers, and relations between them. This file builds the prerequisite infrastructure for the [well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) of the natural numbers, and its direct consequences. -## Definition +- A natural number $n$ is said to be a {{#concept "lower bound" Disambiguation="structured natural numbers" Agda=is-lower-bound-ℕ}} if there is a +function from $P(x)$ to the type $n \leq x$ for all $x : \mathbb{N}$. +- A natural number $n$ is said to be a {{#concept "strict lower bound" Disambiguation="structured natural numbers" Agda=is-strict-lower-bound-ℕ}} if there is a +function from $P(x)$ to the type $n < x$ for all $x : \mathbb{N}$. +- A natural number $n$ is said to be a {{#concept "largest lower bound" Disambiguation="structured natural numbers" Agda=is-largest-lower-bound-ℕ}} if any natural number $x$ is a lower bound if and only if $x \leq n$. +- A natural number $n$ is said to be a {{#concept "largest strict lower bound" Disambiguation="structured natural numbers" Agda=is-largest-strict-lower-bound-ℕ}} if any natural number $x$ is a strict lower bound if and only if $x \leq n$. + +## Definitions + +### Lower bounds ```agda is-lower-bound-ℕ : - {l : Level} (P : ℕ → UU l) (n : ℕ) → UU l + {l : Level} (P : ℕ → UU l) → ℕ → UU l is-lower-bound-ℕ P n = - (m : ℕ) → P m → leq-ℕ n m + (m : ℕ) → P m → n ≤-ℕ m +``` + +### Strict lower bounds + +```agda +is-strict-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → UU l +is-strict-lower-bound-ℕ P n = + (m : ℕ) → P m → n <-ℕ m +``` + +### Largest lower bounds + +```agda +is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-largest-lower-bound-ℕ P n = + (x : ℕ) → is-lower-bound-ℕ P x ↔ x ≤-ℕ n + +is-lower-bound-is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-largest-lower-bound-ℕ P n → is-lower-bound-ℕ P n +is-lower-bound-is-largest-lower-bound-ℕ P n H = + backward-implication (H n) (refl-leq-ℕ n) + +leq-is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-largest-lower-bound-ℕ P n → + (m : ℕ) → is-lower-bound-ℕ P m → m ≤-ℕ n +leq-is-largest-lower-bound-ℕ P n H m = + forward-implication (H m) +``` + +### Largest strict lower bounds + +```agda +is-largest-strict-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-largest-strict-lower-bound-ℕ P n = + (x : ℕ) → is-strict-lower-bound-ℕ P x ↔ x ≤-ℕ n ``` ## Properties @@ -37,15 +96,163 @@ is-lower-bound-ℕ P n = ```agda module _ - {l1 : Level} {P : ℕ → UU l1} + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-lower-bound-ℕ : + (n : ℕ) → is-prop (is-lower-bound-ℕ P n) + is-prop-is-lower-bound-ℕ n = + is-prop-Π (λ x → is-prop-function-type (is-prop-leq-ℕ n x)) + + is-lower-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-lower-bound-ℕ-Prop n) = is-lower-bound-ℕ P n + pr2 (is-lower-bound-ℕ-Prop n) = is-prop-is-lower-bound-ℕ n +``` + +### Being a strict lower bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-strict-lower-bound-ℕ : + (n : ℕ) → is-prop (is-strict-lower-bound-ℕ P n) + is-prop-is-strict-lower-bound-ℕ n = + is-prop-Π (λ x → is-prop-function-type (is-prop-le-ℕ n x)) + + is-strict-lower-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-strict-lower-bound-ℕ-Prop n) = is-strict-lower-bound-ℕ P n + pr2 (is-strict-lower-bound-ℕ-Prop n) = is-prop-is-strict-lower-bound-ℕ n +``` + +### Being a largest lower bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-largest-lower-bound-ℕ : + (n : ℕ) → is-prop (is-largest-lower-bound-ℕ P n) + is-prop-is-largest-lower-bound-ℕ n = + is-prop-Π + ( λ x → is-prop-iff-Prop (is-lower-bound-ℕ-Prop P x) (leq-ℕ-Prop x n)) + + is-largest-lower-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-largest-lower-bound-ℕ-Prop n) = is-largest-lower-bound-ℕ P n + pr2 (is-largest-lower-bound-ℕ-Prop n) = is-prop-is-largest-lower-bound-ℕ n +``` + +### Being a largest strict lower bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) where - abstract - is-prop-is-lower-bound-ℕ : (x : ℕ) → is-prop (is-lower-bound-ℕ P x) - is-prop-is-lower-bound-ℕ x = - is-prop-Π (λ y → is-prop-function-type (is-prop-leq-ℕ x y)) + is-prop-is-largest-strict-lower-bound-ℕ : + (n : ℕ) → is-prop (is-largest-strict-lower-bound-ℕ P n) + is-prop-is-largest-strict-lower-bound-ℕ n = + is-prop-Π + ( λ x → + is-prop-iff-Prop (is-strict-lower-bound-ℕ-Prop P x) (leq-ℕ-Prop x n)) + + is-largest-strict-lower-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-largest-strict-lower-bound-ℕ-Prop n) = + is-largest-strict-lower-bound-ℕ P n + pr2 (is-largest-strict-lower-bound-ℕ-Prop n) = + is-prop-is-largest-strict-lower-bound-ℕ n +``` + +### A strict lower bound is a lower bound + +```agda +is-lower-bound-is-strict-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-strict-lower-bound-ℕ P n → is-lower-bound-ℕ P n +is-lower-bound-is-strict-lower-bound-ℕ P n H x p = + leq-le-ℕ n x (H x p) +``` + +### Any two largest lower bounds are equal + +```agda +eq-is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (m n : ℕ) → + is-largest-lower-bound-ℕ P m → is-largest-lower-bound-ℕ P n → m = n +eq-is-largest-lower-bound-ℕ P m n H K = + antisymmetric-leq-ℕ m n + ( leq-is-largest-lower-bound-ℕ P n K m + ( is-lower-bound-is-largest-lower-bound-ℕ P m H)) + ( leq-is-largest-lower-bound-ℕ P m H n + ( is-lower-bound-is-largest-lower-bound-ℕ P n K)) +``` + +### If $n$ is a largest lower bound, then $n + 1$ is not a largest lower bound - is-lower-bound-ℕ-Prop : (x : ℕ) → Prop l1 - pr1 (is-lower-bound-ℕ-Prop x) = is-lower-bound-ℕ P x - pr2 (is-lower-bound-ℕ-Prop x) = is-prop-is-lower-bound-ℕ x +```agda +is-largest-lower-bound-succ-is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-largest-lower-bound-ℕ P n → ¬ is-largest-lower-bound-ℕ P (succ-ℕ n) +is-largest-lower-bound-succ-is-largest-lower-bound-ℕ P n H K = + has-no-fixed-points-succ-ℕ n + ( eq-is-largest-lower-bound-ℕ P (succ-ℕ n) n K H) +``` + +### If $n + 1$ is a lower bound and $P (n + 1)$ is empty, then $n$ is a lower bound + +```agda +increase-is-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-lower-bound-ℕ P n → ¬ P n → is-lower-bound-ℕ P (succ-ℕ n) +increase-is-lower-bound-ℕ P n H f m p = + map-left-unit-law-coproduct-is-empty + ( n = m) + ( succ-ℕ n ≤-ℕ m) + ( λ α → f (tr P (inv α) p)) + ( decide-leq-ℕ n m (H m p)) +``` + +### Any successor natural number $n + 1$ such that $P (n + 1)$ is empty is not a largest lower bound + +```agda +is-not-largest-lower-bound-is-empty-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + ¬ P n → ¬ is-largest-lower-bound-ℕ P n +is-not-largest-lower-bound-is-empty-ℕ P n f H = + neg-succ-leq-ℕ n + ( leq-is-largest-lower-bound-ℕ P + ( n) + ( H) + ( succ-ℕ n) + ( increase-is-lower-bound-ℕ P n + ( is-lower-bound-is-largest-lower-bound-ℕ P n H) + ( f))) +``` + +### The type $P$ at the largest lower bound is nonempty + +```agda +is-nonempty-structure-is-in-family-is-largest-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-largest-lower-bound-ℕ P n → is-nonempty (P n) +is-nonempty-structure-is-in-family-is-largest-lower-bound-ℕ P n H f = + is-not-largest-lower-bound-is-empty-ℕ P n f H +``` + +### Any largest lower bound of a family of decidable types over $\mathbb{N}$ is structured in $P$ + +```agda +structure-largest-lower-bound-is-decidable-fam-ℕ : + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → + (n : ℕ) → is-largest-lower-bound-ℕ P n → P n +structure-largest-lower-bound-is-decidable-fam-ℕ P d n H = + double-negation-elim-is-decidable + ( d n) + ( is-nonempty-structure-is-in-family-is-largest-lower-bound-ℕ P n H) ``` diff --git a/src/elementary-number-theory/strong-induction-natural-numbers.lagda.md b/src/elementary-number-theory/strong-induction-natural-numbers.lagda.md index 6a0bf777c5..bd4d82a5c3 100644 --- a/src/elementary-number-theory/strong-induction-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strong-induction-natural-numbers.lagda.md @@ -74,7 +74,7 @@ succ-strong-ind-ℕ : {l : Level} (P : ℕ → UU l) → ((k : ℕ) → (□-≤-ℕ P k) → P (succ-ℕ k)) → (k : ℕ) → (□-≤-ℕ P k) → (□-≤-ℕ P (succ-ℕ k)) succ-strong-ind-ℕ P pS k H m p = - cases-succ-strong-ind-ℕ P pS k H m (cases-leq-succ-ℕ p) + cases-succ-strong-ind-ℕ P pS k H m (decide-leq-succ-ℕ m k p) cases-htpy-succ-strong-ind-ℕ : {l : Level} (P : ℕ → UU l) (pS : (k : ℕ) → (□-≤-ℕ P k) → P (succ-ℕ k)) → @@ -93,7 +93,7 @@ htpy-succ-strong-ind-ℕ : ( succ-strong-ind-ℕ P pS k H m p) = ( H m q) htpy-succ-strong-ind-ℕ P pS k H m p q = - cases-htpy-succ-strong-ind-ℕ P pS k H m (cases-leq-succ-ℕ p) q + cases-htpy-succ-strong-ind-ℕ P pS k H m (decide-leq-succ-ℕ m k p) q cases-eq-succ-strong-ind-ℕ : {l : Level} (P : ℕ → UU l) (pS : (k : ℕ) → (□-≤-ℕ P k) → P (succ-ℕ k)) → @@ -113,7 +113,7 @@ eq-succ-strong-ind-ℕ : ( (succ-strong-ind-ℕ P pS k H (succ-ℕ k) p)) = ( pS k H) eq-succ-strong-ind-ℕ P pS k H p = - cases-eq-succ-strong-ind-ℕ P pS k H (cases-leq-succ-ℕ p) + cases-eq-succ-strong-ind-ℕ P pS k H (decide-leq-succ-ℕ (succ-ℕ k) k p) ``` ### The inductive step @@ -161,7 +161,7 @@ cases-eq-compute-succ-strong-ind-ℕ : ( m : ℕ) (p : leq-ℕ m n) → ( induction-strong-ind-ℕ P (zero-strong-ind-ℕ P p0) ( λ k z m₁ z₁ → - cases-succ-strong-ind-ℕ P pS k z m₁ (cases-leq-succ-ℕ z₁)) + cases-succ-strong-ind-ℕ P pS k z m₁ (decide-leq-succ-ℕ m₁ k z₁)) n m p) = ( strong-ind-ℕ P p0 pS m)) → ( m : ℕ) (p : leq-ℕ m (succ-ℕ n)) → @@ -190,10 +190,10 @@ cases-eq-compute-succ-strong-ind-ℕ P p0 pS n α .(succ-ℕ n) p (inr refl) = ( induction-strong-ind-ℕ P ( zero-strong-ind-ℕ P p0) ( λ k H m p₁ → - cases-succ-strong-ind-ℕ P pS k H m (cases-leq-succ-ℕ p₁)) + cases-succ-strong-ind-ℕ P pS k H m (decide-leq-succ-ℕ m k p₁)) n) ( succ-ℕ n)) - ( cases-leq-succ-reflexive-leq-ℕ))) + ( decide-leq-refl-succ-ℕ))) eq-compute-succ-strong-ind-ℕ : { l : Level} (P : ℕ → UU l) (p0 : P zero-ℕ) → @@ -202,14 +202,14 @@ eq-compute-succ-strong-ind-ℕ : ( m : ℕ) (p : leq-ℕ m n) → ( induction-strong-ind-ℕ P (zero-strong-ind-ℕ P p0) ( λ k z m₁ z₁ → - cases-succ-strong-ind-ℕ P pS k z m₁ (cases-leq-succ-ℕ z₁)) + cases-succ-strong-ind-ℕ P pS k z m₁ (decide-leq-succ-ℕ m₁ k z₁)) n m p) = ( strong-ind-ℕ P p0 pS m) eq-compute-succ-strong-ind-ℕ P p0 pS zero-ℕ zero-ℕ _ = refl eq-compute-succ-strong-ind-ℕ P p0 pS (succ-ℕ n) m p = cases-eq-compute-succ-strong-ind-ℕ P p0 pS n ( eq-compute-succ-strong-ind-ℕ P p0 pS n) m p - ( cases-leq-succ-ℕ p) + ( decide-leq-succ-ℕ m n p) compute-succ-strong-ind-ℕ : { l : Level} (P : ℕ → UU l) (p0 : P zero-ℕ) → diff --git a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md index 2a0529164b..1c0694989b 100644 --- a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md @@ -1,4 +1,4 @@ -# Upper bounds for type families over the natural numbers +# Bounds for structured natural numbers ```agda module elementary-number-theory.upper-bounds-natural-numbers where @@ -11,6 +11,18 @@ open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.double-negation +open import foundation.empty-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-extensionality +open import foundation.propositions +open import foundation.transport-along-identifications +open import foundation.type-arithmetic-empty-type open import foundation.universe-levels ``` @@ -18,19 +30,24 @@ open import foundation.universe-levels ## Idea -A type family over the natural numbers has an upper bound `n`, if there is a -function from `P x` to the type `x ≤ n` for all `x : ℕ`. Similar for strict -upper bounds. +Consider a type family $P$ over the [natural numbers](elementary-number-theory.natural-numbers.md). A {{#concept "structured natural number"}} is simply a natural number $n$ equipped with an element $P(n)$. In this file we consider various upper bounds for structured natural numbers, and relations between them. This file builds the prerequisite infrastructure for the [well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) of the natural numbers, and its direct consequences. -## Definition +- A natural number $n$ is said to be an {{#concept "upper bound" Disambiguation="structured natural numbers" Agda=is-upper-bound-ℕ}} if there is a +function from $P(x)$ to the type $x \leq n$ for all $x : \mathbb{N}$. +- A natural number $n$ is said to be a {{#concept "strict upper bound" Disambiguation="structured natural numbers" Agda=is-strict-upper-bound-ℕ}} if there is a +function from $P(x)$ to the type $x < n$ for all $x : \mathbb{N}$. +- A natural number $n$ is said to be a {{#concept "least upper bound" Disambiguation="structured natural numbers" Agda=is-least-upper-bound-ℕ}} if any natural number $x$ is an upper bound if and only if $n \leq x$. +- A natural number $n$ is said to be a {{#concept "least strict upper bound" Disambiguation="structured natural numbers" Agda=is-least-strict-upper-bound-ℕ}} if any natural number $x$ is a strict upper bound if and only if $n \leq x$. + +## Definitions ### Upper bounds ```agda is-upper-bound-ℕ : - {l : Level} (P : ℕ → UU l) (n : ℕ) → UU l + {l : Level} (P : ℕ → UU l) → ℕ → UU l is-upper-bound-ℕ P n = - (m : ℕ) → P m → leq-ℕ m n + (m : ℕ) → P m → m ≤-ℕ n ``` ### Strict upper bounds @@ -39,11 +56,119 @@ is-upper-bound-ℕ P n = is-strict-upper-bound-ℕ : {l : Level} (P : ℕ → UU l) (n : ℕ) → UU l is-strict-upper-bound-ℕ P n = - (m : ℕ) → P m → le-ℕ m n + (m : ℕ) → P m → m <-ℕ n +``` + +### Least upper bounds + +```agda +is-least-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-least-upper-bound-ℕ P n = + (x : ℕ) → is-upper-bound-ℕ P x ↔ n ≤-ℕ x + +is-upper-bound-is-least-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-least-upper-bound-ℕ P n → is-upper-bound-ℕ P n +is-upper-bound-is-least-upper-bound-ℕ P n H = + backward-implication (H n) (refl-leq-ℕ n) + +leq-is-least-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-least-upper-bound-ℕ P n → + (m : ℕ) → is-upper-bound-ℕ P m → n ≤-ℕ m +leq-is-least-upper-bound-ℕ P n H m = + forward-implication (H m) +``` + +### Least strict upper bounds + +```agda +is-least-strict-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-least-strict-upper-bound-ℕ P n = + (x : ℕ) → is-strict-upper-bound-ℕ P x ↔ n ≤-ℕ x ``` ## Properties +### Being an upper bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-upper-bound-ℕ : + (n : ℕ) → is-prop (is-upper-bound-ℕ P n) + is-prop-is-upper-bound-ℕ n = + is-prop-Π (λ x → is-prop-function-type (is-prop-leq-ℕ x n)) + + is-upper-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-upper-bound-ℕ-Prop n) = is-upper-bound-ℕ P n + pr2 (is-upper-bound-ℕ-Prop n) = is-prop-is-upper-bound-ℕ n +``` + +### Being a strict upper bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-strict-upper-bound-ℕ : + (n : ℕ) → is-prop (is-strict-upper-bound-ℕ P n) + is-prop-is-strict-upper-bound-ℕ n = + is-prop-Π (λ x → is-prop-function-type (is-prop-le-ℕ x n)) + + is-strict-upper-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-strict-upper-bound-ℕ-Prop n) = is-strict-upper-bound-ℕ P n + pr2 (is-strict-upper-bound-ℕ-Prop n) = is-prop-is-strict-upper-bound-ℕ n +``` + +### Being a least upper bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-least-upper-bound-ℕ : + (n : ℕ) → is-prop (is-least-upper-bound-ℕ P n) + is-prop-is-least-upper-bound-ℕ n = + is-prop-Π + ( λ x → is-prop-iff-Prop (is-upper-bound-ℕ-Prop P x) (leq-ℕ-Prop n x)) + + is-least-upper-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-least-upper-bound-ℕ-Prop n) = is-least-upper-bound-ℕ P n + pr2 (is-least-upper-bound-ℕ-Prop n) = is-prop-is-least-upper-bound-ℕ n +``` + +### Being a least strict upper bound is a property + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-prop-is-least-strict-upper-bound-ℕ : + (n : ℕ) → is-prop (is-least-strict-upper-bound-ℕ P n) + is-prop-is-least-strict-upper-bound-ℕ n = + is-prop-Π + ( λ x → + is-prop-iff-Prop (is-strict-upper-bound-ℕ-Prop P x) (leq-ℕ-Prop n x)) + + is-least-strict-upper-bound-ℕ-Prop : + (n : ℕ) → Prop l + pr1 (is-least-strict-upper-bound-ℕ-Prop n) = + is-least-strict-upper-bound-ℕ P n + pr2 (is-least-strict-upper-bound-ℕ-Prop n) = + is-prop-is-least-strict-upper-bound-ℕ n +``` + ### A strict upper bound is an upper bound ```agda @@ -53,3 +178,94 @@ is-upper-bound-is-strict-upper-bound-ℕ : is-upper-bound-is-strict-upper-bound-ℕ P n H x p = leq-le-ℕ x n (H x p) ``` + +### Any two least upper bounds are equal + +```agda +eq-is-least-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (m n : ℕ) → + is-least-upper-bound-ℕ P m → is-least-upper-bound-ℕ P n → m = n +eq-is-least-upper-bound-ℕ P m n H K = + antisymmetric-leq-ℕ m n + ( leq-is-least-upper-bound-ℕ P m H n + ( is-upper-bound-is-least-upper-bound-ℕ P n K)) + ( leq-is-least-upper-bound-ℕ P n K m + ( is-upper-bound-is-least-upper-bound-ℕ P m H)) +``` + +### If $n + 1$ is a least upper bound, then $n$ is not a least upper bound + +```agda +is-not-least-upper-bound-is-least-upper-bound-succ-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-least-upper-bound-ℕ P (succ-ℕ n) → ¬ (is-least-upper-bound-ℕ P n) +is-not-least-upper-bound-is-least-upper-bound-succ-ℕ P n H K = + has-no-fixed-points-succ-ℕ n (eq-is-least-upper-bound-ℕ P (succ-ℕ n) n H K) +``` + +### If $n + 1$ is an upper bound and $P (n + 1)$ is empty, then $n$ is an upper bound + +```agda +decrease-is-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-upper-bound-ℕ P (succ-ℕ n) → ¬ P (succ-ℕ n) → is-upper-bound-ℕ P n +decrease-is-upper-bound-ℕ P n H f m p = + map-right-unit-law-coproduct-is-empty + ( m ≤-ℕ n) + ( m = succ-ℕ n) + ( λ α → f (tr P α p)) + ( decide-leq-succ-ℕ m n (H m p)) +``` + +### Any successor natural number $n + 1$ such that $P (n + 1)$ is empty is not a least upper bound + +```agda +is-not-least-upper-bound-is-empty-succ-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + ¬ P (succ-ℕ n) → ¬ is-least-upper-bound-ℕ P (succ-ℕ n) +is-not-least-upper-bound-is-empty-succ-ℕ P n f H = + neg-succ-leq-ℕ n + ( leq-is-least-upper-bound-ℕ P + ( succ-ℕ n) + ( H) + ( n) + ( decrease-is-upper-bound-ℕ P n + ( is-upper-bound-is-least-upper-bound-ℕ P (succ-ℕ n) H) + ( f))) +``` + +### Given a structured natural number, any natural number $n$ such that $P(n)$ is empty is not least upper bound + +```agda +is-not-least-upper-bound-is-empty-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + (m : ℕ) → P m → ¬ (P n) → ¬ is-least-upper-bound-ℕ P n +is-not-least-upper-bound-is-empty-ℕ P zero-ℕ m p f H = + f ( tr P + ( is-zero-leq-zero-ℕ m (is-upper-bound-is-least-upper-bound-ℕ P 0 H m p)) + ( p)) +is-not-least-upper-bound-is-empty-ℕ P (succ-ℕ n) m p f H = + is-not-least-upper-bound-is-empty-succ-ℕ P n f H +``` + +### Given a number with an element in $P$, the type $P$ at the least upper bound is nonempty + +```agda +is-nonempty-structure-is-in-family-is-least-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + is-least-upper-bound-ℕ P n → (m : ℕ) → P m → is-nonempty (P n) +is-nonempty-structure-is-in-family-is-least-upper-bound-ℕ P n H m p f = + is-not-least-upper-bound-is-empty-ℕ P n m p f H +``` + +### Given a structured number in a decidable type family $P$, the least upper bound is structured in $P$ + +```agda +structure-least-upper-bound-is-decidable-fam-ℕ : + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → + (n : ℕ) → is-least-upper-bound-ℕ P n → (m : ℕ) → P m → P n +structure-least-upper-bound-is-decidable-fam-ℕ P d n H m p = + double-negation-elim-is-decidable + ( d n) + ( is-nonempty-structure-is-in-family-is-least-upper-bound-ℕ P n H m p) +``` diff --git a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md index fac7df36c0..3e8f6c3799 100644 --- a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md @@ -7,15 +7,18 @@ module elementary-number-theory.well-ordering-principle-natural-numbers where
Imports ```agda +open import elementary-number-theory.decidable-types open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lower-bounds-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.upper-bounds-natural-numbers open import foundation.cartesian-product-types open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.function-types open import foundation.functoriality-dependent-pair-types open import foundation.hilberts-epsilon-operators open import foundation.identity-types @@ -30,51 +33,55 @@ open import foundation.universe-levels ## Idea -The {{#concept "well-ordering principle of the natural numbers" WDID=Q2488476 WD="well-ordering principle"}} asserts that for every family -of [decidable types](foundation.decidable-types.md) over ℕ equipped with a natural number `n` and an element -`p : P n`, we can find a least natural number `n₀` with an element `p₀ : P n₀`. +The +{{#concept "well-ordering principle of the natural numbers" WDID=Q2488476 WD="well-ordering principle"}} +asserts that for every family of +[decidable types](foundation.decidable-types.md) over ℕ equipped with a natural +number `n` and an element `p : P n`, we can find a least natural number `n₀` +with an element `p₀ : P n₀`. The well-ordering principle has some useful consequences: -- For any family `P` of decidable types equipped with a natural number `n ≥ b` and an element `p : P n` we can find a least natural number `n₀ ≥ b` with an element `p₀ : P n₀`. -- For any family `P` of decidable types equipped with an element `n ≤ b` and an element `p : P n` we can find a largest natural number `n₁ ≤ b` with an element `p₁ : P n₁`. -- For any strictly inflationary map `f : ℕ → ℕ` such that `f 0 ≤ b` we can find a largest natural number `n₁` such that `f n₁ ≤ b`. -## Theorem +- For any family `P` of decidable types equipped with a natural number `n ≥ b` + and an element `p : P n` we can find a least natural number `n₀ ≥ b` with an + element `p₀ : P n₀`. +- For any family `P` of decidable types equipped with an element `n ≤ b` and an + element `p : P n` we can find a largest natural number `n₁ ≤ b` with an + element `p₁ : P n₁`. +- For any + [strictly inflationary map](order-theory.inflationary-maps-strictly-ordered-types.md) + `f : ℕ → ℕ` such that `f 0 ≤ b` we can find a largest natural number `n₁` such + that `f n₁ ≤ b`. + +## Definitions + +### Minimal elements ```agda minimal-element-ℕ : {l : Level} (P : ℕ → UU l) → UU l -minimal-element-ℕ P = Σ ℕ (λ n → (P n) × (is-lower-bound-ℕ P n)) +minimal-element-ℕ P = Σ ℕ (λ n → P n × is-lower-bound-ℕ P n) +``` -module _ - {l1 : Level} (P : ℕ → Prop l1) - where +### Bounded maximal elements - all-elements-equal-minimal-element-ℕ : - all-elements-equal (minimal-element-ℕ (λ n → type-Prop (P n))) - all-elements-equal-minimal-element-ℕ - (pair x (pair p l)) (pair y (pair q k)) = - eq-type-subtype - ( λ n → - product-Prop - ( pair _ (is-prop-type-Prop (P n))) - ( is-lower-bound-ℕ-Prop n)) - ( antisymmetric-leq-ℕ x y (l y q) (k x p)) +```agda +bounded-maximal-element-ℕ : + {l : Level} (P : ℕ → UU l) (b : ℕ) → UU l +bounded-maximal-element-ℕ P b = + Σ ℕ (λ n → n ≤-ℕ b × P n × is-upper-bound-ℕ P n) +``` - is-prop-minimal-element-ℕ : - is-prop (minimal-element-ℕ (λ n → type-Prop (P n))) - is-prop-minimal-element-ℕ = - is-prop-all-elements-equal all-elements-equal-minimal-element-ℕ +## Theorems - minimal-element-ℕ-Prop : Prop l1 - pr1 minimal-element-ℕ-Prop = minimal-element-ℕ (λ n → type-Prop (P n)) - pr2 minimal-element-ℕ-Prop = is-prop-minimal-element-ℕ +### The Well-Ordering Principle +```agda is-minimal-element-succ-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (m : ℕ) (pm : P (succ-ℕ m)) (is-lower-bound-m : is-lower-bound-ℕ (λ x → P (succ-ℕ x)) m) → - ¬ (P zero-ℕ) → is-lower-bound-ℕ P (succ-ℕ m) + ¬ (P 0) → is-lower-bound-ℕ P (succ-ℕ m) is-minimal-element-succ-ℕ P d m pm is-lower-bound-m neg-p0 zero-ℕ p0 = ex-falso (neg-p0 p0) is-minimal-element-succ-ℕ @@ -87,7 +94,7 @@ is-minimal-element-succ-ℕ well-ordering-principle-succ-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (n : ℕ) (p : P (succ-ℕ n)) → - is-decidable (P zero-ℕ) → + is-decidable (P 0) → minimal-element-ℕ (λ m → P (succ-ℕ m)) → minimal-element-ℕ P well-ordering-principle-succ-ℕ P d n p (inl p0) u = ( 0 , p0 , λ m q → leq-zero-ℕ m) @@ -97,32 +104,29 @@ well-ordering-principle-succ-ℕ P d n p (inr neg-p0) (m , pm , is-min-m) = well-ordering-principle-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → Σ ℕ P → minimal-element-ℕ P -pr1 (well-ordering-principle-ℕ P d (pair zero-ℕ p)) = zero-ℕ -pr1 (pr2 (well-ordering-principle-ℕ P d (pair zero-ℕ p))) = p -pr2 (pr2 (well-ordering-principle-ℕ P d (pair zero-ℕ p))) m q = leq-zero-ℕ m -well-ordering-principle-ℕ P d (pair (succ-ℕ n) p) = - well-ordering-principle-succ-ℕ P d n p (d zero-ℕ) - ( well-ordering-principle-ℕ - ( λ m → P (succ-ℕ m)) - ( λ m → d (succ-ℕ m)) - ( pair n p)) - -number-well-ordering-principle-ℕ : +well-ordering-principle-ℕ P d (zero-ℕ , p) = + ( 0 , p , λ m _ → leq-zero-ℕ m) +well-ordering-principle-ℕ P d (succ-ℕ n , p) = + well-ordering-principle-succ-ℕ P d n p + ( d 0) + ( well-ordering-principle-ℕ (P ∘ succ-ℕ) (d ∘ succ-ℕ) (n , p)) + +nat-well-ordering-principle-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (nP : Σ ℕ P) → ℕ -number-well-ordering-principle-ℕ P d nP = +nat-well-ordering-principle-ℕ P d nP = pr1 (well-ordering-principle-ℕ P d nP) ``` -### The well-ordering principle returns `0` if `P 0` holds +## Properties -This is independently of the input `(pair n p) : Σ ℕ P`. +### The well-ordering principle returns `0` if `P 0` holds ```agda is-zero-well-ordering-principle-succ-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) - (n : ℕ) (p : P (succ-ℕ n)) (d0 : is-decidable (P zero-ℕ)) → - (x : minimal-element-ℕ (λ m → P (succ-ℕ m))) (p0 : P zero-ℕ) → - pr1 (well-ordering-principle-succ-ℕ P d n p d0 x) = zero-ℕ + (n : ℕ) (p : P (succ-ℕ n)) (d0 : is-decidable (P 0)) → + (x : minimal-element-ℕ (λ m → P (succ-ℕ m))) (p0 : P 0) → + pr1 (well-ordering-principle-succ-ℕ P d n p d0 x) = 0 is-zero-well-ordering-principle-succ-ℕ P d n p (inl p0) x q0 = refl is-zero-well-ordering-principle-succ-ℕ P d n p (inr np0) x q0 = @@ -130,14 +134,42 @@ is-zero-well-ordering-principle-succ-ℕ P d n p (inr np0) x q0 = is-zero-well-ordering-principle-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → - (x : Σ ℕ P) → P zero-ℕ → is-zero-ℕ (number-well-ordering-principle-ℕ P d x) -is-zero-well-ordering-principle-ℕ P d (pair zero-ℕ p) p0 = refl -is-zero-well-ordering-principle-ℕ P d (pair (succ-ℕ m) p) = - is-zero-well-ordering-principle-succ-ℕ P d m p (d zero-ℕ) + (x : Σ ℕ P) → P 0 → is-zero-ℕ (nat-well-ordering-principle-ℕ P d x) +is-zero-well-ordering-principle-ℕ P d (zero-ℕ , p) p0 = refl +is-zero-well-ordering-principle-ℕ P d (succ-ℕ m , p) = + is-zero-well-ordering-principle-succ-ℕ P d m p (d 0) ( well-ordering-principle-ℕ ( λ z → P (succ-ℕ z)) ( λ x → d (succ-ℕ x)) - ( pair m p)) + ( m , p)) +``` + +### The type of minimal elements of a subtype has at most one element + +```agda +module _ + {l1 : Level} (P : ℕ → Prop l1) + where + + all-elements-equal-minimal-element-ℕ : + all-elements-equal (minimal-element-ℕ (λ n → type-Prop (P n))) + all-elements-equal-minimal-element-ℕ + (x , p , l) (y , q , k) = + eq-type-subtype + ( λ n → + product-Prop + ( _ , is-prop-type-Prop (P n)) + ( is-lower-bound-ℕ-Prop (type-Prop ∘ P) n)) + ( antisymmetric-leq-ℕ x y (l y q) (k x p)) + + is-prop-minimal-element-ℕ : + is-prop (minimal-element-ℕ (λ n → type-Prop (P n))) + is-prop-minimal-element-ℕ = + is-prop-all-elements-equal all-elements-equal-minimal-element-ℕ + + minimal-element-ℕ-Prop : Prop l1 + pr1 minimal-element-ℕ-Prop = minimal-element-ℕ (λ n → type-Prop (P n)) + pr2 minimal-element-ℕ-Prop = is-prop-minimal-element-ℕ ``` ### The ε-operator for decidable subtypes of ℕ @@ -148,8 +180,70 @@ is-zero-well-ordering-principle-ℕ P d (pair (succ-ℕ m) p) = (d : (x : ℕ) → is-decidable (type-Prop (P x))) → ε-operator-Hilbert (type-subtype P) ε-operator-decidable-subtype-ℕ {l1} P d t = - tot ( λ x → pr1) - ( apply-universal-property-trunc-Prop t - ( minimal-element-ℕ-Prop P) - ( well-ordering-principle-ℕ (λ x → type-Prop (P x)) d)) + tot + ( λ x → pr1) + ( apply-universal-property-trunc-Prop t + ( minimal-element-ℕ-Prop P) + ( well-ordering-principle-ℕ (λ x → type-Prop (P x)) d)) +``` + +### Every decidable type family over `ℕ` equipped with an element with an upper bound has a bounded maximal element + +```agda +module _ + {l : Level} (P : ℕ → UU l) (d : (x : ℕ) → is-decidable (P x)) (b : ℕ) + where + + upper-bound-bounded-decidable-family-ℕ : + ℕ → UU l + upper-bound-bounded-decidable-family-ℕ x = + (y : ℕ) → y ≤-ℕ b → P y → y ≤-ℕ x + + is-decidable-upper-bound-bounded-decidable-family-ℕ : + (x : ℕ) → is-decidable (upper-bound-bounded-decidable-family-ℕ x) + is-decidable-upper-bound-bounded-decidable-family-ℕ x = + is-decidable-bounded-Π-ℕ' + ( λ y → P y → y ≤-ℕ x) + ( λ y → is-decidable-function-type (d y) (is-decidable-leq-ℕ y x)) + ( b) + + instance-upper-bound-bounded-decidable-family-ℕ : + upper-bound-bounded-decidable-family-ℕ b + instance-upper-bound-bounded-decidable-family-ℕ y H p = H + + least-upper-bound-bounded-decidable-family-ℕ : + minimal-element-ℕ upper-bound-bounded-decidable-family-ℕ + least-upper-bound-bounded-decidable-family-ℕ = + well-ordering-principle-ℕ + ( upper-bound-bounded-decidable-family-ℕ) + ( is-decidable-upper-bound-bounded-decidable-family-ℕ) + ( b , instance-upper-bound-bounded-decidable-family-ℕ) + + nat-bounded-maximal-element-decidable-family-ℕ : + ℕ + nat-bounded-maximal-element-decidable-family-ℕ = + pr1 least-upper-bound-bounded-decidable-family-ℕ + + upper-bound-nat-bounded-maximal-element-decidable-family-ℕ : + nat-bounded-maximal-element-decidable-family-ℕ ≤-ℕ b + upper-bound-nat-bounded-maximal-element-decidable-family-ℕ = + {!!} + + structure-nat-bounded-maximal-element-decidable-family-ℕ : + P nat-bounded-maximal-element-decidable-family-ℕ + structure-nat-bounded-maximal-element-decidable-family-ℕ = + {!!} + + is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ : + is-upper-bound-ℕ P nat-bounded-maximal-element-decidable-family-ℕ + is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ = + {!!} + + bounded-maximal-element-decidable-family-ℕ : + (n : ℕ) → n ≤-ℕ b → P n → bounded-maximal-element-ℕ P b + bounded-maximal-element-decidable-family-ℕ n H p = + ( nat-bounded-maximal-element-decidable-family-ℕ , + upper-bound-nat-bounded-maximal-element-decidable-family-ℕ , + structure-nat-bounded-maximal-element-decidable-family-ℕ , + is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ) ``` From 03dd0cf8bd2e36028aaa4e169fe69ee8e1f29941 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 10 Jan 2025 12:42:12 -0500 Subject: [PATCH 43/72] factoring out minimal and maximal structure natural numbers --- .../lower-bounds-natural-numbers.lagda.md | 13 +- ...aximal-structured-natural-numbers.lagda.md | 116 ++++++++++++++++++ ...inimal-structured-natural-numbers.lagda.md | 62 ++++++++++ .../upper-bounds-natural-numbers.lagda.md | 13 +- ...rdering-principle-natural-numbers.lagda.md | 29 +---- 5 files changed, 208 insertions(+), 25 deletions(-) create mode 100644 src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md diff --git a/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md index 1603e4d0f5..82a209643b 100644 --- a/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/lower-bounds-natural-numbers.lagda.md @@ -1,4 +1,4 @@ -# Lower bounds of type families over the natural numbers +# Lower bounds for structured natural numbers ```agda module elementary-number-theory.lower-bounds-natural-numbers where @@ -256,3 +256,14 @@ structure-largest-lower-bound-is-decidable-fam-ℕ P d n H = ( d n) ( is-nonempty-structure-is-in-family-is-largest-lower-bound-ℕ P n H) ``` + +### Any lower bound equipped with structure is a largest lower bound + +```agda +is-largest-lower-bound-is-lower-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + P n → is-lower-bound-ℕ P n → is-largest-lower-bound-ℕ P n +pr1 (is-largest-lower-bound-is-lower-bound-ℕ P n p H m) K = K n p +pr2 (is-largest-lower-bound-is-lower-bound-ℕ P n p H m) K x q = + transitive-leq-ℕ m n x (H x q) K +``` diff --git a/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md b/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md new file mode 100644 index 0000000000..bb44326ee4 --- /dev/null +++ b/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md @@ -0,0 +1,116 @@ +# Maximal structured natural numbers + +```agda +module elementary-number-theory.maximal-structured-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.upper-bounds-natural-numbers + +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.universe-levels +``` + +
+ +Consider a type family $P$ over $\mathbb{N}$. A {{#concept "maximal structured natural number" Agda=maximal-element-ℕ}} in $P$ is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is an [upper bound](elementary-number-theory.upper-bounds-natural-numbers.md) for $P$. + +Forthermore, a {{#concept "bounded maximal structured natural number" Agda=bounded-maximal-element-ℕ}} in $P$ with bound $b$ is a natural number $n \leq b$ equipped with an element $p : P(n)$ such that $n$ is an upper bound of for the type family $x \mapsto (x \leq b) × P(x)$. + +## Definitions + +### The predicate of being a maximal structured natural number + +```agda +is-maximal-element-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-maximal-element-ℕ P n = + P n × is-upper-bound-ℕ P n +``` + +### Maximal elements + +```agda +maximal-element-ℕ : + {l : Level} (P : ℕ → UU l) → UU l +maximal-element-ℕ P = + Σ ℕ (is-maximal-element-ℕ P) + +module _ + {l : Level} (P : ℕ → UU l) (n : maximal-element-ℕ P) + where + + nat-maximal-element-ℕ : ℕ + nat-maximal-element-ℕ = pr1 n + + structure-maximal-element-ℕ : P nat-maximal-element-ℕ + structure-maximal-element-ℕ = pr1 (pr2 n) + + is-upper-bound-maximal-element-ℕ : is-upper-bound-ℕ P nat-maximal-element-ℕ + is-upper-bound-maximal-element-ℕ = pr2 (pr2 n) + + is-least-upper-bound-maximal-element-ℕ : + is-least-upper-bound-ℕ P nat-maximal-element-ℕ + is-least-upper-bound-maximal-element-ℕ = + is-least-upper-bound-is-upper-bound-ℕ P + ( nat-maximal-element-ℕ) + ( structure-maximal-element-ℕ) + ( is-upper-bound-maximal-element-ℕ) +``` + +### The predicate of being a maximal bounded structured natural number + +```agda +bounded-family-family-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → ℕ → UU l +bounded-family-family-ℕ P b n = n ≤-ℕ b × P n + +is-maximal-bounded-element-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → ℕ → UU l +is-maximal-bounded-element-ℕ P b = + is-maximal-element-ℕ (bounded-family-family-ℕ P b) +``` + +### Bounded maximal elements + +```agda +bounded-maximal-element-ℕ : + {l : Level} (P : ℕ → UU l) (b : ℕ) → UU l +bounded-maximal-element-ℕ P b = + maximal-element-ℕ (bounded-family-family-ℕ P b) + +module _ + {l : Level} (P : ℕ → UU l) (b : ℕ) (n : bounded-maximal-element-ℕ P b) + where + + nat-bounded-maximal-element-ℕ : ℕ + nat-bounded-maximal-element-ℕ = pr1 n + + upper-bound-bounded-maximal-element-ℕ : nat-bounded-maximal-element-ℕ ≤-ℕ b + upper-bound-bounded-maximal-element-ℕ = pr1 (pr1 (pr2 n)) + + structure-bounded-maximal-element-ℕ : P nat-bounded-maximal-element-ℕ + structure-bounded-maximal-element-ℕ = pr2 (pr1 (pr2 n)) + + is-upper-bound-bounded-maximal-element-ℕ : + is-upper-bound-ℕ (bounded-family-family-ℕ P b) nat-bounded-maximal-element-ℕ + is-upper-bound-bounded-maximal-element-ℕ = pr2 (pr2 n) + + is-least-upper-bound-bounded-maximal-element-ℕ : + is-least-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( nat-bounded-maximal-element-ℕ) + is-least-upper-bound-bounded-maximal-element-ℕ = + is-least-upper-bound-is-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( nat-bounded-maximal-element-ℕ) + ( upper-bound-bounded-maximal-element-ℕ , + structure-bounded-maximal-element-ℕ) + ( is-upper-bound-bounded-maximal-element-ℕ) +``` + diff --git a/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md new file mode 100644 index 0000000000..2577050263 --- /dev/null +++ b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md @@ -0,0 +1,62 @@ +# Minimal structured natural numbers + +```agda +module elementary-number-theory.minimal-structured-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.natural-numbers + +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.universe-levels +``` + +
+ +## Idea + +Consider a type family $P$ over $\mathbb{N}$. A {{#concept "minimal structured natural number" Agda=minimal-element-ℕ}} in $P$ is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is a [lower bound](elementary-number-theory.lower-bounds-natural-numbers.md) for $P$. + +## Definitions + +### The predicate of being a minimal structured natural number + +```agda +is-minimal-element-ℕ : + {l : Level} (P : ℕ → UU l) → ℕ → UU l +is-minimal-element-ℕ P n = P n × is-lower-bound-ℕ P n +``` + +### Minimal elements + +```agda +minimal-element-ℕ : + {l : Level} (P : ℕ → UU l) → UU l +minimal-element-ℕ P = Σ ℕ (is-minimal-element-ℕ P) + +module _ + {l : Level} (P : ℕ → UU l) (n : minimal-element-ℕ P) + where + + nat-minimal-element-ℕ : ℕ + nat-minimal-element-ℕ = pr1 n + + structure-minimal-element-ℕ : P nat-minimal-element-ℕ + structure-minimal-element-ℕ = pr1 (pr2 n) + + is-lower-bound-minimal-element-ℕ : is-lower-bound-ℕ P nat-minimal-element-ℕ + is-lower-bound-minimal-element-ℕ = pr2 (pr2 n) + + is-largest-lower-bound-minimal-element-ℕ : + is-largest-lower-bound-ℕ P nat-minimal-element-ℕ + is-largest-lower-bound-minimal-element-ℕ = + is-largest-lower-bound-is-lower-bound-ℕ P + ( nat-minimal-element-ℕ) + ( structure-minimal-element-ℕ) + ( is-lower-bound-minimal-element-ℕ) +``` + diff --git a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md index 1c0694989b..5b4128ecc8 100644 --- a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md @@ -1,4 +1,4 @@ -# Bounds for structured natural numbers +# Upper bounds for structured natural numbers ```agda module elementary-number-theory.upper-bounds-natural-numbers where @@ -269,3 +269,14 @@ structure-least-upper-bound-is-decidable-fam-ℕ P d n H m p = ( d n) ( is-nonempty-structure-is-in-family-is-least-upper-bound-ℕ P n H m p) ``` + +### Any upper bound equipped with structure is a least upper bound + +```agda +is-least-upper-bound-is-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (n : ℕ) → + P n → is-upper-bound-ℕ P n → is-least-upper-bound-ℕ P n +pr1 (is-least-upper-bound-is-upper-bound-ℕ P n p H m) K = K n p +pr2 (is-least-upper-bound-is-upper-bound-ℕ P n p H m) K x q = + transitive-leq-ℕ x n m K (H x q) +``` diff --git a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md index 3e8f6c3799..2d9dc68325 100644 --- a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md @@ -10,6 +10,8 @@ module elementary-number-theory.well-ordering-principle-natural-numbers where open import elementary-number-theory.decidable-types open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.maximal-structured-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.upper-bounds-natural-numbers @@ -53,25 +55,6 @@ The well-ordering principle has some useful consequences: `f : ℕ → ℕ` such that `f 0 ≤ b` we can find a largest natural number `n₁` such that `f n₁ ≤ b`. -## Definitions - -### Minimal elements - -```agda -minimal-element-ℕ : - {l : Level} (P : ℕ → UU l) → UU l -minimal-element-ℕ P = Σ ℕ (λ n → P n × is-lower-bound-ℕ P n) -``` - -### Bounded maximal elements - -```agda -bounded-maximal-element-ℕ : - {l : Level} (P : ℕ → UU l) (b : ℕ) → UU l -bounded-maximal-element-ℕ P b = - Σ ℕ (λ n → n ≤-ℕ b × P n × is-upper-bound-ℕ P n) -``` - ## Theorems ### The Well-Ordering Principle @@ -242,8 +225,8 @@ module _ bounded-maximal-element-decidable-family-ℕ : (n : ℕ) → n ≤-ℕ b → P n → bounded-maximal-element-ℕ P b bounded-maximal-element-decidable-family-ℕ n H p = - ( nat-bounded-maximal-element-decidable-family-ℕ , - upper-bound-nat-bounded-maximal-element-decidable-family-ℕ , - structure-nat-bounded-maximal-element-decidable-family-ℕ , - is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ) + ( ( nat-bounded-maximal-element-decidable-family-ℕ) , + ( upper-bound-nat-bounded-maximal-element-decidable-family-ℕ , + structure-nat-bounded-maximal-element-decidable-family-ℕ) , + {!!} {- is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ -}) ``` From 948d7c7a669cc4b6b2782b48c83bb3ec7bf79b83 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 10 Jan 2025 18:44:02 -0500 Subject: [PATCH 44/72] bounded maximal elements of decidable families --- ...aximal-structured-natural-numbers.lagda.md | 55 ++++ ...inimal-structured-natural-numbers.lagda.md | 82 ++++++ .../upper-bounds-natural-numbers.lagda.md | 38 +++ ...rdering-principle-natural-numbers.lagda.md | 277 +++++++++--------- 4 files changed, 314 insertions(+), 138 deletions(-) diff --git a/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md b/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md index bb44326ee4..7e426f5226 100644 --- a/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md +++ b/src/elementary-number-theory/maximal-structured-natural-numbers.lagda.md @@ -8,11 +8,15 @@ module elementary-number-theory.maximal-structured-natural-numbers where ```agda open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.lower-bounds-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.upper-bounds-natural-numbers open import foundation.cartesian-product-types open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.propositions +open import foundation.subtypes open import foundation.universe-levels ``` @@ -114,3 +118,54 @@ module _ ( is-upper-bound-bounded-maximal-element-ℕ) ``` +## Properties + +### The type of maximal elements of a subtype has at most one element + +```agda +module _ + {l1 : Level} (P : ℕ → Prop l1) + where + + all-elements-equal-maximal-element-ℕ : + all-elements-equal (maximal-element-ℕ (λ n → type-Prop (P n))) + all-elements-equal-maximal-element-ℕ + (x , p , l) (y , q , k) = + eq-type-subtype + ( λ n → + product-Prop + ( _ , is-prop-type-Prop (P n)) + ( is-upper-bound-ℕ-Prop (type-Prop ∘ P) n)) + ( antisymmetric-leq-ℕ x y (k x p) (l y q)) + + is-prop-maximal-element-ℕ : + is-prop (maximal-element-ℕ (λ n → type-Prop (P n))) + is-prop-maximal-element-ℕ = + is-prop-all-elements-equal all-elements-equal-maximal-element-ℕ + + maximal-element-ℕ-Prop : Prop l1 + pr1 maximal-element-ℕ-Prop = maximal-element-ℕ (λ n → type-Prop (P n)) + pr2 maximal-element-ℕ-Prop = is-prop-maximal-element-ℕ +``` + +### A natural number is a largest lower bound if and only if it is a maximal element of the type of lower bounds + +```agda +module _ + {l : Level} (P : ℕ → UU l) (n : ℕ) + where + + is-maximal-lower-bound-is-largest-lower-bound-ℕ : + is-largest-lower-bound-ℕ P n → is-maximal-element-ℕ (is-lower-bound-ℕ P) n + pr1 (is-maximal-lower-bound-is-largest-lower-bound-ℕ H) = + is-lower-bound-is-largest-lower-bound-ℕ P n H + pr2 (is-maximal-lower-bound-is-largest-lower-bound-ℕ H) m K = + leq-is-largest-lower-bound-ℕ P n H m K + + is-largest-lower-bound-is-maximal-lower-bound-ℕ : + is-maximal-element-ℕ (is-lower-bound-ℕ P) n → is-largest-lower-bound-ℕ P n + pr1 (is-largest-lower-bound-is-maximal-lower-bound-ℕ H m) K = + pr2 H m K + pr2 (is-largest-lower-bound-is-maximal-lower-bound-ℕ H m) K x p = + transitive-leq-ℕ m n x (pr1 H x p) K +``` diff --git a/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md index 2577050263..a08f2c74fa 100644 --- a/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md +++ b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md @@ -7,11 +7,19 @@ module elementary-number-theory.minimal-structured-natural-numbers where
Imports ```agda +open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lower-bounds-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.upper-bounds-natural-numbers open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.function-types +open import foundation.propositions +open import foundation.subtypes open import foundation.universe-levels ``` @@ -60,3 +68,77 @@ module _ ( is-lower-bound-minimal-element-ℕ) ``` +## Properties + +### The type of minimal elements of a subtype has at most one element + +```agda +module _ + {l1 : Level} (P : ℕ → Prop l1) + where + + all-elements-equal-minimal-element-ℕ : + all-elements-equal (minimal-element-ℕ (λ n → type-Prop (P n))) + all-elements-equal-minimal-element-ℕ + (x , p , l) (y , q , k) = + eq-type-subtype + ( λ n → + product-Prop + ( _ , is-prop-type-Prop (P n)) + ( is-lower-bound-ℕ-Prop (type-Prop ∘ P) n)) + ( antisymmetric-leq-ℕ x y (l y q) (k x p)) + + is-prop-minimal-element-ℕ : + is-prop (minimal-element-ℕ (λ n → type-Prop (P n))) + is-prop-minimal-element-ℕ = + is-prop-all-elements-equal all-elements-equal-minimal-element-ℕ + + minimal-element-ℕ-Prop : Prop l1 + pr1 minimal-element-ℕ-Prop = minimal-element-ℕ (λ n → type-Prop (P n)) + pr2 minimal-element-ℕ-Prop = is-prop-minimal-element-ℕ +``` + +### Shifting minimal elements + +If $P(0)$ is a decidable type, then we can obtain a minimal structured element of $P$ from a minimal structured element of $n\mapsto P(n+1)$. + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + shift-minimal-element-ℕ : + is-decidable (P 0) → minimal-element-ℕ (P ∘ succ-ℕ) → minimal-element-ℕ P + shift-minimal-element-ℕ (inl p) m = + ( 0 , p , λ x _ → leq-zero-ℕ x) + pr1 (shift-minimal-element-ℕ (inr f) m) = + succ-ℕ (nat-minimal-element-ℕ (P ∘ succ-ℕ) m) + pr1 (pr2 (shift-minimal-element-ℕ (inr f) m)) = + structure-minimal-element-ℕ (P ∘ succ-ℕ) m + pr2 (pr2 (shift-minimal-element-ℕ (inr f) m)) zero-ℕ p = + ex-falso (f p) + pr2 (pr2 (shift-minimal-element-ℕ (inr f) m)) (succ-ℕ x) p = + is-lower-bound-minimal-element-ℕ (P ∘ succ-ℕ) m x p +``` + +### A natural number is a least upper bound if and only if it is a minimal element of the type of upper bounds + +```agda +module _ + {l : Level} (P : ℕ → UU l) (n : ℕ) + where + + is-minimal-upper-bound-is-least-upper-bound-ℕ : + is-least-upper-bound-ℕ P n → is-minimal-element-ℕ (is-upper-bound-ℕ P) n + pr1 (is-minimal-upper-bound-is-least-upper-bound-ℕ H) = + is-upper-bound-is-least-upper-bound-ℕ P n H + pr2 (is-minimal-upper-bound-is-least-upper-bound-ℕ H) m K = + leq-is-least-upper-bound-ℕ P n H m K + + is-least-upper-bound-is-minimal-upper-bound-ℕ : + is-minimal-element-ℕ (is-upper-bound-ℕ P) n → is-least-upper-bound-ℕ P n + pr1 (is-least-upper-bound-is-minimal-upper-bound-ℕ H m) K = + pr2 H m K + pr2 (is-least-upper-bound-is-minimal-upper-bound-ℕ H m) K x p = + transitive-leq-ℕ x n m K (pr1 H x p) +``` diff --git a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md index 5b4128ecc8..677befd36a 100644 --- a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md @@ -11,6 +11,7 @@ open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.double-negation @@ -280,3 +281,40 @@ pr1 (is-least-upper-bound-is-upper-bound-ℕ P n p H m) K = K n p pr2 (is-least-upper-bound-is-upper-bound-ℕ P n p H m) K x q = transitive-leq-ℕ x n m K (H x q) ``` + +### Any element greater than an upper bound is an upper bound + +```agda +is-upper-bound-leq-is-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) → + (b : ℕ) → is-upper-bound-ℕ P b → + (n : ℕ) → b ≤-ℕ n → is-upper-bound-ℕ P n +is-upper-bound-leq-is-upper-bound-ℕ P b H n K x p = + transitive-leq-ℕ x b n K (H x p) +``` + +### Being an upper bound of a decidable type family is decidable, given an upper bound + +```agda +is-decidable-is-upper-bound-ℕ' : + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → + (b : ℕ) → is-upper-bound-ℕ P b → is-decidable (P b) → + (n : ℕ) → is-decidable (is-upper-bound-ℕ P n) +is-decidable-is-upper-bound-ℕ' P d zero-ℕ H e n = + inl (is-upper-bound-leq-is-upper-bound-ℕ P 0 H n (leq-zero-ℕ n)) +is-decidable-is-upper-bound-ℕ' P d (succ-ℕ b) H (inl p) n = + is-decidable-iff' + ( inv-iff (is-least-upper-bound-is-upper-bound-ℕ P (succ-ℕ b) p H n)) + ( is-decidable-leq-ℕ (succ-ℕ b) n) +is-decidable-is-upper-bound-ℕ' P d (succ-ℕ b) H (inr f) = + is-decidable-is-upper-bound-ℕ' P d b + ( decrease-is-upper-bound-ℕ P b H f) + ( d b) + +is-decidable-is-upper-bound-ℕ : + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → + (b : ℕ) → is-upper-bound-ℕ P b → + (n : ℕ) → is-decidable (is-upper-bound-ℕ P n) +is-decidable-is-upper-bound-ℕ P d b H = + is-decidable-is-upper-bound-ℕ' P d b H (d b) +``` diff --git a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md index 2d9dc68325..87c948ddba 100644 --- a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md @@ -60,44 +60,46 @@ The well-ordering principle has some useful consequences: ### The Well-Ordering Principle ```agda -is-minimal-element-succ-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) - (m : ℕ) (pm : P (succ-ℕ m)) - (is-lower-bound-m : is-lower-bound-ℕ (λ x → P (succ-ℕ x)) m) → - ¬ (P 0) → is-lower-bound-ℕ P (succ-ℕ m) -is-minimal-element-succ-ℕ P d m pm is-lower-bound-m neg-p0 zero-ℕ p0 = - ex-falso (neg-p0 p0) -is-minimal-element-succ-ℕ - P d zero-ℕ pm is-lower-bound-m neg-p0 (succ-ℕ n) psuccn = - leq-zero-ℕ n -is-minimal-element-succ-ℕ - P d (succ-ℕ m) pm is-lower-bound-m neg-p0 (succ-ℕ n) psuccn = - is-lower-bound-m n psuccn - -well-ordering-principle-succ-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) - (n : ℕ) (p : P (succ-ℕ n)) → - is-decidable (P 0) → - minimal-element-ℕ (λ m → P (succ-ℕ m)) → minimal-element-ℕ P -well-ordering-principle-succ-ℕ P d n p (inl p0) u = - ( 0 , p0 , λ m q → leq-zero-ℕ m) -well-ordering-principle-succ-ℕ P d n p (inr neg-p0) (m , pm , is-min-m) = - ( succ-ℕ m , pm , is-minimal-element-succ-ℕ P d m pm is-min-m neg-p0) - -well-ordering-principle-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → - Σ ℕ P → minimal-element-ℕ P -well-ordering-principle-ℕ P d (zero-ℕ , p) = - ( 0 , p , λ m _ → leq-zero-ℕ m) -well-ordering-principle-ℕ P d (succ-ℕ n , p) = - well-ordering-principle-succ-ℕ P d n p - ( d 0) - ( well-ordering-principle-ℕ (P ∘ succ-ℕ) (d ∘ succ-ℕ) (n , p)) - -nat-well-ordering-principle-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) (nP : Σ ℕ P) → ℕ -nat-well-ordering-principle-ℕ P d nP = - pr1 (well-ordering-principle-ℕ P d nP) +well-ordering-principle-ℕ' : + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → is-decidable (P 0) → + (n : ℕ) → P n → minimal-element-ℕ P +well-ordering-principle-ℕ' P d (inl p0) n p = + ( 0 , p0 , λ m _ → leq-zero-ℕ m) +well-ordering-principle-ℕ' P d (inr f) zero-ℕ p = + ex-falso (f p) +well-ordering-principle-ℕ' P d (inr f) (succ-ℕ n) p = + shift-minimal-element-ℕ P + ( inr f) + ( well-ordering-principle-ℕ' (P ∘ succ-ℕ) (d ∘ succ-ℕ) (d 1) n p) + +module _ + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) {n : ℕ} (p : P n) + where + + well-ordering-principle-ℕ : + minimal-element-ℕ P + well-ordering-principle-ℕ = + well-ordering-principle-ℕ' P d (d 0) n p + + nat-well-ordering-principle-ℕ : + ℕ + nat-well-ordering-principle-ℕ = + nat-minimal-element-ℕ P well-ordering-principle-ℕ + + structure-nat-well-ordering-principle-ℕ : + P nat-well-ordering-principle-ℕ + structure-nat-well-ordering-principle-ℕ = + structure-minimal-element-ℕ P well-ordering-principle-ℕ + + is-lower-bound-nat-well-ordering-principle-ℕ : + is-lower-bound-ℕ P nat-well-ordering-principle-ℕ + is-lower-bound-nat-well-ordering-principle-ℕ = + is-lower-bound-minimal-element-ℕ P well-ordering-principle-ℕ + + is-largest-lower-bound-nat-well-ordering-principle-ℕ : + is-largest-lower-bound-ℕ P nat-well-ordering-principle-ℕ + is-largest-lower-bound-nat-well-ordering-principle-ℕ = + is-largest-lower-bound-minimal-element-ℕ P well-ordering-principle-ℕ ``` ## Properties @@ -105,54 +107,114 @@ nat-well-ordering-principle-ℕ P d nP = ### The well-ordering principle returns `0` if `P 0` holds ```agda -is-zero-well-ordering-principle-succ-ℕ : - {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) - (n : ℕ) (p : P (succ-ℕ n)) (d0 : is-decidable (P 0)) → - (x : minimal-element-ℕ (λ m → P (succ-ℕ m))) (p0 : P 0) → - pr1 (well-ordering-principle-succ-ℕ P d n p d0 x) = 0 -is-zero-well-ordering-principle-succ-ℕ P d n p (inl p0) x q0 = - refl -is-zero-well-ordering-principle-succ-ℕ P d n p (inr np0) x q0 = - ex-falso (np0 q0) +is-zero-well-ordering-principle-ℕ' : + {l : Level} (P : ℕ → UU l) + (d : is-decidable-fam P) (d0 : is-decidable (P 0)) → + (n : ℕ) (p : P n) → P 0 → + is-zero-ℕ (nat-minimal-element-ℕ P (well-ordering-principle-ℕ' P d d0 n p)) +is-zero-well-ordering-principle-ℕ' P d (inl x) n p p0 = refl +is-zero-well-ordering-principle-ℕ' P d (inr f) n p p0 = ex-falso (f p0) is-zero-well-ordering-principle-ℕ : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → - (x : Σ ℕ P) → P 0 → is-zero-ℕ (nat-well-ordering-principle-ℕ P d x) -is-zero-well-ordering-principle-ℕ P d (zero-ℕ , p) p0 = refl -is-zero-well-ordering-principle-ℕ P d (succ-ℕ m , p) = - is-zero-well-ordering-principle-succ-ℕ P d m p (d 0) - ( well-ordering-principle-ℕ - ( λ z → P (succ-ℕ z)) - ( λ x → d (succ-ℕ x)) - ( m , p)) + {n : ℕ} (p : P n) → P 0 → is-zero-ℕ (nat-well-ordering-principle-ℕ P d p) +is-zero-well-ordering-principle-ℕ P d p p0 = + is-zero-well-ordering-principle-ℕ' P d (d 0) _ p p0 ``` -### The type of minimal elements of a subtype has at most one element +### Every decidable type family over `ℕ` equipped with an instance of an element with an upper bound has a bounded maximal element ```agda module _ - {l1 : Level} (P : ℕ → Prop l1) + {l : Level} (P : ℕ → UU l) (d : (x : ℕ) → is-decidable (P x)) (b : ℕ) where - all-elements-equal-minimal-element-ℕ : - all-elements-equal (minimal-element-ℕ (λ n → type-Prop (P n))) - all-elements-equal-minimal-element-ℕ - (x , p , l) (y , q , k) = - eq-type-subtype - ( λ n → - product-Prop - ( _ , is-prop-type-Prop (P n)) - ( is-lower-bound-ℕ-Prop (type-Prop ∘ P) n)) - ( antisymmetric-leq-ℕ x y (l y q) (k x p)) - - is-prop-minimal-element-ℕ : - is-prop (minimal-element-ℕ (λ n → type-Prop (P n))) - is-prop-minimal-element-ℕ = - is-prop-all-elements-equal all-elements-equal-minimal-element-ℕ - - minimal-element-ℕ-Prop : Prop l1 - pr1 minimal-element-ℕ-Prop = minimal-element-ℕ (λ n → type-Prop (P n)) - pr2 minimal-element-ℕ-Prop = is-prop-minimal-element-ℕ + minimal-upper-bound-bounded-family-ℕ : + minimal-element-ℕ (is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + minimal-upper-bound-bounded-family-ℕ = + well-ordering-principle-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( is-decidable-is-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( λ x → is-decidable-product (is-decidable-leq-ℕ x b) (d x)) + ( b) + ( λ x → pr1)) + ( λ x → pr1) + + nat-bounded-maximal-element-instance-ℕ : + ℕ + nat-bounded-maximal-element-instance-ℕ = + nat-minimal-element-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( minimal-upper-bound-bounded-family-ℕ) + + is-upper-bound-bounded-maximal-element-instance-ℕ : + is-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( nat-bounded-maximal-element-instance-ℕ) + is-upper-bound-bounded-maximal-element-instance-ℕ = + structure-minimal-element-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( minimal-upper-bound-bounded-family-ℕ) + + is-largest-lower-bound-minimal-upper-bound-bounded-family-ℕ : + is-largest-lower-bound-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( nat-bounded-maximal-element-instance-ℕ) + is-largest-lower-bound-minimal-upper-bound-bounded-family-ℕ = + is-largest-lower-bound-minimal-element-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( minimal-upper-bound-bounded-family-ℕ) + + upper-bound-nat-bounded-maximal-element-instance-ℕ : + nat-bounded-maximal-element-instance-ℕ ≤-ℕ b + upper-bound-nat-bounded-maximal-element-instance-ℕ = + is-lower-bound-is-largest-lower-bound-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( nat-bounded-maximal-element-instance-ℕ) + ( is-largest-lower-bound-minimal-upper-bound-bounded-family-ℕ) + ( b) + ( λ x → pr1) + + is-least-upper-bound-nat-bounded-maximal-element-instance-ℕ : + is-least-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( nat-bounded-maximal-element-instance-ℕ) + pr1 (is-least-upper-bound-nat-bounded-maximal-element-instance-ℕ x) = + is-lower-bound-is-largest-lower-bound-ℕ + ( is-upper-bound-ℕ (bounded-family-family-ℕ P b)) + ( nat-bounded-maximal-element-instance-ℕ) + ( is-largest-lower-bound-minimal-upper-bound-bounded-family-ℕ) + ( x) + pr2 (is-least-upper-bound-nat-bounded-maximal-element-instance-ℕ x) = + is-upper-bound-leq-is-upper-bound-ℕ + ( bounded-family-family-ℕ P b) + ( nat-bounded-maximal-element-instance-ℕ) + ( is-upper-bound-bounded-maximal-element-instance-ℕ) + ( x) + + structure-bounded-maximal-element-instance-ℕ : + (m : ℕ) → m ≤-ℕ b → P m → + P nat-bounded-maximal-element-instance-ℕ + structure-bounded-maximal-element-instance-ℕ m H p = + pr2 + ( structure-least-upper-bound-is-decidable-fam-ℕ + ( bounded-family-family-ℕ P b) + ( λ x → is-decidable-product (is-decidable-leq-ℕ x b) (d x)) + ( nat-bounded-maximal-element-instance-ℕ) + ( is-least-upper-bound-nat-bounded-maximal-element-instance-ℕ) + ( m) + ( H , p)) + + bounded-maximal-element-instance-ℕ : + (m : ℕ) → m ≤-ℕ b → P m → bounded-maximal-element-ℕ P b + pr1 (bounded-maximal-element-instance-ℕ m H p) = + nat-bounded-maximal-element-instance-ℕ + pr1 (pr2 (bounded-maximal-element-instance-ℕ m H p)) = + ( upper-bound-nat-bounded-maximal-element-instance-ℕ , + structure-bounded-maximal-element-instance-ℕ m H p) + pr2 (pr2 (bounded-maximal-element-instance-ℕ m H p)) = + is-upper-bound-bounded-maximal-element-instance-ℕ ``` ### The ε-operator for decidable subtypes of ℕ @@ -167,66 +229,5 @@ module _ ( λ x → pr1) ( apply-universal-property-trunc-Prop t ( minimal-element-ℕ-Prop P) - ( well-ordering-principle-ℕ (λ x → type-Prop (P x)) d)) -``` - -### Every decidable type family over `ℕ` equipped with an element with an upper bound has a bounded maximal element - -```agda -module _ - {l : Level} (P : ℕ → UU l) (d : (x : ℕ) → is-decidable (P x)) (b : ℕ) - where - - upper-bound-bounded-decidable-family-ℕ : - ℕ → UU l - upper-bound-bounded-decidable-family-ℕ x = - (y : ℕ) → y ≤-ℕ b → P y → y ≤-ℕ x - - is-decidable-upper-bound-bounded-decidable-family-ℕ : - (x : ℕ) → is-decidable (upper-bound-bounded-decidable-family-ℕ x) - is-decidable-upper-bound-bounded-decidable-family-ℕ x = - is-decidable-bounded-Π-ℕ' - ( λ y → P y → y ≤-ℕ x) - ( λ y → is-decidable-function-type (d y) (is-decidable-leq-ℕ y x)) - ( b) - - instance-upper-bound-bounded-decidable-family-ℕ : - upper-bound-bounded-decidable-family-ℕ b - instance-upper-bound-bounded-decidable-family-ℕ y H p = H - - least-upper-bound-bounded-decidable-family-ℕ : - minimal-element-ℕ upper-bound-bounded-decidable-family-ℕ - least-upper-bound-bounded-decidable-family-ℕ = - well-ordering-principle-ℕ - ( upper-bound-bounded-decidable-family-ℕ) - ( is-decidable-upper-bound-bounded-decidable-family-ℕ) - ( b , instance-upper-bound-bounded-decidable-family-ℕ) - - nat-bounded-maximal-element-decidable-family-ℕ : - ℕ - nat-bounded-maximal-element-decidable-family-ℕ = - pr1 least-upper-bound-bounded-decidable-family-ℕ - - upper-bound-nat-bounded-maximal-element-decidable-family-ℕ : - nat-bounded-maximal-element-decidable-family-ℕ ≤-ℕ b - upper-bound-nat-bounded-maximal-element-decidable-family-ℕ = - {!!} - - structure-nat-bounded-maximal-element-decidable-family-ℕ : - P nat-bounded-maximal-element-decidable-family-ℕ - structure-nat-bounded-maximal-element-decidable-family-ℕ = - {!!} - - is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ : - is-upper-bound-ℕ P nat-bounded-maximal-element-decidable-family-ℕ - is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ = - {!!} - - bounded-maximal-element-decidable-family-ℕ : - (n : ℕ) → n ≤-ℕ b → P n → bounded-maximal-element-ℕ P b - bounded-maximal-element-decidable-family-ℕ n H p = - ( ( nat-bounded-maximal-element-decidable-family-ℕ) , - ( upper-bound-nat-bounded-maximal-element-decidable-family-ℕ , - structure-nat-bounded-maximal-element-decidable-family-ℕ) , - {!!} {- is-upper-bound-nat-bounded-maximal-element-decidable-family-ℕ -}) + ( λ (n , p) → well-ordering-principle-ℕ (type-Prop ∘ P) d p)) ``` From 08455851b0dcf39770a1b90c45a1477a83096398 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 11 Jan 2025 11:06:17 -0500 Subject: [PATCH 45/72] setting up some files for maps on the natural numbers --- .../decidable-maps-natural-numbers.lagda.md | 17 ++++ .../finite-maps-natural-numbers.lagda.md | 19 ++++ ...tionary-functions-natural-numbers.lagda.md | 81 ++++++++++++++++ ...strict-inequality-natural-numbers.lagda.md | 28 ++++-- ...tionary-functions-natural-numbers.lagda.md | 93 +++++++++++++++++++ .../standard-finite-types.lagda.md | 2 +- 6 files changed, 231 insertions(+), 9 deletions(-) create mode 100644 src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/finite-maps-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/inflationary-functions-natural-numbers.lagda.md create mode 100644 src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md diff --git a/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md b/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md new file mode 100644 index 0000000000..1c0513909c --- /dev/null +++ b/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md @@ -0,0 +1,17 @@ +# Decidable maps into the natural numbers + +```agda +module elementary-number-theory.decidable-maps-natural-numbers where +``` + +
Imports + +```agda + +``` + +
+ +## Idea + +Consider a map $f : A \to \mathbb{N}$ into the [natural numbers](elementary-number-theory.natural-numbers.md). Then $f$ is said to be a {{#concept "decidable map" Disambiguation="natural numbers"}} if its [fibers](foundation-core.fibers-of-maps.md) are [decidable](foundation.decidable-types.md), i.e., if it is a [decidable map](foundation.decidable-maps.md). diff --git a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md new file mode 100644 index 0000000000..a024fc0348 --- /dev/null +++ b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md @@ -0,0 +1,19 @@ +# Finite maps into the natural numbers + +```agda +module elementary-number-theory.finite-maps-natural-numbers where +``` + +
Imports + +```agda + +``` + +
+ +## Idea + +A map $f : A \to \mathbb{N}$ is said to be a {{#concept "finite map" Disambiguation="natural numbers"}} if its [fibers](foundation-core.finite-types.md) are [finite](univalent-combinatorics.finite-types.md). + +Finite maps are [decidable](elementary-number-theory.decidable-maps-natural-numbers.md). Every finite map $f : \mathbb{N} \to \mathbb{N}$ has a definite lowest value, and for every finite map $f : \mathbb{N} \to \mathbb{N}$ that takes a value below $B$ there is a largest number $k$ such that $f(k) \leq b$. diff --git a/src/elementary-number-theory/inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/inflationary-functions-natural-numbers.lagda.md new file mode 100644 index 0000000000..4e8f2a136f --- /dev/null +++ b/src/elementary-number-theory/inflationary-functions-natural-numbers.lagda.md @@ -0,0 +1,81 @@ +# Inflationary functions on the natural numbers + +```agda +module + elementary-number-theory.inflationary-functions-natural-numbers + where +``` + +
Imports + +```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers + +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.universe-levels + +open import order-theory.inflationary-maps-posets +``` + +
+ +## Idea + +A function $f : \mathbb{N} \to \mathbb{N}$ is an {{#concept "inflationary function" Disambiguation="natural numbers" Agda=is-inflationary-ℕ}} if the implication + +$$ + (x \leq y) \to (f(x) \leq f(y)) +$$ + +holds for every $x,y:\mathbb{N}$. If $f$ is an inflationary function on the natural numbers such that $f(0) \leq b$ for some natural number $b$, then there is a maximal natural number $k$ such that $f(k) \leq b$. + +## Definitions + +### The predicate of being an inflationary map on the natural numbers + +```agda +is-inflationary-ℕ : (ℕ → ℕ) → UU lzero +is-inflationary-ℕ = + is-inflationary-map-Poset ℕ-Poset +``` + +### Inflationary maps on the natural numbers + +```agda +inflationary-map-ℕ : UU lzero +inflationary-map-ℕ = + inflationary-map-Poset ℕ-Poset + +map-inflationary-map-ℕ : + inflationary-map-ℕ → ℕ → ℕ +map-inflationary-map-ℕ = + map-inflationary-map-Poset ℕ-Poset + +is-inflationary-inflationary-map-ℕ : + (f : inflationary-map-ℕ) → + is-inflationary-ℕ (map-inflationary-map-ℕ f) +is-inflationary-inflationary-map-ℕ = + is-inflationary-inflationary-map-Poset ℕ-Poset + +fiber-inflationary-map-ℕ : + (f : inflationary-map-ℕ) → ℕ → UU lzero +fiber-inflationary-map-ℕ f = fiber (map-inflationary-map-ℕ f) +``` + +## Properties + +### Being a value of an inflationary map is decidable + +```agda +is-decidable-fiber-inflationary-map-ℕ : + (f : inflationary-map-ℕ) → + is-decidable-fam (fiber-inflationary-map-ℕ f) +is-decidable-fiber-inflationary-map-ℕ = {!!} +``` + +### For any inflationary map $f$ with $f(0) \leq b$, there is a maximal number $k$ such that $f(k) \leq b$ + diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index ec9f968a2c..127c96875a 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -27,6 +27,8 @@ open import foundation.propositions open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels + +open import order-theory.strictly-ordered-types ```
@@ -121,12 +123,12 @@ contradiction-le-one-ℕ zero-ℕ () contradiction-le-one-ℕ (succ-ℕ n) () ``` -### The strict inequality on the natural numbers is anti-reflexive +### The strict inequality on the natural numbers is irreflexive ```agda -anti-reflexive-le-ℕ : (n : ℕ) → ¬ (n <-ℕ n) -anti-reflexive-le-ℕ zero-ℕ () -anti-reflexive-le-ℕ (succ-ℕ n) = anti-reflexive-le-ℕ n +irreflexive-le-ℕ : (n : ℕ) → ¬ (n <-ℕ n) +irreflexive-le-ℕ zero-ℕ () +irreflexive-le-ℕ (succ-ℕ n) = irreflexive-le-ℕ n ``` ### If `x < y` then `x ≠ y` @@ -148,7 +150,7 @@ antisymmetric-le-ℕ (succ-ℕ m) (succ-ℕ n) p q = ### The strict inequality on the natural numbers is transitive ```agda -transitive-le-ℕ : (n m l : ℕ) → (le-ℕ n m) → (le-ℕ m l) → (le-ℕ n l) +transitive-le-ℕ : (n m l : ℕ) → (le-ℕ m l) → (le-ℕ n m) → (le-ℕ n l) transitive-le-ℕ zero-ℕ (succ-ℕ m) (succ-ℕ l) p q = star transitive-le-ℕ (succ-ℕ n) (succ-ℕ m) (succ-ℕ l) p q = transitive-le-ℕ n m l p q @@ -156,7 +158,7 @@ transitive-le-ℕ (succ-ℕ n) (succ-ℕ m) (succ-ℕ l) p q = concatenate-le-eq-le-ℕ : (a b c d : ℕ) → a <-ℕ b → b = c → c <-ℕ d → a <-ℕ d concatenate-le-eq-le-ℕ a b .b d H refl K = - transitive-le-ℕ a b d H K + transitive-le-ℕ a b d K H ``` ### A sharper variant of transitivity @@ -186,6 +188,16 @@ linear-le-ℕ (succ-ℕ x) (succ-ℕ y) = map-coproduct id (map-coproduct (ap succ-ℕ) id) (linear-le-ℕ x y) ``` +### The strictly ordered type of natural numbers + +```agda +ℕ-Strictly-Ordered-Type : Strictly-Ordered-Type lzero lzero +pr1 ℕ-Strictly-Ordered-Type = ℕ +pr1 (pr2 ℕ-Strictly-Ordered-Type) = le-ℕ-Prop +pr1 (pr2 (pr2 ℕ-Strictly-Ordered-Type)) = irreflexive-le-ℕ +pr2 (pr2 (pr2 ℕ-Strictly-Ordered-Type)) = transitive-le-ℕ +``` + ### `n < m` if and only if there exists a nonzero natural number `l` such that `n + l = m` ```agda @@ -219,7 +231,7 @@ succ-le-ℕ (succ-ℕ n) = succ-le-ℕ n preserves-le-succ-ℕ : (m n : ℕ) → le-ℕ m n → le-ℕ m (succ-ℕ n) preserves-le-succ-ℕ m n H = - transitive-le-ℕ m n (succ-ℕ n) H (succ-le-ℕ n) + transitive-le-ℕ m n (succ-ℕ n) (succ-le-ℕ n) H ``` ### Concatenating strict and nonstrict inequalities @@ -410,8 +422,8 @@ preserves-strict-order-add-ℕ a b c d H K = ( a +ℕ c) ( b +ℕ c) ( b +ℕ d) - ( preserves-strict-order-left-add-ℕ a b c H) ( preserves-strict-order-right-add-ℕ b c d K) + ( preserves-strict-order-left-add-ℕ a b c H) ``` ### Multiplication is strictly order preserving diff --git a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md new file mode 100644 index 0000000000..b504d9cbcd --- /dev/null +++ b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md @@ -0,0 +1,93 @@ +# Strictly inflationary functions on the natural numbers + +```agda +module + elementary-number-theory.strictly-inflationary-functions-natural-numbers + where +``` + +
Imports + +```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers + +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.universe-levels + +open import order-theory.inflationary-maps-strictly-ordered-types +``` + +
+ +## Idea + +A function $f : \mathbb{N} \to \mathbb{N}$ is a {{#concept "strictly inflationary function" Disambiguation="natural numbers" Agda=is-strictly-inflationary-ℕ}} if the implication + +$$ + (x < y) \to (f(x) < f(y)) +$$ + +holds for every $x,y:\mathbb{N}$. If $f$ is a strictly inflationary function on the natural numbers such that $f(0) \leq b$ for some natural number $b$, then there is a maximal natural number $k$ such that $f(k) \leq b$. + +## Definitions + +### The predicate of being a strictly inflationary map on the natural numbers + +```agda +is-strictly-inflationary-ℕ : (ℕ → ℕ) → UU lzero +is-strictly-inflationary-ℕ = + is-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type +``` + +### Strictly inflationary maps on the natural numbers + +```agda +strictly-inflationary-map-ℕ : UU lzero +strictly-inflationary-map-ℕ = + inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + +map-strictly-inflationary-map-ℕ : + strictly-inflationary-map-ℕ → ℕ → ℕ +map-strictly-inflationary-map-ℕ = + map-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + +is-strictly-inflationary-strictly-inflationary-map-ℕ : + (f : strictly-inflationary-map-ℕ) → + is-strictly-inflationary-ℕ (map-strictly-inflationary-map-ℕ f) +is-strictly-inflationary-strictly-inflationary-map-ℕ = + is-inflationary-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + +fiber-strictly-inflationary-map-ℕ : + (f : strictly-inflationary-map-ℕ) → ℕ → UU lzero +fiber-strictly-inflationary-map-ℕ f = fiber (map-strictly-inflationary-map-ℕ f) +``` + +## Properties + +### Strictly inflationary maps on the natural numbers are inflationary + +```agda +is-inflationary-strictly-inflationary-map-ℕ : + (f : strictly-inflationary-map-ℕ) + (n : ℕ) → n ≤-ℕ map-strictly-inflationary-map-ℕ f n +is-inflationary-strictly-inflationary-map-ℕ f n = + leq-le-ℕ n + ( map-strictly-inflationary-map-ℕ f n) + ( is-strictly-inflationary-strictly-inflationary-map-ℕ f n) +``` + +### Being a value of a strictly inflationary map is decidable + +```agda +is-decidable-fiber-strictly-inflationary-map-ℕ : + (f : strictly-inflationary-map-ℕ) → + is-decidable-fam (fiber-strictly-inflationary-map-ℕ f) +is-decidable-fiber-strictly-inflationary-map-ℕ = {!!} +``` + +### For any strictly inflationary map $f$ with $f(0) \leq b$, there is a maximal number $k$ such that $f(k) \leq b$ + diff --git a/src/univalent-combinatorics/standard-finite-types.lagda.md b/src/univalent-combinatorics/standard-finite-types.lagda.md index 7c0ab90427..88d54a1cf5 100644 --- a/src/univalent-combinatorics/standard-finite-types.lagda.md +++ b/src/univalent-combinatorics/standard-finite-types.lagda.md @@ -237,8 +237,8 @@ strict-upper-bound-nat-Fin (succ-ℕ k) (inl x) = ( nat-Fin k x) ( k) ( succ-ℕ k) - ( strict-upper-bound-nat-Fin k x) ( succ-le-ℕ k) + ( strict-upper-bound-nat-Fin k x) strict-upper-bound-nat-Fin (succ-ℕ k) (inr star) = succ-le-ℕ k From 9ff514715107b57ee23ce4494148367a7061e84c Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sat, 11 Jan 2025 13:06:38 -0500 Subject: [PATCH 46/72] maximal value-bound input property --- .../finite-maps-natural-numbers.lagda.md | 45 +++++++++++++++++++ .../finite-maps.lagda.md | 45 +++++++++++++++++++ 2 files changed, 90 insertions(+) create mode 100644 src/univalent-combinatorics/finite-maps.lagda.md diff --git a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md index a024fc0348..525c463d37 100644 --- a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md @@ -7,7 +7,14 @@ module elementary-number-theory.finite-maps-natural-numbers where
Imports ```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.maximal-structured-natural-numbers +open import elementary-number-theory.natural-numbers +open import foundation.propositions +open import foundation.universe-levels + +open import univalent-combinatorics.finite-maps ```
@@ -17,3 +24,41 @@ module elementary-number-theory.finite-maps-natural-numbers where A map $f : A \to \mathbb{N}$ is said to be a {{#concept "finite map" Disambiguation="natural numbers"}} if its [fibers](foundation-core.finite-types.md) are [finite](univalent-combinatorics.finite-types.md). Finite maps are [decidable](elementary-number-theory.decidable-maps-natural-numbers.md). Every finite map $f : \mathbb{N} \to \mathbb{N}$ has a definite lowest value, and for every finite map $f : \mathbb{N} \to \mathbb{N}$ that takes a value below $B$ there is a largest number $k$ such that $f(k) \leq b$. + +The condition that $f : \mathbb{N} \to \mathbb{N}$ is finite is equivalent to the condition that for every natural number $n$, if $f$ has a value below $n$ there is a maximal number $k$ such that $f(k)\leq n$. + +## Definitions + +### The predicate of being a finite map into the natural numbers + +```agda +module _ + {l : Level} {A : UU l} (f : A → ℕ) + where + + is-finite-prop-map-ℕ : Prop l + is-finite-prop-map-ℕ = is-finite-prop-map f + + is-finite-map-ℕ : UU l + is-finite-map-ℕ = is-finite-map f + + is-prop-is-finite-map-ℕ : is-prop is-finite-map-ℕ + is-prop-is-finite-map-ℕ = is-prop-is-finite-map f +``` + +### The maximal value-bound input property of a function on the natural numbers + +The maximal value-bound input property asserts that for every natural number $n$ there is a maximal number $k$ for which $f(k)\leq n$. Note that it is not necessarily the case that the value $f(k)$ is itself maximal. + +This property doesn't seem to have a widely recognized name, so we use an explicit descriptor. + +```agda +maximal-value-bound-input-property-map-ℕ : + (ℕ → ℕ) → UU lzero +maximal-value-bound-input-property-map-ℕ f = + (n : ℕ) → maximal-element-ℕ (λ k → f k ≤-ℕ n) +``` + +## Properties + +### Any finite map satisfies the maximal value-bound input property diff --git a/src/univalent-combinatorics/finite-maps.lagda.md b/src/univalent-combinatorics/finite-maps.lagda.md new file mode 100644 index 0000000000..1b77f62b1e --- /dev/null +++ b/src/univalent-combinatorics/finite-maps.lagda.md @@ -0,0 +1,45 @@ +# Finite maps + +```agda +module univalent-combinatorics.finite-maps where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.fibers-of-maps +open import foundation.propositions +open import foundation.universe-levels + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +A map $f : A \to B$ is said to be a {{#concept "finite map"}} if its [fibers](foundation-core.finite-types.md) are [finite](univalent-combinatorics.finite-types.md). + +Finite maps are [decidable](elementary-number-theory.decidable-maps-natural-numbers.md). + +## Definitions + +### The predicate of being a finite map + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where + + is-finite-prop-map : Prop (l1 ⊔ l2) + is-finite-prop-map = + Π-Prop B (λ y → is-finite-Prop (fiber f y)) + + is-finite-map : UU (l1 ⊔ l2) + is-finite-map = type-Prop is-finite-prop-map + + is-prop-is-finite-map : is-prop is-finite-map + is-prop-is-finite-map = is-prop-type-Prop is-finite-prop-map +``` From f12069811028e3696d8e4ba52784556f6333ced9 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 12 Jan 2025 11:35:54 -0500 Subject: [PATCH 47/72] finite subtypes of the natural numbers have maximal elements --- ...ecidable-subtypes-natural-numbers.lagda.md | 133 ++++++++++++++++++ .../finite-maps-natural-numbers.lagda.md | 22 ++- .../maximum-natural-numbers.lagda.md | 116 +++++++++++++-- .../upper-bounds-natural-numbers.lagda.md | 8 ++ .../standard-finite-trees.lagda.md | 3 +- 5 files changed, 271 insertions(+), 11 deletions(-) create mode 100644 src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md diff --git a/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md new file mode 100644 index 0000000000..4dee6666c6 --- /dev/null +++ b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md @@ -0,0 +1,133 @@ +# Decidable subtypes of the natural numbers + +```agda +module elementary-number-theory.decidable-subtypes-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.maximal-structured-natural-numbers +open import elementary-number-theory.maximum-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.well-ordering-principle-natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.decidable-subtypes +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.equivalences +open import foundation.function-types +open import foundation.inhabited-subtypes +open import foundation.propositional-truncations +open import foundation.transport-along-identifications +open import foundation.universe-levels + +open import univalent-combinatorics.counting +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +In this file we study [decidable subtypes](foundation.decidable-subtypes.md) of the [natural numbers](elementary-number-theory.natural-numbers.md). The type of all decidable subtypes of the natural numbers is the *decidable powerset of the natural numbers*. + +As a direct consequence of the [well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) of the natural numbers, which is formulated for decidable structures over the natural numbers, it follows that every [inhabited](foundation.inhabited-subtypes.md) decidable subtype of the natural numbers has a least element. We also show that every finite decidable subtype has a largest element. + +## Properties + +### Inhabited decidable subtypes of the natural numbers have a minimal element + +```agda +module _ + {l : Level} (P : decidable-subtype l ℕ) + where + + well-ordering-principle-decidable-subtype-ℕ : + is-inhabited-subtype (subtype-decidable-subtype P) → + minimal-element-ℕ (is-in-decidable-subtype P) + well-ordering-principle-decidable-subtype-ℕ H = + apply-universal-property-trunc-Prop H + ( minimal-element-ℕ-Prop (subtype-decidable-subtype P)) + ( λ (n , p) → + well-ordering-principle-ℕ + ( is-in-decidable-subtype P) + ( is-decidable-decidable-subtype P) + ( p)) +``` + +### Finite decidable subtypes of the natural numbers have a maximal element + +```agda +module _ + {l : Level} (P : decidable-subtype l ℕ) + where + + nat-maximal-element-count-type-decidable-subtype-ℕ : + count (type-decidable-subtype P) → ℕ + nat-maximal-element-count-type-decidable-subtype-ℕ (n , e) = + max-standard-finite-family-ℕ n + ( inclusion-decidable-subtype P ∘ map-equiv e) + + is-in-subtype-maximal-element-count-type-decidable-subtype-ℕ : + (e : count (type-decidable-subtype P)) → + type-decidable-subtype P → + is-in-decidable-subtype P + ( nat-maximal-element-count-type-decidable-subtype-ℕ e) + is-in-subtype-maximal-element-count-type-decidable-subtype-ℕ (zero-ℕ , e) x = + ex-falso (map-inv-equiv e x) + is-in-subtype-maximal-element-count-type-decidable-subtype-ℕ + ( succ-ℕ n , e) x = + tr + ( is-in-decidable-subtype P) + ( pr2 + ( is-attained-max-standard-finite-family-succ-ℕ n + ( inclusion-decidable-subtype P ∘ map-equiv e))) + ( pr2 + ( map-equiv e + ( pr1 + ( is-attained-max-standard-finite-family-succ-ℕ n + ( inclusion-decidable-subtype P ∘ map-equiv e))))) + + is-upper-bound-maximal-element-count-type-decidable-subtype-ℕ : + (e : count (type-decidable-subtype P)) → + (i : ℕ) → is-in-decidable-subtype P i → + i ≤-ℕ nat-maximal-element-count-type-decidable-subtype-ℕ e + is-upper-bound-maximal-element-count-type-decidable-subtype-ℕ + (zero-ℕ , e) i p = + ex-falso (map-inv-equiv e (i , p)) + is-upper-bound-maximal-element-count-type-decidable-subtype-ℕ + (succ-ℕ n , e) i p = + tr + ( λ x → x ≤-ℕ _) + ( ap pr1 (is-section-map-inv-equiv e (i , p))) + ( is-upper-bound-max-standard-finite-family-ℕ + ( succ-ℕ n) + ( inclusion-decidable-subtype P ∘ map-equiv e) + ( map-inv-equiv e (i , p))) + + maximal-element-count-type-decidable-subtype-ℕ : + count (type-decidable-subtype P) → + type-decidable-subtype P → + maximal-element-ℕ (is-in-decidable-subtype P) + pr1 (maximal-element-count-type-decidable-subtype-ℕ e K) = + nat-maximal-element-count-type-decidable-subtype-ℕ e + pr1 (pr2 (maximal-element-count-type-decidable-subtype-ℕ e K)) = + is-in-subtype-maximal-element-count-type-decidable-subtype-ℕ e K + pr2 (pr2 (maximal-element-count-type-decidable-subtype-ℕ e K)) i p = + is-upper-bound-maximal-element-count-type-decidable-subtype-ℕ e i p + + maximal-element-is-finite-decidable-subtype-ℕ : + is-finite (type-decidable-subtype P) → + is-inhabited-subtype (subtype-decidable-subtype P) → + maximal-element-ℕ (is-in-decidable-subtype P) + maximal-element-is-finite-decidable-subtype-ℕ H K = + apply-twice-universal-property-trunc-Prop H K + ( maximal-element-ℕ-Prop (subtype-decidable-subtype P)) + ( maximal-element-count-type-decidable-subtype-ℕ) +``` + diff --git a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md index 525c463d37..0bb6303a91 100644 --- a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md @@ -56,9 +56,27 @@ This property doesn't seem to have a widely recognized name, so we use an explic maximal-value-bound-input-property-map-ℕ : (ℕ → ℕ) → UU lzero maximal-value-bound-input-property-map-ℕ f = - (n : ℕ) → maximal-element-ℕ (λ k → f k ≤-ℕ n) + (n k : ℕ) → f k ≤-ℕ n → maximal-element-ℕ (λ k → f k ≤-ℕ n) ``` ## Properties -### Any finite map satisfies the maximal value-bound input property +### Any finite map on the natural numbers satisfies the maximal value-bound input property + +**Proof.** Consider a map $f : \mathbb{N} \to \mathbb{N}$ with finite fibers, and consider natural numbers $n$ and $k$ such that $f(k) \leq n$. Since $f$ has finite fibers, it follows that the type + +$$ + \sum_{i:\mathbb{N}}f(i)\leq n +$$ + +is a finite decidable subtype of the natural numbers + +```agda +module _ + (f : ℕ → ℕ) + where + + maximal-value-bound-input-property-is-finite-map-ℕ : + is-finite-map-ℕ f → maximal-value-bound-input-property-map-ℕ f + maximal-value-bound-input-property-is-finite-map-ℕ H = {!!} +``` diff --git a/src/elementary-number-theory/maximum-natural-numbers.lagda.md b/src/elementary-number-theory/maximum-natural-numbers.lagda.md index 97038be80f..e5483b75b9 100644 --- a/src/elementary-number-theory/maximum-natural-numbers.lagda.md +++ b/src/elementary-number-theory/maximum-natural-numbers.lagda.md @@ -8,14 +8,23 @@ module elementary-number-theory.maximum-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.natural-numbers open import foundation.action-on-identifications-binary-functions open import foundation.action-on-identifications-functions open import foundation.coproduct-types +open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.coproduct-types +open import foundation.functoriality-coproduct-types +open import foundation.functoriality-dependent-pair-types open import foundation.identity-types +open import foundation.negated-equality +open import foundation.type-arithmetic-empty-type open import foundation.unit-type open import order-theory.least-upper-bounds-posets @@ -27,7 +36,15 @@ open import univalent-combinatorics.standard-finite-types ## Idea -We define the operation of maximum (least upper bound) for the natural numbers. +The {{#concept "maximum" Disambiguation="natural numbers" Agda=max-ℕ}} of two [natural numbers](elementary-number-theory.natural-numbers.md) is a binary operation that returns the largest of two natural numbers. It is defined recursively by + +```text + max-ℕ 0 n = n + max-ℕ (m + 1) 0 = m + 1 + max-ℕ (m + 1) (n + 1) = (max m n) + 1. +``` + +Furthermore, we define the maximum of any [standard finite](univalent-combinatorics.standard-finite-types.md) family of natural numbers. ## Definition @@ -49,9 +66,10 @@ max-ℕ' x y = max-ℕ y x ### Maximum of elements of standard finite types ```agda -max-Fin-ℕ : (n : ℕ) → (Fin n → ℕ) → ℕ -max-Fin-ℕ zero-ℕ f = zero-ℕ -max-Fin-ℕ (succ-ℕ n) f = max-ℕ (f (inr star)) (max-Fin-ℕ n (λ k → f (inl k))) +max-standard-finite-family-ℕ : (n : ℕ) → (Fin n → ℕ) → ℕ +max-standard-finite-family-ℕ zero-ℕ f = zero-ℕ +max-standard-finite-family-ℕ (succ-ℕ n) f = + max-ℕ (max-standard-finite-family-ℕ n (λ k → f (inl k))) (f (inr star)) ``` ## Properties @@ -68,7 +86,7 @@ leq-max-ℕ (succ-ℕ k) (succ-ℕ m) zero-ℕ H K = H leq-max-ℕ (succ-ℕ k) (succ-ℕ m) (succ-ℕ n) H K = leq-max-ℕ k m n H K leq-left-leq-max-ℕ : - (k m n : ℕ) → (max-ℕ m n) ≤-ℕ k → m ≤-ℕ k + (k m n : ℕ) → max-ℕ m n ≤-ℕ k → m ≤-ℕ k leq-left-leq-max-ℕ k zero-ℕ zero-ℕ H = star leq-left-leq-max-ℕ k zero-ℕ (succ-ℕ n) H = star leq-left-leq-max-ℕ k (succ-ℕ m) zero-ℕ H = H @@ -76,18 +94,18 @@ leq-left-leq-max-ℕ (succ-ℕ k) (succ-ℕ m) (succ-ℕ n) H = leq-left-leq-max-ℕ k m n H leq-right-leq-max-ℕ : - (k m n : ℕ) → (max-ℕ m n) ≤-ℕ k → n ≤-ℕ k + (k m n : ℕ) → max-ℕ m n ≤-ℕ k → n ≤-ℕ k leq-right-leq-max-ℕ k zero-ℕ zero-ℕ H = star leq-right-leq-max-ℕ k zero-ℕ (succ-ℕ n) H = H leq-right-leq-max-ℕ k (succ-ℕ m) zero-ℕ H = star leq-right-leq-max-ℕ (succ-ℕ k) (succ-ℕ m) (succ-ℕ n) H = leq-right-leq-max-ℕ k m n H -left-leq-max-ℕ : (m n : ℕ) → leq-ℕ m (max-ℕ m n) +left-leq-max-ℕ : (m n : ℕ) → m ≤-ℕ max-ℕ m n left-leq-max-ℕ m n = leq-left-leq-max-ℕ (max-ℕ m n) m n (refl-leq-ℕ (max-ℕ m n)) -right-leq-max-ℕ : (m n : ℕ) → leq-ℕ n (max-ℕ m n) +right-leq-max-ℕ : (m n : ℕ) → n ≤-ℕ max-ℕ m n right-leq-max-ℕ m n = leq-right-leq-max-ℕ (max-ℕ m n) m n (refl-leq-ℕ (max-ℕ m n)) @@ -181,3 +199,85 @@ right-distributive-add-max-ℕ x y z = ( ( left-distributive-add-max-ℕ z x y) ∙ ( ap-max-ℕ (commutative-add-ℕ z x) (commutative-add-ℕ z y))) ``` + +### Decide the maximum value + +```agda +decide-max-ℕ : + (m n : ℕ) → (max-ℕ m n = m) + (max-ℕ m n = n) +decide-max-ℕ zero-ℕ n = + inr refl +decide-max-ℕ (succ-ℕ m) zero-ℕ = + inl refl +decide-max-ℕ (succ-ℕ m) (succ-ℕ n) = + map-coproduct (ap succ-ℕ) (ap succ-ℕ) (decide-max-ℕ m n) + +is-left-max-ℕ : + (m n : ℕ) → max-ℕ m n ≠ n → max-ℕ m n = m +is-left-max-ℕ m n H = + map-right-unit-law-coproduct-is-empty + ( max-ℕ m n = m) + ( max-ℕ m n = n) + ( H) + ( decide-max-ℕ m n) + +is-right-max-ℕ : + (m n : ℕ) → max-ℕ m n ≠ m → max-ℕ m n = n +is-right-max-ℕ m n H = + map-left-unit-law-coproduct-is-empty + ( max-ℕ m n = m) + ( max-ℕ m n = n) + ( H) + ( decide-max-ℕ m n) +``` + +### The maximum of a standard finite family of natural numbers with at least one element is attained + +```agda +is-attained-max-standard-finite-family-succ-ℕ' : + (n : ℕ) (f : Fin (succ-ℕ n) → ℕ) → + is-decidable (max-standard-finite-family-ℕ (succ-ℕ n) f = f (inr star)) → + fiber f (max-standard-finite-family-ℕ (succ-ℕ n) f) +is-attained-max-standard-finite-family-succ-ℕ' zero-ℕ f d = + ( inr star , refl) +is-attained-max-standard-finite-family-succ-ℕ' (succ-ℕ n) f (inl p) = + ( inr star , inv p) +is-attained-max-standard-finite-family-succ-ℕ' (succ-ℕ n) f (inr np) = + map-Σ + ( λ k → f k = _) + ( inl) + ( λ k p → + ( p) ∙ + ( inv + ( is-left-max-ℕ + ( max-standard-finite-family-ℕ (succ-ℕ n) (f ∘ inl)) + ( f (inr star)) + ( np)))) + ( is-attained-max-standard-finite-family-succ-ℕ' n + ( f ∘ inl) + ( has-decidable-equality-ℕ _ _)) + +is-attained-max-standard-finite-family-succ-ℕ : + (n : ℕ) (f : Fin (succ-ℕ n) → ℕ) → + fiber f (max-standard-finite-family-ℕ (succ-ℕ n) f) +is-attained-max-standard-finite-family-succ-ℕ n f = + is-attained-max-standard-finite-family-succ-ℕ' n f + ( has-decidable-equality-ℕ _ _) +``` + +### The maximum of a standard finite family of natural numbers is an upper bound + +```agda +is-upper-bound-max-standard-finite-family-ℕ : + (n : ℕ) (f : Fin n → ℕ) → (i : Fin n) → + f i ≤-ℕ max-standard-finite-family-ℕ n f +is-upper-bound-max-standard-finite-family-ℕ (succ-ℕ n) f (inl x) = + transitive-leq-ℕ + ( f (inl x)) + ( max-standard-finite-family-ℕ n (f ∘ inl)) + ( max-standard-finite-family-ℕ (succ-ℕ n) f) + ( left-leq-max-ℕ (max-standard-finite-family-ℕ n (f ∘ inl)) (f (inr star))) + ( is-upper-bound-max-standard-finite-family-ℕ n (f ∘ inl) x) +is-upper-bound-max-standard-finite-family-ℕ (succ-ℕ n) f (inr star) = + right-leq-max-ℕ (max-standard-finite-family-ℕ n (f ∘ inl)) (f (inr star)) +``` diff --git a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md index 677befd36a..330b0a6084 100644 --- a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md @@ -295,6 +295,14 @@ is-upper-bound-leq-is-upper-bound-ℕ P b H n K x p = ### Being an upper bound of a decidable type family is decidable, given an upper bound +The type family `is-upper-bound-ℕ P` is an example of a type family `Q` over the natural numbers satisfying + +```text + Q b → Π (n : ℕ), is-decidable (Q n) +``` + +for any natural number `b`. + ```agda is-decidable-is-upper-bound-ℕ' : {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) → diff --git a/src/univalent-combinatorics/standard-finite-trees.lagda.md b/src/univalent-combinatorics/standard-finite-trees.lagda.md index 9ef39e079f..7416c54e1b 100644 --- a/src/univalent-combinatorics/standard-finite-trees.lagda.md +++ b/src/univalent-combinatorics/standard-finite-trees.lagda.md @@ -50,7 +50,8 @@ number-nodes-Tree-Fin (tree-Fin (succ-ℕ n) f) = height-Tree-Fin : Tree-Fin → ℕ height-Tree-Fin (tree-Fin zero-ℕ f) = zero-ℕ height-Tree-Fin (tree-Fin (succ-ℕ n) f) = - succ-ℕ (max-Fin-ℕ (succ-ℕ n) (λ k → height-Tree-Fin (f k))) + succ-ℕ + ( max-standard-finite-family-ℕ (succ-ℕ n) (λ k → height-Tree-Fin (f k))) is-leaf-Tree-Fin : Tree-Fin → UU lzero is-leaf-Tree-Fin (tree-Fin zero-ℕ _) = unit From 1d66e6e89862d00aeb7de0ce89776be03a2fecbd Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 13 Jan 2025 22:56:10 -0500 Subject: [PATCH 48/72] structured maximal value-bound inputs --- .../bounded-natural-numbers.lagda.md | 6 + ...ecidable-subtypes-natural-numbers.lagda.md | 27 +- .../finite-maps-natural-numbers.lagda.md | 251 +++++++++++++++++- .../inequality-natural-numbers.lagda.md | 80 ++++-- .../maximum-natural-numbers.lagda.md | 2 +- ...rithmetic-cartesian-product-types.lagda.md | 15 ++ 6 files changed, 352 insertions(+), 29 deletions(-) diff --git a/src/elementary-number-theory/bounded-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-natural-numbers.lagda.md index c95c04bc67..99ee9966fa 100644 --- a/src/elementary-number-theory/bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bounded-natural-numbers.lagda.md @@ -23,6 +23,7 @@ open import foundation.subtypes open import foundation.universe-levels open import univalent-combinatorics.counting +open import univalent-combinatorics.finite-types open import univalent-combinatorics.standard-finite-types ``` @@ -114,6 +115,11 @@ count-bounded-ℕ : (n : ℕ) → count (bounded-ℕ n) pr1 (count-bounded-ℕ n) = succ-ℕ n pr2 (count-bounded-ℕ n) = equiv-count-bounded-ℕ n + +is-finite-bounded-ℕ : + (n : ℕ) → is-finite (bounded-ℕ n) +is-finite-bounded-ℕ n = + is-finite-count (count-bounded-ℕ n) ``` ## See also diff --git a/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md index 4dee6666c6..25c268a519 100644 --- a/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md +++ b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md @@ -7,6 +7,7 @@ module elementary-number-theory.decidable-subtypes-natural-numbers where
Imports ```agda +open import elementary-number-theory.bounded-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.maximal-structured-natural-numbers open import elementary-number-theory.maximum-natural-numbers @@ -22,12 +23,16 @@ open import foundation.equivalences open import foundation.function-types open import foundation.inhabited-subtypes open import foundation.propositional-truncations +open import foundation.propositions open import foundation.transport-along-identifications +open import foundation.type-arithmetic-dependent-pair-types open import foundation.universe-levels open import univalent-combinatorics.counting +open import univalent-combinatorics.dependent-pair-types open import univalent-combinatorics.finite-types open import univalent-combinatorics.standard-finite-types +open import univalent-combinatorics.decidable-subtypes ```
@@ -60,7 +65,7 @@ module _ ( p)) ``` -### Finite decidable subtypes of the natural numbers have a maximal element +### Inhatbited finite decidable subtypes of the natural numbers have a maximal element ```agda module _ @@ -131,3 +136,23 @@ module _ ( maximal-element-count-type-decidable-subtype-ℕ) ``` +### Decidable subtypes with a maximal element are finite + +```agda +module _ + {l : Level} (P : decidable-subtype l ℕ) + where + + is-finite-maximal-element-decidable-subtype-ℕ : + maximal-element-ℕ (is-in-decidable-subtype P) → + is-finite (type-decidable-subtype P) + is-finite-maximal-element-decidable-subtype-ℕ (m , H , K) = + is-finite-equiv + ( ( right-unit-law-Σ-is-contr + ( λ (x , p) → is-proof-irrelevant-leq-ℕ x m (K x p))) ∘e + ( equiv-right-swap-Σ)) + ( is-finite-type-decidable-subtype + ( P ∘ pr1) + ( is-finite-bounded-ℕ m)) +``` + diff --git a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md index 0bb6303a91..e0f3d918ee 100644 --- a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md @@ -7,14 +7,34 @@ module elementary-number-theory.finite-maps-natural-numbers where
Imports ```agda +open import elementary-number-theory.decidable-subtypes-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.maximal-structured-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.well-ordering-principle-natural-numbers +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-subtypes +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.logical-equivalences +open import foundation.propositional-truncations open import foundation.propositions +open import foundation.subtypes +open import foundation.type-arithmetic-cartesian-product-types +open import foundation.type-arithmetic-coproduct-types open import foundation.universe-levels +open import univalent-combinatorics.coproduct-types open import univalent-combinatorics.finite-maps +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.decidable-subtypes ```
@@ -50,13 +70,118 @@ module _ The maximal value-bound input property asserts that for every natural number $n$ there is a maximal number $k$ for which $f(k)\leq n$. Note that it is not necessarily the case that the value $f(k)$ is itself maximal. -This property doesn't seem to have a widely recognized name, so we use an explicit descriptor. +This property doesn't seem to have a widely recognized name, so we use an explicit descriptor. Suggestions are welcome. ```agda -maximal-value-bound-input-property-map-ℕ : +module _ + (f : ℕ → ℕ) (n : ℕ) + where + + is-value-bound-input-ℕ : + (k : ℕ) → UU lzero + is-value-bound-input-ℕ k = + f k ≤-ℕ n + + is-prop-is-value-bound-input-ℕ : + (k : ℕ) → is-prop (is-value-bound-input-ℕ k) + is-prop-is-value-bound-input-ℕ k = + is-prop-leq-ℕ (f k) n + + is-decidable-is-value-bound-input-ℕ : + (k : ℕ) → is-decidable (is-value-bound-input-ℕ k) + is-decidable-is-value-bound-input-ℕ k = + is-decidable-leq-ℕ (f k) n + + is-value-bound-input-subtype-ℕ : + subtype lzero ℕ + pr1 (is-value-bound-input-subtype-ℕ k) = + is-value-bound-input-ℕ k + pr2 (is-value-bound-input-subtype-ℕ k) = + is-prop-is-value-bound-input-ℕ k + + is-value-bound-input-decidable-subtype-ℕ : + decidable-subtype lzero ℕ + pr1 (is-value-bound-input-decidable-subtype-ℕ k) = + is-value-bound-input-ℕ k + pr1 (pr2 (is-value-bound-input-decidable-subtype-ℕ k)) = + is-prop-is-value-bound-input-ℕ k + pr2 (pr2 (is-value-bound-input-decidable-subtype-ℕ k)) = + is-decidable-is-value-bound-input-ℕ k + +maximal-value-bound-input-property-ℕ : (ℕ → ℕ) → UU lzero -maximal-value-bound-input-property-map-ℕ f = - (n k : ℕ) → f k ≤-ℕ n → maximal-element-ℕ (λ k → f k ≤-ℕ n) +maximal-value-bound-input-property-ℕ f = + (n k : ℕ) → f k ≤-ℕ n → maximal-element-ℕ (is-value-bound-input-ℕ f n) +``` + +### The maximal structured value-bound input property of a function on the natural numbers + +```agda +module _ + {l : Level} (P : ℕ → UU l) + where + + is-structured-value-bound-input-ℕ : + (f : ℕ → ℕ) (n k : ℕ) → UU l + is-structured-value-bound-input-ℕ f n k = + f k ≤-ℕ n × P (f k) + + is-prop-is-structured-value-bound-input-ℕ : + ((x : ℕ) → is-prop (P x)) → + (f : ℕ → ℕ) (n k : ℕ) → is-prop (is-structured-value-bound-input-ℕ f n k) + is-prop-is-structured-value-bound-input-ℕ H f n k = + is-prop-product (is-prop-leq-ℕ (f k) n) (H (f k)) + + is-decidable-is-structured-value-bound-input-ℕ : + is-decidable-fam P → (f : ℕ → ℕ) (n k : ℕ) → + is-decidable (is-structured-value-bound-input-ℕ f n k) + is-decidable-is-structured-value-bound-input-ℕ d f n k = + is-decidable-product (is-decidable-leq-ℕ (f k) n) (d (f k)) + +module _ + {l : Level} (P : ℕ → Prop l) + where + + is-structured-value-bound-input-subtype-ℕ : + (f : ℕ → ℕ) (n : ℕ) → subtype l ℕ + pr1 (is-structured-value-bound-input-subtype-ℕ f n k) = + is-structured-value-bound-input-ℕ (type-Prop ∘ P) f n k + pr2 (is-structured-value-bound-input-subtype-ℕ f n k) = + is-prop-is-structured-value-bound-input-ℕ + ( type-Prop ∘ P) + ( is-prop-type-Prop ∘ P) + ( f) + ( n) + ( k) + +module _ + {l : Level} (P : decidable-subtype l ℕ) + where + + is-structured-value-bound-input-decidable-subtype-ℕ : + (f : ℕ → ℕ) (n : ℕ) → decidable-subtype l ℕ + pr1 (is-structured-value-bound-input-decidable-subtype-ℕ f n k) = + is-structured-value-bound-input-ℕ (is-in-decidable-subtype P) f n k + pr1 (pr2 (is-structured-value-bound-input-decidable-subtype-ℕ f n k)) = + is-prop-is-structured-value-bound-input-ℕ + ( is-in-decidable-subtype P) + ( is-prop-is-in-decidable-subtype P) + ( f) + ( n) + ( k) + pr2 (pr2 (is-structured-value-bound-input-decidable-subtype-ℕ f n k)) = + is-decidable-is-structured-value-bound-input-ℕ + ( is-in-decidable-subtype P) + ( is-decidable-decidable-subtype P) + ( f) + ( n) + ( k) + +maximal-structured-value-bound-input-property-ℕ : + {l : Level} (P : ℕ → UU l) → (ℕ → ℕ) → UU l +maximal-structured-value-bound-input-property-ℕ P f = + (n k : ℕ) → f k ≤-ℕ n → P (f k) → + maximal-element-ℕ (is-structured-value-bound-input-ℕ P f n) ``` ## Properties @@ -69,14 +194,124 @@ $$ \sum_{i:\mathbb{N}}f(i)\leq n $$ -is a finite decidable subtype of the natural numbers +is a finite decidable subtype of the natural numbers, with at least one element. + +For the converse, suppose that $f$ satisfies the maximal value-bound input property. The fiber of $f$ at $n$ is a decidable subtype of the type + +$$ + \sum_{i:\mathbb{N}}f(i)\leq n. +$$ + +Since the type above itself is a decidable subtype of the natural numbers with a maximal element, it is finite. It therefore follows that the fiber of $f$ at $n$ is finite. ```agda module _ - (f : ℕ → ℕ) + (f : ℕ → ℕ) (H : is-finite-map-ℕ f) where + is-finite-is-value-bound-input-decidable-subtype-ℕ : + (n : ℕ) → is-finite (type-subtype (is-value-bound-input-subtype-ℕ f n)) + is-finite-is-value-bound-input-decidable-subtype-ℕ zero-ℕ = + is-finite-equiv' + ( equiv-tot + ( λ x → + equiv-iff-is-prop + ( is-prop-is-value-bound-input-ℕ f 0 x) + ( is-set-ℕ (f x) 0) + ( is-zero-leq-zero-ℕ (f x)) + ( leq-eq-ℕ (f x) 0))) + ( H 0) + is-finite-is-value-bound-input-decidable-subtype-ℕ (succ-ℕ n) = + is-finite-equiv' + ( left-distributive-Σ-coproduct ℕ _ _ ∘e + ( equiv-tot (λ x → compute-leq-succ-ℕ (f x) n))) + ( is-finite-coproduct + ( is-finite-is-value-bound-input-decidable-subtype-ℕ n) + ( H (succ-ℕ n))) + maximal-value-bound-input-property-is-finite-map-ℕ : - is-finite-map-ℕ f → maximal-value-bound-input-property-map-ℕ f - maximal-value-bound-input-property-is-finite-map-ℕ H = {!!} + maximal-value-bound-input-property-ℕ f + maximal-value-bound-input-property-is-finite-map-ℕ n k K = + maximal-element-is-finite-decidable-subtype-ℕ + ( is-value-bound-input-decidable-subtype-ℕ f n) + ( is-finite-is-value-bound-input-decidable-subtype-ℕ n) + ( unit-trunc-Prop (k , K)) + + nat-maximal-value-bound-input-property-is-finite-map-ℕ : + (n k : ℕ) → f k ≤-ℕ n → ℕ + nat-maximal-value-bound-input-property-is-finite-map-ℕ n k K = + nat-maximal-element-ℕ + ( is-value-bound-input-ℕ f n) + ( maximal-value-bound-input-property-is-finite-map-ℕ n k K) + + is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ : + (n k : ℕ) (K : f k ≤-ℕ n) (x : ℕ) → f x ≤-ℕ n → + x ≤-ℕ nat-maximal-value-bound-input-property-is-finite-map-ℕ n k K + is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ n k K = + is-upper-bound-maximal-element-ℕ + ( is-value-bound-input-ℕ f n) + ( maximal-value-bound-input-property-is-finite-map-ℕ n k K) + + is-value-bound-input-maximal-value-bound-input-property-is-finite-map-ℕ : + (n k : ℕ) (K : f k ≤-ℕ n) → + is-value-bound-input-ℕ f n + ( nat-maximal-value-bound-input-property-is-finite-map-ℕ n k K) + is-value-bound-input-maximal-value-bound-input-property-is-finite-map-ℕ + n k K = + structure-maximal-element-ℕ + ( is-value-bound-input-ℕ f n) + ( maximal-value-bound-input-property-is-finite-map-ℕ n k K) +``` + +### Any finite map on the natural numbers satisfies the maximal structured value-bound input property + +Consider the following data: + +- A family $P$ of decidable types over the natural numbers +- A finite map $f$ +- A natural number $n$ +- A natural number $k$ such that $f(k) \leq n$ equipped with an element of type $P(f(k))$. + +Then there is a largest natural number $m$ such that $f(m) \leq n$ equipped with an element of type $P(f(m))$. + +**Example.** For any natural number $1 \leq n$ there exists a largest number $m$ such that $m^2 \mid n$. Indeed, if we let $P(x)$ be the decidable predicate that $x$ divides $n$ and $f$ the [squaring function](elementary-number-theory.squares-natural-numbers.md). The squaring function has finite fibers and we have $1^2 \leq n$ and $1^2 \mid n$. The theorem thus gives us a largest number $m$ such that $m^2 \mid n$. We carry out this example in the module about [Square-free decompositions](elementary-number-theory.square-free-decompositions-natural-numbers.md). + +```agda +module _ + {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) + (f : ℕ → ℕ) (H : is-finite-map-ℕ f) + where + + maximal-structured-value-bound-input-property-is-finite-map-ℕ : + maximal-structured-value-bound-input-property-ℕ P f + maximal-structured-value-bound-input-property-is-finite-map-ℕ n k K p = + ( nat-bounded-maximal-element-ℕ + ( λ x → is-value-bound-input-ℕ f n x × P (f x)) + ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) + ( b) , + ( structure-bounded-maximal-element-ℕ + ( λ x → is-value-bound-input-ℕ f n x × P (f x)) + ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) + ( b)) , + ( λ x (u , v) → + is-upper-bound-bounded-maximal-element-ℕ + ( λ x → is-value-bound-input-ℕ f n x × P (f x)) + ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) + ( b) + ( x) + ( is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ + f H n k K x u , + u , + v))) + where + + b = + bounded-maximal-element-instance-ℕ + ( λ x → is-value-bound-input-ℕ f n x × P (f x)) + ( λ x → is-decidable-product (is-decidable-leq-ℕ (f x) n) (d (f x))) + ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) + ( k) + ( is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ + f H n k K k K) + ( K , p) ``` diff --git a/src/elementary-number-theory/inequality-natural-numbers.lagda.md b/src/elementary-number-theory/inequality-natural-numbers.lagda.md index 49ed531c42..53b6fdbeb5 100644 --- a/src/elementary-number-theory/inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/inequality-natural-numbers.lagda.md @@ -10,21 +10,26 @@ module elementary-number-theory.inequality-natural-numbers where open import category-theory.precategories open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import foundation.action-on-identifications-functions open import foundation.binary-relations open import foundation.cartesian-product-types +open import foundation.contractible-types open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.equivalences open import foundation.function-types open import foundation.functoriality-coproduct-types open import foundation.identity-types +open import foundation.logical-equivalences open import foundation.negation open import foundation.propositions +open import foundation.sections open import foundation.unit-type open import foundation.universe-levels @@ -77,6 +82,11 @@ is-prop-leq-ℕ (succ-ℕ m) (succ-ℕ n) = is-prop-leq-ℕ m n leq-ℕ-Prop : ℕ → ℕ → Prop lzero pr1 (leq-ℕ-Prop m n) = leq-ℕ m n pr2 (leq-ℕ-Prop m n) = is-prop-leq-ℕ m n + +is-proof-irrelevant-leq-ℕ : + (m n : ℕ) → m ≤-ℕ n → is-contr (m ≤-ℕ n) +is-proof-irrelevant-leq-ℕ m n = + is-proof-irrelevant-is-prop (is-prop-leq-ℕ m n) ``` ### Inequality on the natural numbers is decidable @@ -247,6 +257,25 @@ leq-succ-leq-ℕ : leq-succ-leq-ℕ m n p = transitive-leq-ℕ m n (succ-ℕ n) (succ-leq-ℕ n) p ``` +### The successor of `n` is not less than or equal to `n` + +```agda +neg-succ-leq-ℕ : + (n : ℕ) → ¬ (succ-ℕ n ≤-ℕ n) +neg-succ-leq-ℕ zero-ℕ = id +neg-succ-leq-ℕ (succ-ℕ n) = neg-succ-leq-ℕ n +``` + +### `m ≤ n` if and only if `n + 1 ≰ m` + +```agda +contradiction-leq-ℕ : (m n : ℕ) → m ≤-ℕ n → ¬ (succ-ℕ n ≤-ℕ m) +contradiction-leq-ℕ (succ-ℕ m) (succ-ℕ n) H K = contradiction-leq-ℕ m n H K + +contradiction-leq-ℕ' : (m n : ℕ) → succ-ℕ n ≤-ℕ m → ¬ (m ≤-ℕ n) +contradiction-leq-ℕ' m n K H = contradiction-leq-ℕ m n H K +``` + ### Any natural number less than or equal to `n+1` is either less than or equal to `n` or it is `n+1` ```agda @@ -259,6 +288,38 @@ decide-leq-succ-ℕ (succ-ℕ m) zero-ℕ l = decide-leq-succ-ℕ (succ-ℕ m) (succ-ℕ n) l = map-coproduct id (ap succ-ℕ) (decide-leq-succ-ℕ m n l) +inv-decide-leq-succ-ℕ : + (m n : ℕ) → (m ≤-ℕ n) + (m = succ-ℕ n) → m ≤-ℕ succ-ℕ n +inv-decide-leq-succ-ℕ m n (inl H) = + transitive-leq-ℕ m n (succ-ℕ n) (leq-succ-leq-ℕ n n (refl-leq-ℕ n)) H +inv-decide-leq-succ-ℕ m n (inr p) = + leq-eq-ℕ m (succ-ℕ n) p + +all-elements-equal-cases-leq-succ-ℕ : + (m n : ℕ) → all-elements-equal ((m ≤-ℕ n) + (m = succ-ℕ n)) +all-elements-equal-cases-leq-succ-ℕ m n (inl H) (inl K) = + ap inl (eq-is-prop (is-prop-leq-ℕ m n)) +all-elements-equal-cases-leq-succ-ℕ .(succ-ℕ n) n (inl H) (inr refl) = + ex-falso (neg-succ-leq-ℕ n H) +all-elements-equal-cases-leq-succ-ℕ .(succ-ℕ n) n (inr refl) (inl K) = + ex-falso (neg-succ-leq-ℕ n K) +all-elements-equal-cases-leq-succ-ℕ .(succ-ℕ n) n (inr refl) (inr q) = + ap inr (eq-is-prop (is-set-ℕ (succ-ℕ n) (succ-ℕ n))) + +is-prop-cases-leq-succ-ℕ : + (m n : ℕ) → is-prop ((m ≤-ℕ n) + (m = succ-ℕ n)) +is-prop-cases-leq-succ-ℕ m n = + is-prop-all-elements-equal (all-elements-equal-cases-leq-succ-ℕ m n) + +compute-leq-succ-ℕ : + (m n : ℕ) → m ≤-ℕ succ-ℕ n ≃ (m ≤-ℕ n) + (m = succ-ℕ n) +compute-leq-succ-ℕ m n = + equiv-iff-is-prop + ( is-prop-leq-ℕ m (succ-ℕ n)) + ( is-prop-cases-leq-succ-ℕ m n) + ( decide-leq-succ-ℕ m n) + ( inv-decide-leq-succ-ℕ m n) + decide-leq-refl-succ-ℕ : {n : ℕ} → decide-leq-succ-ℕ (succ-ℕ n) n (refl-leq-ℕ n) = inr refl decide-leq-refl-succ-ℕ {zero-ℕ} = refl @@ -277,25 +338,6 @@ decide-leq-ℕ (succ-ℕ m) (succ-ℕ n) H = map-coproduct (ap succ-ℕ) id (decide-leq-ℕ m n H) ``` -### The successor of `n` is not less than or equal to `n` - -```agda -neg-succ-leq-ℕ : - (n : ℕ) → ¬ (succ-ℕ n ≤-ℕ n) -neg-succ-leq-ℕ zero-ℕ = id -neg-succ-leq-ℕ (succ-ℕ n) = neg-succ-leq-ℕ n -``` - -### `m ≤ n` if and only if `n + 1 ≰ m` - -```agda -contradiction-leq-ℕ : (m n : ℕ) → m ≤-ℕ n → ¬ (succ-ℕ n ≤-ℕ m) -contradiction-leq-ℕ (succ-ℕ m) (succ-ℕ n) H K = contradiction-leq-ℕ m n H K - -contradiction-leq-ℕ' : (m n : ℕ) → succ-ℕ n ≤-ℕ m → ¬ (m ≤-ℕ n) -contradiction-leq-ℕ' m n K H = contradiction-leq-ℕ m n H K -``` - ### Addition preserves inequality of natural numbers ```agda diff --git a/src/elementary-number-theory/maximum-natural-numbers.lagda.md b/src/elementary-number-theory/maximum-natural-numbers.lagda.md index e5483b75b9..22e54f6d82 100644 --- a/src/elementary-number-theory/maximum-natural-numbers.lagda.md +++ b/src/elementary-number-theory/maximum-natural-numbers.lagda.md @@ -269,7 +269,7 @@ is-attained-max-standard-finite-family-succ-ℕ n f = ```agda is-upper-bound-max-standard-finite-family-ℕ : - (n : ℕ) (f : Fin n → ℕ) → (i : Fin n) → + (n : ℕ) (f : Fin n → ℕ) (i : Fin n) → f i ≤-ℕ max-standard-finite-family-ℕ n f is-upper-bound-max-standard-finite-family-ℕ (succ-ℕ n) f (inl x) = transitive-leq-ℕ diff --git a/src/foundation/type-arithmetic-cartesian-product-types.lagda.md b/src/foundation/type-arithmetic-cartesian-product-types.lagda.md index 4d0d338b3b..798dae6897 100644 --- a/src/foundation/type-arithmetic-cartesian-product-types.lagda.md +++ b/src/foundation/type-arithmetic-cartesian-product-types.lagda.md @@ -110,6 +110,13 @@ module _ inv-right-unit-law-product-is-contr = inv-equiv right-unit-law-product-is-contr +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (H : A → is-contr B) + where + + right-unit-law-product-is-contr' : A × B ≃ A + right-unit-law-product-is-contr' = right-unit-law-Σ-is-contr H + module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} (is-contr-A : is-contr A) where @@ -132,6 +139,14 @@ module _ equiv-pr2-product-is-contr : (A × B) ≃ B pr1 equiv-pr2-product-is-contr = pr2 pr2 equiv-pr2-product-is-contr = is-equiv-pr2-product-is-contr + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (H : B → is-contr A) + where + + left-unit-law-product-is-contr' : A × B ≃ B + left-unit-law-product-is-contr' = + right-unit-law-product-is-contr' H ∘e commutative-product ``` ### Adding redundant properties From 89abe8486ce4ab5a784655469746f786ca76f727 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 14 Jan 2025 22:10:49 -0500 Subject: [PATCH 49/72] any two distinct Fermat numbers are relatively prime --- references.bib | 12 +++ .../distance-natural-numbers.lagda.md | 72 +++++++------- .../divisibility-natural-numbers.lagda.md | 3 + .../exponentiation-natural-numbers.lagda.md | 27 ++++-- .../factorials.lagda.md | 33 ++++++- .../fermat-numbers.lagda.md | 83 +++++++++++++++- .../finite-maps-natural-numbers.lagda.md | 12 +-- ...st-common-divisor-natural-numbers.lagda.md | 3 +- ...-arithmetic-standard-finite-types.lagda.md | 16 ++++ .../parity-natural-numbers.lagda.md | 26 ++++- .../products-of-natural-numbers.lagda.md | 33 +++++++ .../pronic-numbers.lagda.md | 4 +- .../proper-divisors-natural-numbers.lagda.md | 43 ++++++++- .../squares-natural-numbers.lagda.md | 6 +- ...strict-inequality-natural-numbers.lagda.md | 94 +++++++++++-------- .../strictly-bounded-natural-numbers.lagda.md | 10 +- .../sums-of-natural-numbers.lagda.md | 36 +++++++ 17 files changed, 400 insertions(+), 113 deletions(-) diff --git a/references.bib b/references.bib index 80ab6d619f..64520bd954 100644 --- a/references.bib +++ b/references.bib @@ -608,6 +608,18 @@ @misc{Mye21 primaryclass = {math.CT} } +@book {NZM, + author = {Niven, Ivan and Zuckerman, Herbert S. and Montgomery, Hugh L.}, + title = {An introduction to the theory of numbers}, + edition = {Fifth}, + publisher = {John Wiley \& Sons, Inc., New York}, + year = {1991}, + pages = {xiv+529}, + isbn = {0-471-62546-9}, + mrclass = {11-01}, + mrnumber = {1083765}, +} + @online{oeis, title = {The {{On-Line Encyclopedia}} of {{Integer Sequences}}}, author = {OEIS Foundation Inc.}, diff --git a/src/elementary-number-theory/distance-natural-numbers.lagda.md b/src/elementary-number-theory/distance-natural-numbers.lagda.md index a0121c276e..07288f0299 100644 --- a/src/elementary-number-theory/distance-natural-numbers.lagda.md +++ b/src/elementary-number-theory/distance-natural-numbers.lagda.md @@ -29,10 +29,12 @@ open import foundation.universe-levels ## Idea The -{{#concept "distance function" Disambiguation="between natural numbers" Agda=dist-ℕ}} +{{#concept "distance function" Disambiguation="natural numbers" Agda=dist-ℕ}} between [natural numbers](elementary-number-theory.natural-numbers.md) measures -how far two natural numbers are apart. In the agda-unimath library we often -prefer to work with `dist-ℕ` over the partially defined subtraction operation. +how far two natural numbers are apart. If the [inequality](elementary-number-theory.inequality-natural-numbers.md) $x \leq y$ holds, then the distance between $x$ and $y$ is the unique natural number $d$ equipped with an [identification](foundation-core.identity-types.md) $x + d = y$. + +**Note.** In the agda-unimath library, we often +prefer to work with `dist-ℕ` over the subtraction operation, which is either partially defined or it returns nonsensical values. Not only is the distance function sensibly defined for any pair of natural numbers, but it also satisfies the more pleasant and predictable properties of a [metric](metric-spaces.metric-spaces.md). ## Definition @@ -104,39 +106,39 @@ right-unit-law-dist-ℕ (succ-ℕ n) = refl ```agda triangle-inequality-dist-ℕ : - (m n k : ℕ) → (dist-ℕ m n) ≤-ℕ ((dist-ℕ m k) +ℕ (dist-ℕ k n)) + (m n k : ℕ) → dist-ℕ m n ≤-ℕ dist-ℕ m k +ℕ dist-ℕ k n triangle-inequality-dist-ℕ zero-ℕ zero-ℕ zero-ℕ = star triangle-inequality-dist-ℕ zero-ℕ zero-ℕ (succ-ℕ k) = star triangle-inequality-dist-ℕ zero-ℕ (succ-ℕ n) zero-ℕ = tr - ( leq-ℕ (succ-ℕ n)) + ( succ-ℕ n ≤-ℕ_) ( inv (left-unit-law-add-ℕ (succ-ℕ n))) ( refl-leq-ℕ (succ-ℕ n)) triangle-inequality-dist-ℕ zero-ℕ (succ-ℕ n) (succ-ℕ k) = concatenate-eq-leq-eq-ℕ ( inv (ap succ-ℕ (left-unit-law-dist-ℕ n))) ( triangle-inequality-dist-ℕ zero-ℕ n k) - ( ( ap (succ-ℕ ∘ (_+ℕ (dist-ℕ k n))) (left-unit-law-dist-ℕ k)) ∙ + ( ( ap (succ-ℕ ∘ (_+ℕ dist-ℕ k n)) (left-unit-law-dist-ℕ k)) ∙ ( inv (left-successor-law-add-ℕ k (dist-ℕ k n)))) triangle-inequality-dist-ℕ (succ-ℕ m) zero-ℕ zero-ℕ = refl-leq-ℕ (succ-ℕ m) triangle-inequality-dist-ℕ (succ-ℕ m) zero-ℕ (succ-ℕ k) = concatenate-eq-leq-eq-ℕ ( inv (ap succ-ℕ (right-unit-law-dist-ℕ m))) ( triangle-inequality-dist-ℕ m zero-ℕ k) - ( ap (succ-ℕ ∘ ((dist-ℕ m k) +ℕ_)) (right-unit-law-dist-ℕ k)) + ( ap (succ-ℕ ∘ (dist-ℕ m k +ℕ_)) (right-unit-law-dist-ℕ k)) triangle-inequality-dist-ℕ (succ-ℕ m) (succ-ℕ n) zero-ℕ = concatenate-leq-eq-ℕ ( dist-ℕ m n) ( transitive-leq-ℕ ( dist-ℕ m n) - ( succ-ℕ ((dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n))) - ( succ-ℕ (succ-ℕ ((dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n)))) - ( succ-leq-ℕ (succ-ℕ ((dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n)))) + ( succ-ℕ (dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n)) + ( succ-ℕ (succ-ℕ (dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n))) + ( succ-leq-ℕ (succ-ℕ (dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n))) ( transitive-leq-ℕ ( dist-ℕ m n) - ( (dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n)) - ( succ-ℕ ((dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n))) - ( succ-leq-ℕ ((dist-ℕ m zero-ℕ) +ℕ (dist-ℕ zero-ℕ n))) + ( dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n) + ( succ-ℕ (dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n)) + ( succ-leq-ℕ (dist-ℕ m zero-ℕ +ℕ dist-ℕ zero-ℕ n)) ( triangle-inequality-dist-ℕ m n zero-ℕ))) ( ( ap ( succ-ℕ ∘ succ-ℕ) @@ -185,7 +187,7 @@ rewrite-right-dist-add-ℕ x .(dist-ℕ x z) z H refl = is-additive-right-inverse-dist-ℕ x z H is-difference-dist-ℕ : - (x y : ℕ) → x ≤-ℕ y → x +ℕ (dist-ℕ x y) = y + (x y : ℕ) → x ≤-ℕ y → x +ℕ dist-ℕ x y = y is-difference-dist-ℕ zero-ℕ zero-ℕ H = refl is-difference-dist-ℕ zero-ℕ (succ-ℕ y) H = left-unit-law-add-ℕ (succ-ℕ y) is-difference-dist-ℕ (succ-ℕ x) (succ-ℕ y) H = @@ -193,7 +195,7 @@ is-difference-dist-ℕ (succ-ℕ x) (succ-ℕ y) H = ( ap succ-ℕ (is-difference-dist-ℕ x y H)) is-difference-dist-ℕ' : - (x y : ℕ) → x ≤-ℕ y → (dist-ℕ x y) +ℕ x = y + (x y : ℕ) → x ≤-ℕ y → dist-ℕ x y +ℕ x = y is-difference-dist-ℕ' x y H = ( commutative-add-ℕ (dist-ℕ x y) x) ∙ ( is-difference-dist-ℕ x y H) @@ -213,8 +215,8 @@ dist-add-ℕ' x y = symmetric-dist-ℕ (x +ℕ y) x ∙ dist-add-ℕ x y ```agda triangle-equality-dist-ℕ : - (x y z : ℕ) → (x ≤-ℕ y) → (y ≤-ℕ z) → - (dist-ℕ x y) +ℕ (dist-ℕ y z) = dist-ℕ x z + (x y z : ℕ) → x ≤-ℕ y → y ≤-ℕ z → + dist-ℕ x y +ℕ dist-ℕ y z = dist-ℕ x z triangle-equality-dist-ℕ zero-ℕ zero-ℕ zero-ℕ H1 H2 = refl triangle-equality-dist-ℕ zero-ℕ zero-ℕ (succ-ℕ z) star star = ap succ-ℕ (left-unit-law-add-ℕ z) @@ -227,40 +229,40 @@ triangle-equality-dist-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) H1 H2 = cases-dist-ℕ : (x y z : ℕ) → UU lzero cases-dist-ℕ x y z = - ( (dist-ℕ x y) +ℕ (dist-ℕ y z) = dist-ℕ x z) + - ( ( (dist-ℕ y z) +ℕ (dist-ℕ x z) = dist-ℕ x y) + - ( (dist-ℕ x z) +ℕ (dist-ℕ x y) = dist-ℕ y z)) + ( dist-ℕ x y +ℕ dist-ℕ y z = dist-ℕ x z) + + ( ( dist-ℕ y z +ℕ dist-ℕ x z = dist-ℕ x y) + + ( dist-ℕ x z +ℕ dist-ℕ x y = dist-ℕ y z)) is-total-dist-ℕ : (x y z : ℕ) → cases-dist-ℕ x y z is-total-dist-ℕ x y z with order-three-elements-ℕ x y z -is-total-dist-ℕ x y z | inl (inl (pair H1 H2)) = +is-total-dist-ℕ x y z | inl (inl (H1 , H2)) = inl (triangle-equality-dist-ℕ x y z H1 H2) -is-total-dist-ℕ x y z | inl (inr (pair H1 H2)) = +is-total-dist-ℕ x y z | inl (inr (H1 , H2)) = inr ( inl ( ( commutative-add-ℕ (dist-ℕ y z) (dist-ℕ x z)) ∙ - ( ( ap ((dist-ℕ x z) +ℕ_) (symmetric-dist-ℕ y z)) ∙ + ( ( ap (dist-ℕ x z +ℕ_) (symmetric-dist-ℕ y z)) ∙ ( triangle-equality-dist-ℕ x z y H1 H2)))) -is-total-dist-ℕ x y z | inr (inl (inl (pair H1 H2))) = +is-total-dist-ℕ x y z | inr (inl (inl (H1 , H2))) = inr ( inl - ( ( ap ((dist-ℕ y z) +ℕ_) (symmetric-dist-ℕ x z)) ∙ + ( ( ap (dist-ℕ y z +ℕ_) (symmetric-dist-ℕ x z)) ∙ ( ( triangle-equality-dist-ℕ y z x H1 H2) ∙ ( symmetric-dist-ℕ y x)))) -is-total-dist-ℕ x y z | inr (inl (inr (pair H1 H2))) = +is-total-dist-ℕ x y z | inr (inl (inr (H1 , H2))) = inr ( inr - ( ( ap ((dist-ℕ x z) +ℕ_) (symmetric-dist-ℕ x y)) ∙ + ( ( ap (dist-ℕ x z +ℕ_) (symmetric-dist-ℕ x y)) ∙ ( ( commutative-add-ℕ (dist-ℕ x z) (dist-ℕ y x)) ∙ ( triangle-equality-dist-ℕ y x z H1 H2)))) -is-total-dist-ℕ x y z | inr (inr (inl (pair H1 H2))) = +is-total-dist-ℕ x y z | inr (inr (inl (H1 , H2))) = inr ( inr - ( ( ap (_+ℕ (dist-ℕ x y)) (symmetric-dist-ℕ x z)) ∙ + ( ( ap (_+ℕ dist-ℕ x y) (symmetric-dist-ℕ x z)) ∙ ( ( triangle-equality-dist-ℕ z x y H1 H2) ∙ ( symmetric-dist-ℕ z y)))) -is-total-dist-ℕ x y z | inr (inr (inr (pair H1 H2))) = +is-total-dist-ℕ x y z | inr (inr (inr (H1 , H2))) = inl ( ( ap-add-ℕ (symmetric-dist-ℕ x y) (symmetric-dist-ℕ y z)) ∙ ( ( commutative-add-ℕ (dist-ℕ y x) (dist-ℕ z y)) ∙ @@ -283,7 +285,7 @@ leq-dist-ℕ (succ-ℕ x) (succ-ℕ y) H = ```agda strict-upper-bound-dist-ℕ : - (b x y : ℕ) → le-ℕ x b → le-ℕ y b → le-ℕ (dist-ℕ x y) b + (b x y : ℕ) → x <-ℕ b → y <-ℕ b → dist-ℕ x y <-ℕ b strict-upper-bound-dist-ℕ (succ-ℕ b) zero-ℕ y H K = K strict-upper-bound-dist-ℕ (succ-ℕ b) (succ-ℕ x) zero-ℕ H K = H strict-upper-bound-dist-ℕ (succ-ℕ b) (succ-ℕ x) (succ-ℕ y) H K = @@ -294,14 +296,14 @@ strict-upper-bound-dist-ℕ (succ-ℕ b) (succ-ℕ x) (succ-ℕ y) H K = ```agda right-successor-law-dist-ℕ : - (x y : ℕ) → leq-ℕ x y → dist-ℕ x (succ-ℕ y) = succ-ℕ (dist-ℕ x y) + (x y : ℕ) → x ≤-ℕ y → dist-ℕ x (succ-ℕ y) = succ-ℕ (dist-ℕ x y) right-successor-law-dist-ℕ zero-ℕ zero-ℕ star = refl right-successor-law-dist-ℕ zero-ℕ (succ-ℕ y) star = refl right-successor-law-dist-ℕ (succ-ℕ x) (succ-ℕ y) H = right-successor-law-dist-ℕ x y H left-successor-law-dist-ℕ : - (x y : ℕ) → leq-ℕ y x → dist-ℕ (succ-ℕ x) y = succ-ℕ (dist-ℕ x y) + (x y : ℕ) → y ≤-ℕ x → dist-ℕ (succ-ℕ x) y = succ-ℕ (dist-ℕ x y) left-successor-law-dist-ℕ zero-ℕ zero-ℕ star = refl left-successor-law-dist-ℕ (succ-ℕ x) zero-ℕ star = refl left-successor-law-dist-ℕ (succ-ℕ x) (succ-ℕ y) H = @@ -330,7 +332,7 @@ translation-invariant-dist-ℕ' k m n = ```agda left-distributive-mul-dist-ℕ : - (m n k : ℕ) → k *ℕ (dist-ℕ m n) = dist-ℕ (k *ℕ m) (k *ℕ n) + (m n k : ℕ) → k *ℕ dist-ℕ m n = dist-ℕ (k *ℕ m) (k *ℕ n) left-distributive-mul-dist-ℕ zero-ℕ zero-ℕ zero-ℕ = refl left-distributive-mul-dist-ℕ zero-ℕ zero-ℕ (succ-ℕ k) = left-distributive-mul-dist-ℕ zero-ℕ zero-ℕ k @@ -357,7 +359,7 @@ left-distributive-mul-dist-ℕ (succ-ℕ m) (succ-ℕ n) (succ-ℕ k) = ( inv (left-distributive-mul-dist-ℕ m n (succ-ℕ k)))) right-distributive-mul-dist-ℕ : - (x y k : ℕ) → (dist-ℕ x y) *ℕ k = dist-ℕ (x *ℕ k) (y *ℕ k) + (x y k : ℕ) → dist-ℕ x y *ℕ k = dist-ℕ (x *ℕ k) (y *ℕ k) right-distributive-mul-dist-ℕ x y k = ( commutative-mul-ℕ (dist-ℕ x y) k) ∙ ( ( left-distributive-mul-dist-ℕ x y k) ∙ diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 10c6c595f0..ef046f0ad2 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -325,6 +325,9 @@ pr2 (refl-div-ℕ x) = left-unit-law-mul-ℕ x div-eq-ℕ : (x y : ℕ) → x = y → div-ℕ x y div-eq-ℕ x .x refl = refl-div-ℕ x +neq-neg-div-ℕ : (x y : ℕ) → ¬ div-ℕ x y → x ≠ y +neq-neg-div-ℕ x y H p = H (div-eq-ℕ x y p) + antisymmetric-div-ℕ : is-antisymmetric div-ℕ antisymmetric-div-ℕ zero-ℕ zero-ℕ H K = refl antisymmetric-div-ℕ zero-ℕ (succ-ℕ y) (k , p) K = diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index 3dacba5c3f..4272144754 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -12,6 +12,7 @@ open import commutative-algebra.powers-of-elements-commutative-semirings open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.commutative-semiring-of-natural-numbers +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -150,10 +151,24 @@ compute-constant-product-ℕ m (succ-ℕ (succ-ℕ n)) = ### The base of the exponent divides its successor exponents -```text -bounded-div-exp-succ-ℕ : - (m n : ℕ) → bounded-div-ℕ m (m ^ℕ succ-ℕ n) -pr1 (bounded-div-exp-succ-ℕ m n) = m ^ℕ n -pr1 (pr2 (bounded-div-exp-succ-ℕ m n)) = {!!} -pr2 (pr2 (bounded-div-exp-succ-ℕ m n)) = {!!} +```agda +div-exp-succ-ℕ : + (m n : ℕ) → div-ℕ m (m ^ℕ succ-ℕ n) +pr1 (div-exp-succ-ℕ m n) = m ^ℕ n +pr2 (div-exp-succ-ℕ m n) = inv (exp-succ-ℕ n m) + +div-exp-is-successor-ℕ : + (m n : ℕ) → is-successor-ℕ n → div-ℕ m (m ^ℕ n) +div-exp-is-successor-ℕ m .(succ-ℕ k) (k , refl) = + div-exp-succ-ℕ m k + +div-exp-ℕ : + (m n : ℕ) → is-nonzero-ℕ n → div-ℕ m (m ^ℕ n) +div-exp-ℕ m n H = + div-exp-is-successor-ℕ m n (is-successor-is-nonzero-ℕ H) + +bounded-div-exp-ℕ : + (m n : ℕ) → is-nonzero-ℕ n → bounded-div-ℕ m (m ^ℕ n) +bounded-div-exp-ℕ m n H = + bounded-div-div-ℕ m (m ^ℕ n) (div-exp-ℕ m n H) ``` diff --git a/src/elementary-number-theory/factorials.lagda.md b/src/elementary-number-theory/factorials.lagda.md index 05851e426e..a6b3dd39b7 100644 --- a/src/elementary-number-theory/factorials.lagda.md +++ b/src/elementary-number-theory/factorials.lagda.md @@ -12,6 +12,7 @@ open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.relatively-prime-natural-numbers open import foundation.coproduct-types open import foundation.dependent-pair-types @@ -27,7 +28,9 @@ The {{#concept "factorial" WD="factorial" WDID=Q120976 Agda=factorial-ℕ}} `n!` of a number `n`, counts the number of ways to linearly order `n` distinct objects. -## Definition +## Definitions + +### Factorial numbers ```agda factorial-ℕ : ℕ → ℕ @@ -35,6 +38,13 @@ factorial-ℕ 0 = 1 factorial-ℕ (succ-ℕ m) = (factorial-ℕ m) *ℕ (succ-ℕ m) ``` +### Successors of factorial numbers + +```agda +succ-factorial-ℕ : ℕ → ℕ +succ-factorial-ℕ n = succ-ℕ (factorial-ℕ n) +``` + ## Properties ```agda @@ -73,7 +83,28 @@ leq-factorial-ℕ (succ-ℕ n) = ( is-nonzero-factorial-ℕ n) ``` +### Any two consecutive successor factorials are relatively prime + +This property is stated as exercise 53 in section 1.2 of {{#cite NZM}}. + +```agda +is-one-is-common-divisor-consecutive-succ-factorial-ℕ : + (n d : ℕ) → + div-ℕ d (succ-factorial-ℕ n) → div-ℕ d (succ-factorial-ℕ (succ-ℕ n)) → + is-one-ℕ d +is-one-is-common-divisor-consecutive-succ-factorial-ℕ n d H K = {!!} + +is-relatively-prime-consecutive-succ-factorial-ℕ : + (n : ℕ) → + is-relatively-prime-ℕ (succ-factorial-ℕ n) (succ-factorial-ℕ (succ-ℕ n)) +is-relatively-prime-consecutive-succ-factorial-ℕ n = {!!} +``` + ## External links - [Factorial](https://en.wikipedia.org/wiki/Factorial) at Wikipedia - [A000142](https://oeis.org/A000142) in the OEIS + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index 58e31b2127..32adc85b44 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -8,19 +8,29 @@ module elementary-number-theory.fermat-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.congruence-natural-numbers open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers +open import elementary-number-theory.greatest-common-divisor-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.products-of-natural-numbers +open import elementary-number-theory.relatively-prime-natural-numbers open import elementary-number-theory.squares-natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import elementary-number-theory.strong-induction-natural-numbers open import foundation.action-on-identifications-functions +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.function-types open import foundation.identity-types open import foundation.negated-equality +open import foundation.transport-along-identifications open import foundation.unit-type open import univalent-combinatorics.standard-finite-types @@ -111,7 +121,18 @@ leq-two-fermat-number-ℕ : leq-two-fermat-number-ℕ n = leq-one-exp-ℕ 2 (2 ^ℕ n) star ``` -### The Fermat number of a successor +### Every Fermat number is odd + +```agda +is-odd-fermat-number-ℕ : + (n : ℕ) → is-odd-ℕ (fermat-number-ℕ n) +is-odd-fermat-number-ℕ n = + is-odd-succ-is-even-ℕ + ( 2 ^ℕ (2 ^ℕ n)) + ( div-exp-ℕ 2 (2 ^ℕ n) (is-nonzero-exp-ℕ 2 n (is-nonzero-succ-ℕ 1))) +``` + +### The distance from a Fermat number to the number $2$ Any Fermat number of the form $F(n+1)=2^{2^{n+1}}+1$ can be computed as @@ -129,7 +150,11 @@ dist-fermat-number-two-ℕ (succ-ℕ n) = ap (λ x → dist-ℕ x 1) (ap (2 ^ℕ_) (exp-succ-ℕ n 2) ∙ exp-mul-ℕ (2 ^ℕ n) 2) ∙ distance-of-squares-ℕ (2 ^ℕ 2 ^ℕ n) 1 ∙ ap (_*ℕ fermat-number-ℕ n) (dist-fermat-number-two-ℕ n) +``` +### Computing the $n$-th Fermat number as the product of all the previous Fermat numbers plus $2$ + +```agda compute-fermat-number-ℕ : (n : ℕ) → fermat-number-ℕ n = Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)) +ℕ 2 @@ -145,7 +170,7 @@ compute-fermat-number-ℕ n = ( commutative-add-ℕ 2 (Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)))) ``` -### The two definitions of the Fermat numbers agree +### The $n$th Fermat number is equal to the $n$th recursive Fermat number **Proof.** The proof is by strong induction on the natural numbers. The base case holds by reflexivity. For the inductive step, assume that @@ -174,10 +199,58 @@ $F_{n+1}=(\prod_{n=0}^n F_n) + 2$. This implies that if $d$ divides $F_m$ and $F_n$ for some $n>m$, then $d|2$. However, the Fermat numbers are odd, so $d=1$. ```agda -is-one-div-fermat-number-ℕ : - (m n d : ℕ) → m ≠ n → +is-one-div-fermat-number-nat-Fin-ℕ : + (n d : ℕ) (i : Fin n) → + div-ℕ d (fermat-number-ℕ (nat-Fin n i)) → + div-ℕ d (fermat-number-ℕ n) → + is-one-ℕ d +is-one-div-fermat-number-nat-Fin-ℕ (succ-ℕ n) d i H K = + is-one-is-odd-div-two-ℕ d + ( div-right-summand-ℕ d + ( Π-ℕ (succ-ℕ n) (λ j → fermat-number-ℕ (nat-Fin (succ-ℕ n) j))) + ( 2) + ( transitive-div-ℕ d + ( fermat-number-ℕ (nat-Fin (succ-ℕ n) i)) + ( Π-ℕ (succ-ℕ n) (λ j → fermat-number-ℕ (nat-Fin (succ-ℕ n) j))) + ( div-factor-Π-ℕ (succ-ℕ n) (fermat-number-ℕ ∘ nat-Fin (succ-ℕ n)) i) + ( H)) + ( concatenate-div-eq-ℕ K (compute-fermat-number-ℕ (succ-ℕ n)))) + ( is-odd-div-is-odd-ℕ d + ( fermat-number-ℕ (succ-ℕ n)) + ( is-odd-fermat-number-ℕ (succ-ℕ n)) + ( K)) + +is-one-div-fermat-number-le-ℕ : + (m n d : ℕ) → m <-ℕ n → div-ℕ d (fermat-number-ℕ m) → div-ℕ d (fermat-number-ℕ n) → is-one-ℕ d -is-one-div-fermat-number-ℕ m n d H K L = {!!} +is-one-div-fermat-number-le-ℕ m (succ-ℕ n) d H K L = + is-one-div-fermat-number-nat-Fin-ℕ + ( succ-ℕ n) + ( d) + ( mod-succ-ℕ n m) + ( tr + ( div-ℕ d ∘ fermat-number-ℕ) + ( inv (eq-nat-mod-succ-ℕ n m H)) + ( K)) + ( L) + +is-one-is-common-divisor-fermat-number-ℕ : + (m n d : ℕ) → m ≠ n → + is-common-divisor-ℕ (fermat-number-ℕ m) (fermat-number-ℕ n) d → is-one-ℕ d +is-one-is-common-divisor-fermat-number-ℕ m n d H (K , L) + with + decide-le-neq-ℕ m n H +... | inl u = is-one-div-fermat-number-le-ℕ m n d u K L +... | inr u = is-one-div-fermat-number-le-ℕ n m d u L K + +is-relatively-prime-fermat-number-ℕ : + (m n : ℕ) → m ≠ n → + is-relatively-prime-ℕ (fermat-number-ℕ m) (fermat-number-ℕ n) +is-relatively-prime-fermat-number-ℕ m n H = + is-one-is-common-divisor-fermat-number-ℕ m n + ( gcd-ℕ (fermat-number-ℕ m) (fermat-number-ℕ n)) + ( H) + ( is-common-divisor-gcd-ℕ (fermat-number-ℕ m) (fermat-number-ℕ n)) ``` ## External link diff --git a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md index e0f3d918ee..db2426a1dc 100644 --- a/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finite-maps-natural-numbers.lagda.md @@ -285,10 +285,10 @@ module _ maximal-structured-value-bound-input-property-is-finite-map-ℕ : maximal-structured-value-bound-input-property-ℕ P f maximal-structured-value-bound-input-property-is-finite-map-ℕ n k K p = - ( nat-bounded-maximal-element-ℕ + ( ( nat-bounded-maximal-element-ℕ ( λ x → is-value-bound-input-ℕ f n x × P (f x)) ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) - ( b) , + ( b)) , ( structure-bounded-maximal-element-ℕ ( λ x → is-value-bound-input-ℕ f n x × P (f x)) ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) @@ -299,10 +299,10 @@ module _ ( nat-maximal-value-bound-input-property-is-finite-map-ℕ f H n k K) ( b) ( x) - ( is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ - f H n k K x u , - u , - v))) + ( ( is-upper-bound-nat-maximal-value-bound-input-property-is-finite-map-ℕ + f H n k K x u) , + ( u) , + ( v)))) where b = diff --git a/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md b/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md index 1134193c8f..0ae0e6fd99 100644 --- a/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md +++ b/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md @@ -15,6 +15,7 @@ open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.euclidean-division-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -169,7 +170,7 @@ abstract well-ordering-principle-ℕ ( is-multiple-of-gcd-ℕ a b) ( is-decidable-is-multiple-of-gcd-ℕ a b) - ( pair (a +ℕ b) (sum-is-multiple-of-gcd-ℕ a b)) + ( sum-is-multiple-of-gcd-ℕ a b) gcd-ℕ : ℕ → ℕ → ℕ gcd-ℕ a b = pr1 (GCD-ℕ a b) diff --git a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md index 2291f16380..aca1cbe6bc 100644 --- a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md +++ b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md @@ -15,6 +15,7 @@ open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-binary-functions open import foundation.action-on-identifications-functions @@ -192,6 +193,21 @@ leq-nat-mod-succ-ℕ k (succ-ℕ x) = ( leq-nat-succ-Fin (succ-ℕ k) (mod-succ-ℕ k x)) ``` +### If `x < k + 1`, then `x` is equal to its residue modulo `k + 1` + +```agda +eq-nat-mod-succ-ℕ : + (k x : ℕ) → x <-ℕ succ-ℕ k → nat-Fin (succ-ℕ k) (mod-succ-ℕ k x) = x +eq-nat-mod-succ-ℕ k x H = + eq-cong-le-ℕ + ( succ-ℕ k) + ( nat-Fin (succ-ℕ k) (mod-succ-ℕ k x)) + ( x) + ( strict-upper-bound-nat-Fin (succ-ℕ k) (mod-succ-ℕ k x)) + ( H) + ( cong-nat-mod-succ-ℕ k x) +``` + ## Operations ### Addition on the standard finite sets diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 991964b988..c240daf58f 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -18,6 +18,8 @@ open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.prime-numbers +open import elementary-number-theory.proper-divisors-natural-numbers open import foundation.action-on-identifications-functions open import foundation.cartesian-product-types @@ -278,13 +280,33 @@ has-odd-expansion-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = ```agda is-even-div-is-even-ℕ : - (n m : ℕ) → is-even-ℕ m → div-ℕ m n → is-even-ℕ n + (m n : ℕ) → is-even-ℕ m → div-ℕ m n → is-even-ℕ n is-even-div-is-even-ℕ ._ ._ (m' , refl) (k , refl) = k *ℕ m' , associative-mul-ℕ k m' 2 is-even-div-4-ℕ : (n : ℕ) → div-ℕ 4 n → is-even-ℕ n -is-even-div-4-ℕ n = is-even-div-is-even-ℕ n 4 (2 , refl) +is-even-div-4-ℕ n = is-even-div-is-even-ℕ 4 n (2 , refl) +``` + +### Any divisor of an odd number is odd + +```agda +is-odd-div-is-odd-ℕ : + (m n : ℕ) → is-odd-ℕ n → div-ℕ m n → is-odd-ℕ m +is-odd-div-is-odd-ℕ m n H K L = + H (is-even-div-is-even-ℕ m n L K) +``` + +### Any odd divisor of `2` is equal to `1` + +```agda +is-one-is-odd-div-two-ℕ : + (n : ℕ) → div-ℕ n 2 → is-odd-ℕ n → is-one-ℕ n +is-one-is-odd-div-two-ℕ n H K = + is-one-is-proper-divisor-two-ℕ n + ( ( neq-neg-div-ℕ 2 n K ∘ inv) , + ( H)) ``` ### A number is even if and only if it is congruent to `0` modulo `2` diff --git a/src/elementary-number-theory/products-of-natural-numbers.lagda.md b/src/elementary-number-theory/products-of-natural-numbers.lagda.md index 915ea27987..145ab2e7e7 100644 --- a/src/elementary-number-theory/products-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/products-of-natural-numbers.lagda.md @@ -7,12 +7,14 @@ module elementary-number-theory.products-of-natural-numbers where
Imports ```agda +open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions open import foundation.coproduct-types +open import foundation.dependent-pair-types open import foundation.function-extensionality open import foundation.function-types open import foundation.homotopies @@ -21,6 +23,7 @@ open import foundation.unit-type open import lists.lists +open import univalent-combinatorics.skipping-element-standard-finite-types open import univalent-combinatorics.standard-finite-types ``` @@ -84,3 +87,33 @@ preserves-htpy-Π-ℕ : preserves-htpy-Π-ℕ k H = ap (Π-ℕ k) (eq-htpy H) ``` + +### Any factor of a product of natural numbers divides the product + +```agda +quotient-div-factor-Π-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → ℕ +quotient-div-factor-Π-ℕ (succ-ℕ k) f i = + Π-ℕ k (f ∘ skip-Fin k i) + +eq-quotient-div-factor-Π-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → + quotient-div-factor-Π-ℕ k f i *ℕ f i = Π-ℕ k f +eq-quotient-div-factor-Π-ℕ (succ-ℕ zero-ℕ) f (inr star) = + refl +eq-quotient-div-factor-Π-ℕ (succ-ℕ (succ-ℕ k)) f (inl i) = + ( right-swap-mul-ℕ + ( Π-ℕ k (f ∘ inl ∘ skip-Fin k i)) + ( f (inr star)) + ( f (inl i))) ∙ + ( ap (_*ℕ f (inr star)) (eq-quotient-div-factor-Π-ℕ (succ-ℕ k) (f ∘ inl) i)) +eq-quotient-div-factor-Π-ℕ (succ-ℕ (succ-ℕ k)) f (inr star) = + refl + +div-factor-Π-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → div-ℕ (f i) (Π-ℕ k f) +pr1 (div-factor-Π-ℕ k f i) = + quotient-div-factor-Π-ℕ k f i +pr2 (div-factor-Π-ℕ k f i) = + eq-quotient-div-factor-Π-ℕ k f i +``` diff --git a/src/elementary-number-theory/pronic-numbers.lagda.md b/src/elementary-number-theory/pronic-numbers.lagda.md index bc50942eac..fea16a1b64 100644 --- a/src/elementary-number-theory/pronic-numbers.lagda.md +++ b/src/elementary-number-theory/pronic-numbers.lagda.md @@ -79,14 +79,14 @@ abstract with is-even-or-is-even-succ-ℕ n ... | inl H = is-even-div-is-even-ℕ - ( pronic-number-ℕ n) ( n) + ( pronic-number-ℕ n) ( H) ( succ-ℕ n , commutative-mul-ℕ (succ-ℕ n) n) ... | inr H = is-even-div-is-even-ℕ - ( pronic-number-ℕ n) ( succ-ℕ n) + ( pronic-number-ℕ n) ( H) ( n , refl) ``` diff --git a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md index 4deb6c75a3..833bae1a15 100644 --- a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md @@ -32,32 +32,48 @@ open import foundation.universe-levels ## Idea -A proper divisor of a natural number `n` is a natural number `d ≠ n` that -divides `n`. +A {{#concept "proper divisor" Disambituation="natural numbers" Agda=is-proper-divisor-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a natural number `d ≠ n` that +[divides](elementary-number-theory.divisbility-natural-numbers.md) `n`. + +## Definitions + +### The predicate of being a proper divisor of a natural number ```agda is-proper-divisor-ℕ : ℕ → ℕ → UU lzero is-proper-divisor-ℕ n d = (d ≠ n) × (div-ℕ d n) +``` + +## Properties + +### Being a proper divisor of a natural number is decidable +```agda is-decidable-is-proper-divisor-ℕ : (n d : ℕ) → is-decidable (is-proper-divisor-ℕ n d) is-decidable-is-proper-divisor-ℕ n d = is-decidable-product ( is-decidable-neg (has-decidable-equality-ℕ d n)) ( is-decidable-div-ℕ d n) +``` + +### Any successor is a proper divisor of $0$ +```agda is-proper-divisor-zero-succ-ℕ : (n : ℕ) → is-proper-divisor-ℕ zero-ℕ (succ-ℕ n) pr1 (is-proper-divisor-zero-succ-ℕ n) = is-nonzero-succ-ℕ n pr2 (is-proper-divisor-zero-succ-ℕ n) = div-zero-ℕ (succ-ℕ n) +``` +### Any proper divisor of a natural number is strictly smaller than its dividend + +```agda le-is-proper-divisor-ℕ : (x y : ℕ) → is-nonzero-ℕ y → is-proper-divisor-ℕ y x → le-ℕ x y le-is-proper-divisor-ℕ x y H K = le-leq-neq-ℕ (leq-div-ℕ x y H (pr2 K)) (pr1 K) ``` -## Properties - ### Being a proper divisor is a property ```agda @@ -97,3 +113,22 @@ le-one-quotient-div-is-proper-divisor-ℕ d x f H g = ( quotient-div-ℕ d x H) ( leq-one-quotient-div-ℕ d x H (leq-one-is-nonzero-ℕ x f))) ``` + +### A divisor `d` of a number `n` is a proper divisor if and only if `x ∤ d` + +```agda +neg-div-is-proper-divisor-ℕ : + (d n : ℕ) → is-proper-divisor-ℕ n d → ¬ div-ℕ n d +neg-div-is-proper-divisor-ℕ d n (np , H) K = + np (antisymmetric-div-ℕ d n H K) + +is-proper-divisor-neg-div-ℕ : + (d n : ℕ) → div-ℕ d n → ¬ div-ℕ n d → is-proper-divisor-ℕ n d +pr1 (is-proper-divisor-neg-div-ℕ d .d H K) refl = K H +pr2 (is-proper-divisor-neg-div-ℕ d n H K) = H +``` + +## See also + +- [Prime numbers](elementary-number-theory.prime-numbers.md) +- [Nontrivial divisors](elementary-number-theory.nontrivial-divisors-natural-numbers.md) diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index 3afee61c53..d9d98e9d3c 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -270,13 +270,13 @@ strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ (succ-ℕ n))) H = ( n +ℕ 2) ( square-ℕ (n +ℕ 2)) ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) - ( strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ n)) star) ( transitive-le-ℕ ( square-ℕ (n +ℕ 2)) ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) ( square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) - ( le-add-succ-ℕ (square-ℕ (n +ℕ 2)) (n +ℕ 1)) - ( le-add-succ-ℕ (square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) (n +ℕ 1)))) + ( le-add-succ-ℕ (square-ℕ (n +ℕ 2) +ℕ (n +ℕ 2)) (n +ℕ 1)) + ( le-add-succ-ℕ (square-ℕ (n +ℕ 2)) (n +ℕ 1))) + ( strict-is-inflationary-square-ℕ (succ-ℕ (succ-ℕ n)) star)) ( inv (square-succ-ℕ' (n +ℕ 2))) ``` diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 127c96875a..91c094166d 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -46,8 +46,8 @@ le-ℕ-Prop (succ-ℕ n) (succ-ℕ m) = le-ℕ-Prop n m le-ℕ : ℕ → ℕ → UU lzero le-ℕ n m = type-Prop (le-ℕ-Prop n m) -is-prop-le-ℕ : (n : ℕ) → (m : ℕ) → is-prop (le-ℕ n m) -is-prop-le-ℕ n m = is-prop-type-Prop (le-ℕ-Prop n m) +is-prop-le-ℕ : (m n : ℕ) → is-prop (le-ℕ m n) +is-prop-le-ℕ m n = is-prop-type-Prop (le-ℕ-Prop m n) infix 30 _<-ℕ_ _<-ℕ_ = le-ℕ @@ -59,23 +59,23 @@ _<-ℕ_ = le-ℕ ```agda concatenate-eq-le-eq-ℕ : - (x y z w : ℕ) → x = y → le-ℕ y z → z = w → le-ℕ x w -concatenate-eq-le-eq-ℕ x y z .z refl p refl = p + (m n k l : ℕ) → m = n → n <-ℕ k → k = l → m <-ℕ l +concatenate-eq-le-eq-ℕ m n k .k refl p refl = p concatenate-eq-le-ℕ : - (x y z : ℕ) → x = y → le-ℕ y z → le-ℕ x z -concatenate-eq-le-ℕ x .x z refl p = p + (m n k : ℕ) → m = n → n <-ℕ k → m <-ℕ k +concatenate-eq-le-ℕ m .m k refl p = p concatenate-le-eq-ℕ : - (x y z : ℕ) → le-ℕ x y → y = z → le-ℕ x z -concatenate-le-eq-ℕ x y .y p refl = p + (m n k : ℕ) → m <-ℕ n → n = k → m <-ℕ k +concatenate-le-eq-ℕ m n .n p refl = p ``` ### Strict inequality on the natural numbers is decidable ```agda is-decidable-le-ℕ : - (m n : ℕ) → is-decidable (le-ℕ m n) + (m n : ℕ) → is-decidable (m <-ℕ n) is-decidable-le-ℕ zero-ℕ zero-ℕ = inr id is-decidable-le-ℕ zero-ℕ (succ-ℕ n) = inl star is-decidable-le-ℕ (succ-ℕ m) zero-ℕ = inr id @@ -85,14 +85,14 @@ is-decidable-le-ℕ (succ-ℕ m) (succ-ℕ n) = is-decidable-le-ℕ m n ### If `m < n` then `n` must be nonzero ```agda -is-nonzero-le-ℕ : (m n : ℕ) → le-ℕ m n → is-nonzero-ℕ n +is-nonzero-le-ℕ : (m n : ℕ) → m <-ℕ n → is-nonzero-ℕ n is-nonzero-le-ℕ m .zero-ℕ () refl ``` ### Any nonzero natural number is strictly greater than `0` ```agda -le-zero-is-nonzero-ℕ : (n : ℕ) → is-nonzero-ℕ n → le-ℕ zero-ℕ n +le-zero-is-nonzero-ℕ : (n : ℕ) → is-nonzero-ℕ n → zero-ℕ <-ℕ n le-zero-is-nonzero-ℕ zero-ℕ H = H refl le-zero-is-nonzero-ℕ (succ-ℕ n) H = star ``` @@ -101,9 +101,9 @@ le-zero-is-nonzero-ℕ (succ-ℕ n) H = star ```agda contradiction-le-zero-ℕ : - (m : ℕ) → (le-ℕ m zero-ℕ) → empty + (n : ℕ) → ¬ (n <-ℕ zero-ℕ) contradiction-le-zero-ℕ zero-ℕ () -contradiction-le-zero-ℕ (succ-ℕ m) () +contradiction-le-zero-ℕ (succ-ℕ n) () ``` ### Any natural number strictly greater than another natural number is strictly greater than `0` @@ -118,7 +118,7 @@ le-zero-le-ℕ m (succ-ℕ n) H = star ```agda contradiction-le-one-ℕ : - (n : ℕ) → le-ℕ (succ-ℕ n) 1 → empty + (n : ℕ) → ¬ (succ-ℕ n <-ℕ 1) contradiction-le-one-ℕ zero-ℕ () contradiction-le-one-ℕ (succ-ℕ n) () ``` @@ -134,15 +134,15 @@ irreflexive-le-ℕ (succ-ℕ n) = irreflexive-le-ℕ n ### If `x < y` then `x ≠ y` ```agda -neq-le-ℕ : {x y : ℕ} → le-ℕ x y → x ≠ y -neq-le-ℕ {zero-ℕ} {succ-ℕ y} H = is-nonzero-succ-ℕ y ∘ inv -neq-le-ℕ {succ-ℕ x} {succ-ℕ y} H p = neq-le-ℕ H (is-injective-succ-ℕ p) +neq-le-ℕ : {m n : ℕ} → m <-ℕ n → m ≠ n +neq-le-ℕ {zero-ℕ} {succ-ℕ n} H = is-nonzero-succ-ℕ n ∘ inv +neq-le-ℕ {succ-ℕ m} {succ-ℕ n} H p = neq-le-ℕ H (is-injective-succ-ℕ p) ``` ### The strict inequality on the natural numbers is antisymmetric ```agda -antisymmetric-le-ℕ : (m n : ℕ) → le-ℕ m n → le-ℕ n m → m = n +antisymmetric-le-ℕ : (m n : ℕ) → m <-ℕ n → n <-ℕ m → m = n antisymmetric-le-ℕ (succ-ℕ m) (succ-ℕ n) p q = ap succ-ℕ (antisymmetric-le-ℕ m n p q) ``` @@ -150,7 +150,7 @@ antisymmetric-le-ℕ (succ-ℕ m) (succ-ℕ n) p q = ### The strict inequality on the natural numbers is transitive ```agda -transitive-le-ℕ : (n m l : ℕ) → (le-ℕ m l) → (le-ℕ n m) → (le-ℕ n l) +transitive-le-ℕ : (n m l : ℕ) → m <-ℕ l → n <-ℕ m → n <-ℕ l transitive-le-ℕ zero-ℕ (succ-ℕ m) (succ-ℕ l) p q = star transitive-le-ℕ (succ-ℕ n) (succ-ℕ m) (succ-ℕ l) p q = transitive-le-ℕ n m l p q @@ -165,7 +165,7 @@ concatenate-le-eq-le-ℕ a b .b d H refl K = ```agda transitive-le-ℕ' : - (k l m : ℕ) → (le-ℕ k l) → (le-ℕ l (succ-ℕ m)) → le-ℕ k m + (k l m : ℕ) → k <-ℕ l → l <-ℕ succ-ℕ m → k <-ℕ m transitive-le-ℕ' zero-ℕ zero-ℕ m () s transitive-le-ℕ' (succ-ℕ k) zero-ℕ m () s transitive-le-ℕ' zero-ℕ (succ-ℕ l) zero-ℕ star s = @@ -180,7 +180,7 @@ transitive-le-ℕ' (succ-ℕ k) (succ-ℕ l) (succ-ℕ m) t s = ### The strict inequality on the natural numbers is linear ```agda -linear-le-ℕ : (x y : ℕ) → (le-ℕ x y) + ((x = y) + (le-ℕ y x)) +linear-le-ℕ : (x y : ℕ) → x <-ℕ y + ((x = y) + y <-ℕ x) linear-le-ℕ zero-ℕ zero-ℕ = inr (inl refl) linear-le-ℕ zero-ℕ (succ-ℕ y) = inl star linear-le-ℕ (succ-ℕ x) zero-ℕ = inr (inr star) @@ -188,6 +188,18 @@ linear-le-ℕ (succ-ℕ x) (succ-ℕ y) = map-coproduct id (map-coproduct (ap succ-ℕ) id) (linear-le-ℕ x y) ``` +### Given two unequal natural numbers, we can decide their strict ordering + +```agda +decide-le-neq-ℕ : + (m n : ℕ) → m ≠ n → (m <-ℕ n) + (n <-ℕ m) +decide-le-neq-ℕ zero-ℕ zero-ℕ H = ex-falso (H refl) +decide-le-neq-ℕ zero-ℕ (succ-ℕ n) H = inl star +decide-le-neq-ℕ (succ-ℕ m) zero-ℕ H = inr star +decide-le-neq-ℕ (succ-ℕ m) (succ-ℕ n) H = + decide-le-neq-ℕ m n (λ p → H (ap succ-ℕ p)) +``` + ### The strictly ordered type of natural numbers ```agda @@ -202,15 +214,17 @@ pr2 (pr2 (pr2 ℕ-Strictly-Ordered-Type)) = transitive-le-ℕ ```agda subtraction-le-ℕ : - (n m : ℕ) → le-ℕ n m → Σ ℕ (λ l → (is-nonzero-ℕ l) × (l +ℕ n = m)) -subtraction-le-ℕ zero-ℕ m p = pair m (pair (is-nonzero-le-ℕ zero-ℕ m p) refl) + (n m : ℕ) → n <-ℕ m → Σ ℕ (λ l → is-nonzero-ℕ l × (l +ℕ n = m)) +subtraction-le-ℕ zero-ℕ m p = + ( m , is-nonzero-le-ℕ zero-ℕ m p , refl) subtraction-le-ℕ (succ-ℕ n) (succ-ℕ m) p = - pair (pr1 P) (pair (pr1 (pr2 P)) (ap succ-ℕ (pr2 (pr2 P)))) + ( pr1 P , pr1 (pr2 P) , ap succ-ℕ (pr2 (pr2 P))) where + P : Σ ℕ (λ l' → (is-nonzero-ℕ l') × (l' +ℕ n = m)) P = subtraction-le-ℕ n m p -le-subtraction-ℕ : (n m l : ℕ) → is-nonzero-ℕ l → l +ℕ n = m → le-ℕ n m +le-subtraction-ℕ : (n m l : ℕ) → is-nonzero-ℕ l → l +ℕ n = m → n <-ℕ m le-subtraction-ℕ zero-ℕ m l q p = tr (λ x → le-ℕ zero-ℕ x) p (le-zero-is-nonzero-ℕ l q) le-subtraction-ℕ (succ-ℕ n) (succ-ℕ m) l q p = @@ -220,7 +234,7 @@ le-subtraction-ℕ (succ-ℕ n) (succ-ℕ m) l q p = ### Any natural number is strictly less than its successor ```agda -succ-le-ℕ : (n : ℕ) → le-ℕ n (succ-ℕ n) +succ-le-ℕ : (n : ℕ) → n <-ℕ succ-ℕ n succ-le-ℕ zero-ℕ = star succ-le-ℕ (succ-ℕ n) = succ-le-ℕ n ``` @@ -229,7 +243,7 @@ succ-le-ℕ (succ-ℕ n) = succ-le-ℕ n ```agda preserves-le-succ-ℕ : - (m n : ℕ) → le-ℕ m n → le-ℕ m (succ-ℕ n) + (m n : ℕ) → m <-ℕ n → m <-ℕ succ-ℕ n preserves-le-succ-ℕ m n H = transitive-le-ℕ m n (succ-ℕ n) (succ-le-ℕ n) H ``` @@ -238,14 +252,14 @@ preserves-le-succ-ℕ m n H = ```agda concatenate-leq-le-ℕ : - (x y z : ℕ) → x ≤-ℕ y → le-ℕ y z → le-ℕ x z + (x y z : ℕ) → x ≤-ℕ y → y <-ℕ z → x <-ℕ z concatenate-leq-le-ℕ zero-ℕ zero-ℕ (succ-ℕ z) H K = star concatenate-leq-le-ℕ zero-ℕ (succ-ℕ y) (succ-ℕ z) H K = star concatenate-leq-le-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) H K = concatenate-leq-le-ℕ x y z H K concatenate-le-leq-ℕ : - (x y z : ℕ) → le-ℕ x y → y ≤-ℕ z → le-ℕ x z + (x y z : ℕ) → x <-ℕ y → y ≤-ℕ z → x <-ℕ z concatenate-le-leq-ℕ zero-ℕ (succ-ℕ y) (succ-ℕ z) H K = star concatenate-le-leq-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) H K = concatenate-le-leq-ℕ x y z H K @@ -254,7 +268,7 @@ concatenate-le-leq-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) H K = ### If `m < n` then `n ≰ m` ```agda -contradiction-le-ℕ : (m n : ℕ) → le-ℕ m n → ¬ (n ≤-ℕ m) +contradiction-le-ℕ : (m n : ℕ) → m <-ℕ n → ¬ (n ≤-ℕ m) contradiction-le-ℕ zero-ℕ (succ-ℕ n) H K = K contradiction-le-ℕ (succ-ℕ m) (succ-ℕ n) H = contradiction-le-ℕ m n H ``` @@ -262,14 +276,14 @@ contradiction-le-ℕ (succ-ℕ m) (succ-ℕ n) H = contradiction-le-ℕ m n H ### If `n ≤ m` then `m ≮ n` ```agda -contradiction-le-ℕ' : (m n : ℕ) → n ≤-ℕ m → ¬ (le-ℕ m n) +contradiction-le-ℕ' : (m n : ℕ) → n ≤-ℕ m → ¬ (m <-ℕ n) contradiction-le-ℕ' m n K H = contradiction-le-ℕ m n H K ``` ### If `m ≮ n` then `n ≤ m` ```agda -leq-not-le-ℕ : (m n : ℕ) → ¬ (le-ℕ m n) → n ≤-ℕ m +leq-not-le-ℕ : (m n : ℕ) → ¬ (m <-ℕ n) → n ≤-ℕ m leq-not-le-ℕ zero-ℕ zero-ℕ H = star leq-not-le-ℕ zero-ℕ (succ-ℕ n) H = ex-falso (H star) leq-not-le-ℕ (succ-ℕ m) zero-ℕ H = star @@ -290,7 +304,7 @@ le-not-leq-ℕ (succ-ℕ m) (succ-ℕ n) H = le-not-leq-ℕ m n H ```agda leq-le-ℕ : - (x y : ℕ) → le-ℕ x y → x ≤-ℕ y + (x y : ℕ) → x <-ℕ y → x ≤-ℕ y leq-le-ℕ zero-ℕ (succ-ℕ y) H = star leq-le-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-le-ℕ x y H ``` @@ -299,7 +313,7 @@ leq-le-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-le-ℕ x y H ```agda leq-le-succ-ℕ : - (x y : ℕ) → le-ℕ x (succ-ℕ y) → x ≤-ℕ y + (x y : ℕ) → x <-ℕ succ-ℕ y → x ≤-ℕ y leq-le-succ-ℕ zero-ℕ y H = star leq-le-succ-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-le-succ-ℕ x y H ``` @@ -308,7 +322,7 @@ leq-le-succ-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-le-succ-ℕ x y H ```agda leq-succ-le-ℕ : - (x y : ℕ) → le-ℕ x y → leq-ℕ (succ-ℕ x) y + (x y : ℕ) → x <-ℕ y → succ-ℕ x ≤-ℕ y leq-succ-le-ℕ zero-ℕ (succ-ℕ y) H = star leq-succ-le-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-succ-le-ℕ x y H ``` @@ -317,7 +331,7 @@ leq-succ-le-ℕ (succ-ℕ x) (succ-ℕ y) H = leq-succ-le-ℕ x y H ```agda le-succ-leq-ℕ : - (x y : ℕ) → leq-ℕ x y → le-ℕ x (succ-ℕ y) + (x y : ℕ) → x ≤-ℕ y → x <-ℕ succ-ℕ y le-succ-leq-ℕ zero-ℕ zero-ℕ H = star le-succ-leq-ℕ zero-ℕ (succ-ℕ y) H = star le-succ-leq-ℕ (succ-ℕ x) (succ-ℕ y) H = le-succ-leq-ℕ x y H @@ -327,18 +341,18 @@ le-succ-leq-ℕ (succ-ℕ x) (succ-ℕ y) H = le-succ-leq-ℕ x y H ```agda eq-or-le-leq-ℕ : - (x y : ℕ) → leq-ℕ x y → ((x = y) + (le-ℕ x y)) + (x y : ℕ) → x ≤-ℕ y → ((x = y) + (x <-ℕ y)) eq-or-le-leq-ℕ zero-ℕ zero-ℕ H = inl refl eq-or-le-leq-ℕ zero-ℕ (succ-ℕ y) H = inr star eq-or-le-leq-ℕ (succ-ℕ x) (succ-ℕ y) H = map-coproduct (ap succ-ℕ) id (eq-or-le-leq-ℕ x y H) eq-or-le-leq-ℕ' : - (x y : ℕ) → leq-ℕ x y → ((y = x) + (le-ℕ x y)) + (x y : ℕ) → x ≤-ℕ y → ((y = x) + (x <-ℕ y)) eq-or-le-leq-ℕ' x y H = map-coproduct inv id (eq-or-le-leq-ℕ x y H) leq-eq-or-le-ℕ : - (x y : ℕ) → ((x = y) + (le-ℕ x y)) → leq-ℕ x y + (x y : ℕ) → ((x = y) + (x <-ℕ y)) → x ≤-ℕ y leq-eq-or-le-ℕ x .x (inl refl) = refl-leq-ℕ x leq-eq-or-le-ℕ x y (inr l) = leq-le-ℕ x y l ``` @@ -346,7 +360,7 @@ leq-eq-or-le-ℕ x y (inr l) = leq-le-ℕ x y l ### If `x ≤ y` and `x ≠ y` then `x < y` ```agda -le-leq-neq-ℕ : {x y : ℕ} → x ≤-ℕ y → x ≠ y → le-ℕ x y +le-leq-neq-ℕ : {x y : ℕ} → x ≤-ℕ y → x ≠ y → x <-ℕ y le-leq-neq-ℕ {zero-ℕ} {zero-ℕ} l f = ex-falso (f refl) le-leq-neq-ℕ {zero-ℕ} {succ-ℕ y} l f = star le-leq-neq-ℕ {succ-ℕ x} {succ-ℕ y} l f = diff --git a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md index 9e7dc01d97..24af1afb01 100644 --- a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md @@ -87,15 +87,9 @@ is-section-fin-strictly-bounded-ℕ : is-section-fin-strictly-bounded-ℕ (succ-ℕ n) i = eq-type-subtype ( λ x → le-ℕ-Prop x (succ-ℕ n)) - ( eq-cong-le-ℕ - ( succ-ℕ n) - ( nat-Fin (succ-ℕ n) (fin-strictly-bounded-ℕ (succ-ℕ n) i)) + ( eq-nat-mod-succ-ℕ n ( nat-strictly-bounded-ℕ (succ-ℕ n) i) - ( strict-upper-bound-nat-Fin - ( succ-ℕ n) - ( fin-strictly-bounded-ℕ (succ-ℕ n) i)) - ( strict-upper-bound-nat-strictly-bounded-ℕ (succ-ℕ n) i) - ( cong-nat-mod-succ-ℕ n (nat-strictly-bounded-ℕ (succ-ℕ n) i))) + ( strict-upper-bound-nat-strictly-bounded-ℕ (succ-ℕ n) i)) is-retraction-fin-strictly-bounded-ℕ : (n : ℕ) → diff --git a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md index b2562fc4ee..68497fe7c1 100644 --- a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md @@ -8,6 +8,7 @@ module elementary-number-theory.sums-of-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -28,6 +29,7 @@ open import foundation.whiskering-homotopies-composition open import lists.lists open import univalent-combinatorics.counting +open import univalent-combinatorics.skipping-element-standard-finite-types open import univalent-combinatorics.standard-finite-types ``` @@ -170,3 +172,37 @@ leq-sum-Fin-ℕ (succ-ℕ k) f (inl x) = leq-sum-Fin-ℕ (succ-ℕ k) f (inr x) = leq-add-ℕ' (f (inr x)) (sum-Fin-ℕ k (f ∘ inl)) ``` + +### The difference between a summand and the sum of natural numbers + +```agda +sum-skip-Fin-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → ℕ +sum-skip-Fin-ℕ (succ-ℕ k) f i = + sum-Fin-ℕ k (f ∘ skip-Fin k i) + +eq-sum-skip-Fin-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → + sum-skip-Fin-ℕ k f i +ℕ f i = sum-Fin-ℕ k f +eq-sum-skip-Fin-ℕ (succ-ℕ zero-ℕ) f (inr star) = + refl +eq-sum-skip-Fin-ℕ (succ-ℕ (succ-ℕ k)) f (inl i) = + ( right-swap-add-ℕ + ( sum-Fin-ℕ k (f ∘ inl ∘ skip-Fin k i)) + ( f (inr star)) + ( f (inl i))) ∙ + ( ap (_+ℕ f (inr star)) (eq-sum-skip-Fin-ℕ (succ-ℕ k) (f ∘ inl) i)) +eq-sum-skip-Fin-ℕ (succ-ℕ (succ-ℕ k)) f (inr star) = + refl + +compute-dist-summand-sum-Fin-ℕ : + (k : ℕ) (f : Fin k → ℕ) (i : Fin k) → + dist-ℕ (f i) (sum-Fin-ℕ k f) = sum-skip-Fin-ℕ k f i +compute-dist-summand-sum-Fin-ℕ k f i = + inv + ( rewrite-left-add-dist-ℕ + ( sum-skip-Fin-ℕ k f i) + ( f i) + ( sum-Fin-ℕ k f) + ( eq-sum-skip-Fin-ℕ k f i)) +``` From a686f18d2e6038b4385f1d690a6241b4554e07a7 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 15 Jan 2025 17:43:15 -0500 Subject: [PATCH 50/72] Farey fractions --- .../2-adic-decomposition.lagda.md | 10 +- .../exponentiation-natural-numbers.lagda.md | 149 ++++++++++++++++-- .../farey-fractions.lagda.md | 80 ++++++++++ .../fermat-numbers.lagda.md | 10 +- .../sums-of-natural-numbers.lagda.md | 3 +- 5 files changed, 233 insertions(+), 19 deletions(-) create mode 100644 src/elementary-number-theory/farey-fractions.lagda.md diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index cc6d24cae8..c836a630e4 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -46,6 +46,8 @@ $$ (m , n) \mapsto 2^m(2n + 1) - 1. $$ +The exponent $k$ such that the 2-adic decomposition of $n$ is $2^k \cdot m=n$ is called the {{#concept "2-adic valuation" Disambiguation="natural numbers" Agda=valuation-2-adic-decomposition-ℕ}} of $n$. + ## Definitions ### The $2$-adic composition function @@ -66,15 +68,15 @@ module _ (n : ℕ) (d : 2-adic-decomposition-ℕ n) where - exponent-2-adic-decomposition-ℕ : ℕ - exponent-2-adic-decomposition-ℕ = pr1 d + valuation-2-adic-decomposition-ℕ : ℕ + valuation-2-adic-decomposition-ℕ = pr1 d index-odd-factor-2-adic-decomposition-ℕ : ℕ index-odd-factor-2-adic-decomposition-ℕ = pr1 (pr2 d) eq-2-adic-decomposition-ℕ : 2-adic-composition-ℕ - exponent-2-adic-decomposition-ℕ + valuation-2-adic-decomposition-ℕ index-odd-factor-2-adic-decomposition-ℕ = n eq-2-adic-decomposition-ℕ = pr2 (pr2 d) @@ -147,7 +149,7 @@ module _ ℕ exponent-even-case-2-adic-decomposition-is-even-or-odd-ℕ = succ-ℕ - ( exponent-2-adic-decomposition-ℕ + ( valuation-2-adic-decomposition-ℕ ( q) ( 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ)) diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index 4272144754..6ea46dc614 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -13,6 +13,7 @@ open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.commutative-semiring-of-natural-numbers open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -21,8 +22,14 @@ open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types +open import foundation.embeddings +open import foundation.empty-types open import foundation.identity-types +open import foundation.injective-maps open import foundation.unit-type + +open import order-theory.order-preserving-maps-posets +open import order-theory.strict-order-preserving-maps ```
@@ -59,22 +66,22 @@ annihilation-law-exp-ℕ : (n : ℕ) → 1 ^ℕ n = 1 annihilation-law-exp-ℕ = power-one-Commutative-Semiring ℕ-Commutative-Semiring ``` -### `xⁿ⁺¹ = xⁿx` +### `mⁿ⁺¹ = mⁿm` ```agda exp-succ-ℕ : - (n : ℕ) (x : ℕ) → x ^ℕ succ-ℕ n = x ^ℕ n *ℕ x -exp-succ-ℕ = - power-succ-Commutative-Semiring ℕ-Commutative-Semiring + (m n : ℕ) → m ^ℕ succ-ℕ n = m ^ℕ n *ℕ m +exp-succ-ℕ m n = + power-succ-Commutative-Semiring ℕ-Commutative-Semiring n m ``` -### `xⁿ⁺¹ = xxⁿ` +### `mⁿ⁺¹ = mmⁿ` ```agda exp-succ-ℕ' : - (n : ℕ) (x : ℕ) → x ^ℕ succ-ℕ n = x *ℕ x ^ℕ n -exp-succ-ℕ' = - power-succ-Commutative-Semiring' ℕ-Commutative-Semiring + (m n : ℕ) → m ^ℕ succ-ℕ n = m *ℕ m ^ℕ n +exp-succ-ℕ' m n = + power-succ-Commutative-Semiring' ℕ-Commutative-Semiring n m ``` ### Powers by sums of natural numbers are products of powers @@ -138,6 +145,28 @@ leq-one-exp-ℕ m n H = ( is-nonzero-exp-ℕ m n (is-nonzero-leq-one-ℕ m H)) ``` +### The exponent $m^n$ of a number $m>1$ by a nonzero number $n$ is strictly greater than $0$ + +```agda +le-one-exp-ℕ : + (m n : ℕ) → 1 <-ℕ m → is-nonzero-ℕ n → 1 <-ℕ m ^ℕ n +le-one-exp-ℕ (succ-ℕ (succ-ℕ m)) zero-ℕ H K = ex-falso (K refl) +le-one-exp-ℕ (succ-ℕ (succ-ℕ m)) (succ-ℕ zero-ℕ) H K = star +le-one-exp-ℕ (succ-ℕ (succ-ℕ m)) (succ-ℕ (succ-ℕ n)) H K = + concatenate-le-leq-ℕ 1 + ( succ-ℕ (succ-ℕ m) ^ℕ succ-ℕ n) + ( succ-ℕ (succ-ℕ m) ^ℕ succ-ℕ (succ-ℕ n)) + ( le-one-exp-ℕ (succ-ℕ (succ-ℕ m)) (succ-ℕ n) H (is-nonzero-succ-ℕ n)) + ( concatenate-eq-leq-eq-ℕ + ( inv (right-unit-law-mul-ℕ (succ-ℕ (succ-ℕ m) ^ℕ succ-ℕ n))) + ( preserves-order-right-mul-ℕ + ( succ-ℕ (succ-ℕ m) ^ℕ succ-ℕ n) + ( 1) + ( succ-ℕ (succ-ℕ m)) + ( star)) + ( exp-succ-ℕ (succ-ℕ (succ-ℕ m)) (succ-ℕ n))) +``` + ### The exponent $m^n$ is equal to the $n$-fold product of $m$ ```agda @@ -155,7 +184,7 @@ compute-constant-product-ℕ m (succ-ℕ (succ-ℕ n)) = div-exp-succ-ℕ : (m n : ℕ) → div-ℕ m (m ^ℕ succ-ℕ n) pr1 (div-exp-succ-ℕ m n) = m ^ℕ n -pr2 (div-exp-succ-ℕ m n) = inv (exp-succ-ℕ n m) +pr2 (div-exp-succ-ℕ m n) = inv (exp-succ-ℕ m n) div-exp-is-successor-ℕ : (m n : ℕ) → is-successor-ℕ n → div-ℕ m (m ^ℕ n) @@ -172,3 +201,105 @@ bounded-div-exp-ℕ : bounded-div-exp-ℕ m n H = bounded-div-div-ℕ m (m ^ℕ n) (div-exp-ℕ m n H) ``` + +### If an exponential of a number greater than $1$ is equal to $1$, then its exponent is $0$ + +```agda +is-zero-exponent-is-one-exp-ℕ : + (m n : ℕ) → 1 <-ℕ m → is-one-ℕ (m ^ℕ n) → is-zero-ℕ n +is-zero-exponent-is-one-exp-ℕ m zero-ℕ H K = + refl +is-zero-exponent-is-one-exp-ℕ m (succ-ℕ n) H p = + ex-falso + ( neq-le-ℕ + ( concatenate-leq-le-ℕ 1 + ( m ^ℕ n) + ( m ^ℕ n *ℕ m) + ( leq-one-exp-ℕ m n (leq-le-ℕ 1 m H)) + ( concatenate-eq-le-ℕ + ( m ^ℕ n) + ( m ^ℕ n *ℕ 1) + ( m ^ℕ n *ℕ m) + ( inv (right-unit-law-mul-ℕ (m ^ℕ n))) + ( preserves-strict-order-left-mul-ℕ + ( m ^ℕ n) + ( 1) + ( m) + ( is-nonzero-exp-ℕ m n (is-nonzero-le-ℕ 1 m H)) H))) + ( inv p ∙ exp-succ-ℕ m n)) +``` + +### The function $n \mapsto m^n$ is injective and an embedding for any $m>1$ + +```agda +is-injective-exp-ℕ : + (m : ℕ) → 1 <-ℕ m → is-injective (m ^ℕ_) +is-injective-exp-ℕ m H {zero-ℕ} {zero-ℕ} p = + refl +is-injective-exp-ℕ m H {zero-ℕ} {succ-ℕ l} p = + inv (is-zero-exponent-is-one-exp-ℕ m (succ-ℕ l) H (inv p)) +is-injective-exp-ℕ m H {succ-ℕ k} {zero-ℕ} p = + is-zero-exponent-is-one-exp-ℕ m (succ-ℕ k) H p +is-injective-exp-ℕ m H {succ-ℕ k} {succ-ℕ l} p = + ap + ( succ-ℕ) + ( is-injective-exp-ℕ m H {k} {l} + ( is-injective-right-mul-ℕ m + ( is-nonzero-le-ℕ 1 m H) + ( inv (exp-succ-ℕ m k) ∙ p ∙ exp-succ-ℕ m l))) + +is-emb-exp-ℕ : + (m : ℕ) → 1 <-ℕ m → is-emb (m ^ℕ_) +is-emb-exp-ℕ m H = + is-emb-is-injective is-set-ℕ (is-injective-exp-ℕ m H) +``` + +### Exponentiating a nonzero number is an order preserving function + +```agda +preserves-order-exponent-exp-ℕ : + (m : ℕ) → is-nonzero-ℕ m → + preserves-order-Poset ℕ-Poset ℕ-Poset (m ^ℕ_) +preserves-order-exponent-exp-ℕ m H zero-ℕ zero-ℕ K = + refl-leq-ℕ 1 +preserves-order-exponent-exp-ℕ m H zero-ℕ (succ-ℕ k) K = + leq-one-exp-ℕ m (succ-ℕ k) (leq-one-is-nonzero-ℕ m H) +preserves-order-exponent-exp-ℕ m H (succ-ℕ n) (succ-ℕ k) K = + concatenate-eq-leq-eq-ℕ + ( exp-succ-ℕ m n) + ( preserves-order-mul-ℕ + ( m ^ℕ n) + ( m ^ℕ k) + ( m) + ( m) + ( preserves-order-exponent-exp-ℕ m H n k K) + ( refl-leq-ℕ m)) + ( inv (exp-succ-ℕ m k)) +``` + +### Exponentiating a number strictly greater than $1$ is a strictly order preserving function + +```agda +preserves-strict-order-exponent-exp-ℕ : + (m : ℕ) → 1 <-ℕ m → + preserves-strict-order-map-Strictly-Ordered-Type + ( ℕ-Strictly-Ordered-Type) + ( ℕ-Strictly-Ordered-Type) + ( m ^ℕ_) +preserves-strict-order-exponent-exp-ℕ m H zero-ℕ (succ-ℕ k) K = + le-one-exp-ℕ m (succ-ℕ k) H (is-nonzero-succ-ℕ k) +preserves-strict-order-exponent-exp-ℕ m H (succ-ℕ n) (succ-ℕ k) K = + concatenate-eq-le-eq-ℕ + ( m ^ℕ succ-ℕ n) + ( m ^ℕ n *ℕ m) + ( m ^ℕ k *ℕ m) + ( m ^ℕ succ-ℕ k) + ( exp-succ-ℕ m n) + ( preserves-strict-order-right-mul-ℕ + ( m) + ( m ^ℕ n) + ( m ^ℕ k) + ( is-nonzero-le-ℕ 1 m H) + ( preserves-strict-order-exponent-exp-ℕ m H n k K)) + ( inv (exp-succ-ℕ m k)) +``` diff --git a/src/elementary-number-theory/farey-fractions.lagda.md b/src/elementary-number-theory/farey-fractions.lagda.md new file mode 100644 index 0000000000..4649d950de --- /dev/null +++ b/src/elementary-number-theory/farey-fractions.lagda.md @@ -0,0 +1,80 @@ +# Farey fractions + +```agda +module elementary-number-theory.farey-fractions where +``` + +
Imports + +```agda +open import elementary-number-theory.integer-fractions + +open import foundation.universe-levels +``` + +
+ +## Idea + +A {{#concept "Farey fraction" Agda=farey-fraction}} is a [reduced](elementary-number-theory.reduced-integer-fractions.md) [integer fraction](elementary-number-theory.integer-fractions.md) between $0$ and $1$ inclusive. More specifically, a Farey fraction of order $n$ is a reduced integer fraction between $0$ and $1$ whose denominator does not exceed $n$. + +The Farey fractions ℱ can be inductively generated mutually with a binary relation $R$ with the following constructors: + +```text + 0 : ℱ + 1 : ℱ + 𝓂 : (x y : ℱ) → ℛ x y → ℱ + + 𝓈 : ℛ 0 1 + 𝓇 : (x y : ℱ) (r : ℛ x y) → ℛ x (𝓂 x y r) + 𝓁 : (x y : ℱ) (r : ℛ x y) → ℛ (𝓂 x y r) y +``` + +The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions"}} of two adjacent Farey fractions. + +Farey fractions appear in Chapter 3 of {{#cite HW08}}, but they are covered in more detail in Chapter 6 of {{#cite NZM}}. + +## Definitions + +### The inductively generated Farey fractions + +```agda +mutual + + data + farey-fraction : UU lzero + where + + zero-farey-fraction : farey-fraction + one-farey-fraction : farey-fraction + + mediant-farey-fraction : + (x y : farey-fraction) → adjacent-farey-fraction x y → farey-fraction + + data + adjacent-farey-fraction : farey-fraction → farey-fraction → UU lzero + where + + adjacent-zero-one-farey-fraction : + adjacent-farey-fraction zero-farey-fraction one-farey-fraction + + right-adjacent-mediant-farey-fraction : + (x y : farey-fraction) (H : adjacent-farey-fraction x y) → + adjacent-farey-fraction x (mediant-farey-fraction x y H) + + left-adjacent-mediant-farey-fraction : + (x y : farey-fraction) (H : adjacent-farey-fraction x y) → + adjacent-farey-fraction (mediant-farey-fraction x y H) y +``` + +### The inclusion of Farey fractions into the integer fractions + +```agda +integer-fraction-farey-fraction : + farey-fraction → fraction-ℤ +integer-fraction-farey-fraction = ? +``` + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index 32adc85b44..45eabce82a 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -237,11 +237,11 @@ is-one-div-fermat-number-le-ℕ m (succ-ℕ n) d H K L = is-one-is-common-divisor-fermat-number-ℕ : (m n d : ℕ) → m ≠ n → is-common-divisor-ℕ (fermat-number-ℕ m) (fermat-number-ℕ n) d → is-one-ℕ d -is-one-is-common-divisor-fermat-number-ℕ m n d H (K , L) - with - decide-le-neq-ℕ m n H -... | inl u = is-one-div-fermat-number-le-ℕ m n d u K L -... | inr u = is-one-div-fermat-number-le-ℕ n m d u L K +is-one-is-common-divisor-fermat-number-ℕ m n d H (K , L) = + rec-coproduct + ( λ u → is-one-div-fermat-number-le-ℕ m n d u K L) + ( λ u → is-one-div-fermat-number-le-ℕ n m d u L K) + ( decide-le-neq-ℕ m n H) is-relatively-prime-fermat-number-ℕ : (m n : ℕ) → m ≠ n → diff --git a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md index 68497fe7c1..341fd6a89f 100644 --- a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md @@ -99,7 +99,8 @@ numbers up to, and including, a given upper bound. ```agda bounded-sum-ℕ : (N : ℕ) → ((i : ℕ) → leq-ℕ i N → ℕ) → ℕ -bounded-sum-ℕ zero-ℕ f = zero-ℕ +bounded-sum-ℕ zero-ℕ f = + f 0 (refl-leq-ℕ 0) bounded-sum-ℕ (succ-ℕ N) f = add-ℕ ( bounded-sum-ℕ N (λ x H → f x (leq-succ-leq-ℕ x N H))) From fa5ced321a8503ce2d5dcbd3be01d2aee9369a12 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 15 Jan 2025 17:53:26 -0500 Subject: [PATCH 51/72] unbounded Farey fractions --- .../farey-fractions.lagda.md | 49 ++++++++++++++++++- 1 file changed, 47 insertions(+), 2 deletions(-) diff --git a/src/elementary-number-theory/farey-fractions.lagda.md b/src/elementary-number-theory/farey-fractions.lagda.md index 4649d950de..aebb910897 100644 --- a/src/elementary-number-theory/farey-fractions.lagda.md +++ b/src/elementary-number-theory/farey-fractions.lagda.md @@ -7,7 +7,8 @@ module elementary-number-theory.farey-fractions where
Imports ```agda -open import elementary-number-theory.integer-fractions +-- open import elementary-number-theory.integer-fractions +open import elementary-number-theory.integers open import foundation.universe-levels ``` @@ -67,9 +68,53 @@ mutual adjacent-farey-fraction (mediant-farey-fraction x y H) y ``` -### The inclusion of Farey fractions into the integer fractions +### Unbounded Farey fractions + +Egbert conjectures that the type of unbounded Farey fractions is equivalent to the type of rational numbers in a canonical way. ```agda +mutual + + data + unbounded-farey-fraction : UU lzero + where + + farey-integer : + ℤ → unbounded-farey-fraction + + mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) → + adjacent-unbounded-farey-fraction x y → unbounded-farey-fraction + + data + adjacent-unbounded-farey-fraction : + unbounded-farey-fraction → unbounded-farey-fraction → UU lzero + where + + adjacent-farey-integer-succ : + (a : ℤ) → + adjacent-unbounded-farey-fraction + ( farey-integer a) + ( farey-integer (succ-ℤ a)) + + right-adjacent-mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) + (H : adjacent-unbounded-farey-fraction x y) → + adjacent-unbounded-farey-fraction + ( x) + ( mediant-unbounded-farey-fraction x y H) + + left-adjacent-mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) + (H : adjacent-unbounded-farey-fraction x y) → + adjacent-unbounded-farey-fraction + ( mediant-unbounded-farey-fraction x y H) + ( y) +``` + +### The inclusion of Farey fractions into the integer fractions + +```text integer-fraction-farey-fraction : farey-fraction → fraction-ℤ integer-fraction-farey-fraction = ? From 24c32f7c24d0ca25a8d1e95712c667ee099b5bbb Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Thu, 16 Jan 2025 14:22:04 -0500 Subject: [PATCH 52/72] some properties of Farey fractions --- .../2-adic-decomposition.lagda.md | 2 +- .../exponentiation-natural-numbers.lagda.md | 105 +++++++++++++- .../farey-fractions.lagda.md | 132 ++++++++++++------ ...strict-inequality-natural-numbers.lagda.md | 6 +- ...tionary-functions-natural-numbers.lagda.md | 20 ++- .../unbounded-farey-fractions.lagda.md | 74 ++++++++++ .../standard-finite-types.lagda.md | 14 +- 7 files changed, 297 insertions(+), 56 deletions(-) create mode 100644 src/elementary-number-theory/unbounded-farey-fractions.lagda.md diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index c836a630e4..7638b6edc7 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -129,7 +129,7 @@ $$ 2^{k+1}(2l + 1) = 2(2^k(2l + 1)) = 2*y = x. $$ -```agda +```text module _ (x : ℕ) (H : 1 ≤-ℕ x) (f : diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index 6ea46dc614..2f56b58d34 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -15,18 +15,25 @@ open import elementary-number-theory.commutative-semiring-of-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.strictly-inflationary-functions-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.products-of-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import elementary-number-theory.strong-induction-natural-numbers open import foundation.action-on-identifications-functions +open import foundation.coproduct-types +open import foundation.decidable-maps open import foundation.dependent-pair-types open import foundation.embeddings open import foundation.empty-types +open import foundation.fibers-of-maps open import foundation.identity-types open import foundation.injective-maps +open import foundation.negated-equality open import foundation.unit-type +open import foundation.universe-levels open import order-theory.order-preserving-maps-posets open import order-theory.strict-order-preserving-maps @@ -38,18 +45,41 @@ open import order-theory.strict-order-preserving-maps The {{#concept "exponential" Agda=exp-ℕ WD="exponentiation" WDID=Q33456}} $m^n$ of two [natural numbers](elementary-number-theory.natural-numbers.md) is the -number obtained by multiplying $m$ with itself $n$ times. +number obtained by multiplying $m$ with itself $n$ times. We use the following terminology in our formalization and naming scheme: + +- The number $m^n$ is the **$n$th power** of $m$. +- The number $m^n$ is an **exponential**. +- The number $n$ in $m^n$ is called the **exponent** of the exponential. +- The number $m$ in $m^n$ is the **base** of the exponential. +- The operation $m,n \mapsto m^n$ is called **exponentiation**. +- The function $n \mapsto m^n$ is the **exponentiation function** with base $m$. +- The function $m \mapsto m^n$ is the **$n$th power function**. Specific instances are [squaring](elementary-number-theory.squares-natural-numbers.md) and [cubing](elementary-number-theory.cubes-natural-numbers.md). The natural numbers satisfy Tarski's high school arithmetic laws for exponentiation. ## Definition -### Exponentiation of natural numbers +### Powers of natural numbers + +The function `power-ℕ n : ℕ → ℕ` defines $n$th power of a natural number $m$. ```agda power-ℕ : ℕ → ℕ → ℕ power-ℕ = power-Commutative-Semiring ℕ-Commutative-Semiring +``` +### The predicate of being an $n$th power + +```agda +is-power-ℕ : ℕ → ℕ → UU lzero +is-power-ℕ m n = fiber (power-ℕ n) m +``` + +### Exponentiation of natural numbers + +The function `exp-ℕ : ℕ → ℕ → ℕ` defines the exponentiation map $m,n \mapsto m^n$. Note that this is just the power function, with the order of its arguments swapped. + +```agda exp-ℕ : ℕ → ℕ → ℕ exp-ℕ m n = power-ℕ n m @@ -57,6 +87,13 @@ infixr 45 _^ℕ_ _^ℕ_ = exp-ℕ ``` +### The predicate of being an exponential of a given natural number + +```agda +is-exponential-ℕ : ℕ → ℕ → UU lzero +is-exponential-ℕ m n = fiber (exp-ℕ m) n +``` + ## Properties ### `1ⁿ = 1` @@ -212,6 +249,8 @@ is-zero-exponent-is-one-exp-ℕ m zero-ℕ H K = is-zero-exponent-is-one-exp-ℕ m (succ-ℕ n) H p = ex-falso ( neq-le-ℕ + ( 1) + ( m ^ℕ n *ℕ m) ( concatenate-leq-le-ℕ 1 ( m ^ℕ n) ( m ^ℕ n *ℕ m) @@ -303,3 +342,65 @@ preserves-strict-order-exponent-exp-ℕ m H (succ-ℕ n) (succ-ℕ k) K = ( preserves-strict-order-exponent-exp-ℕ m H n k K)) ( inv (exp-succ-ℕ m k)) ``` + +### The function $n \mapsto m^n$ is strictly inflationary for any $1Imports ```agda --- open import elementary-number-theory.integer-fractions -open import elementary-number-theory.integers +open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.multiplication-natural-numbers +open import elementary-number-theory.natural-numbers +open import foundation.action-on-identifications-functions +open import foundation.identity-types open import foundation.universe-levels ``` @@ -31,7 +34,13 @@ The Farey fractions ℱ can be inductively generated mutually with a binary rela 𝓁 : (x y : ℱ) (r : ℛ x y) → ℛ (𝓂 x y r) y ``` -The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions"}} of two adjacent Farey fractions. +The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions"}} of two adjacent Farey fractions. The elements $0$ and $1$ in the type of Farey fractions represent the Farey fractions $0/1$ and $1/1$. Given two adjacent Farey fractions representing $a/b$ and $c/d$, the mediant of $a/b$ and $c/d$ is the Farey fraction representing + +$$ + \frac{a+c}{b+d}. +$$ + +The mediant of any two adjacent Farey fractions representing reduced fractions $a/b$ and $c/d$ represents again a reduced fraction. Farey fractions appear in Chapter 3 of {{#cite HW08}}, but they are covered in more detail in Chapter 6 of {{#cite NZM}}. @@ -68,57 +77,92 @@ mutual adjacent-farey-fraction (mediant-farey-fraction x y H) y ``` -### Unbounded Farey fractions - -Egbert conjectures that the type of unbounded Farey fractions is equivalent to the type of rational numbers in a canonical way. +### The inclusion of Farey fractions into the integer fractions ```agda -mutual +numerator-farey-fraction : + farey-fraction → ℕ +numerator-farey-fraction zero-farey-fraction = + 0 +numerator-farey-fraction one-farey-fraction = + 1 +numerator-farey-fraction (mediant-farey-fraction x y H) = + numerator-farey-fraction x +ℕ numerator-farey-fraction y + +denominator-farey-fraction : + farey-fraction → ℕ +denominator-farey-fraction zero-farey-fraction = + 1 +denominator-farey-fraction one-farey-fraction = + 1 +denominator-farey-fraction (mediant-farey-fraction x y H) = + denominator-farey-fraction x +ℕ denominator-farey-fraction y +``` - data - unbounded-farey-fraction : UU lzero - where +## Properties - farey-integer : - ℤ → unbounded-farey-fraction +### Any two adjacent Farey fractions $a/b$ and $c/d$ satisfy $bc = ad + 1$ - mediant-unbounded-farey-fraction : - (x y : unbounded-farey-fraction) → - adjacent-unbounded-farey-fraction x y → unbounded-farey-fraction +Equivalently, two adjacent Farey fractions $a/b$ and $c/d$ satisfy the relation - data - adjacent-unbounded-farey-fraction : - unbounded-farey-fraction → unbounded-farey-fraction → UU lzero - where +$$ + bc - ad = 1. +$$ - adjacent-farey-integer-succ : - (a : ℤ) → - adjacent-unbounded-farey-fraction - ( farey-integer a) - ( farey-integer (succ-ℤ a)) - - right-adjacent-mediant-unbounded-farey-fraction : - (x y : unbounded-farey-fraction) - (H : adjacent-unbounded-farey-fraction x y) → - adjacent-unbounded-farey-fraction - ( x) - ( mediant-unbounded-farey-fraction x y H) - - left-adjacent-mediant-unbounded-farey-fraction : - (x y : unbounded-farey-fraction) - (H : adjacent-unbounded-farey-fraction x y) → - adjacent-unbounded-farey-fraction - ( mediant-unbounded-farey-fraction x y H) - ( y) +This is also known as the **characteristic property of adjacent Farey fractions**. + +```agda +characteristic-property-adjacent-farey-fraction : + (x y : farey-fraction) (r : adjacent-farey-fraction x y) → + denominator-farey-fraction x *ℕ numerator-farey-fraction y = + numerator-farey-fraction x *ℕ denominator-farey-fraction y +ℕ 1 +characteristic-property-adjacent-farey-fraction ._ ._ + adjacent-zero-one-farey-fraction = + refl +characteristic-property-adjacent-farey-fraction x ._ + ( right-adjacent-mediant-farey-fraction .x y r) = + left-distributive-mul-add-ℕ + ( denominator-farey-fraction x) + ( numerator-farey-fraction x) + ( numerator-farey-fraction y) ∙ + ap-add-ℕ + ( commutative-mul-ℕ + ( denominator-farey-fraction x) + ( numerator-farey-fraction x)) + ( characteristic-property-adjacent-farey-fraction x y r) ∙ + ap + ( succ-ℕ) + ( inv + ( left-distributive-mul-add-ℕ + ( numerator-farey-fraction x) + ( denominator-farey-fraction x) + ( denominator-farey-fraction y))) +characteristic-property-adjacent-farey-fraction ._ y + ( left-adjacent-mediant-farey-fraction x .y r) = + right-distributive-mul-add-ℕ + ( denominator-farey-fraction x) + ( denominator-farey-fraction y) + ( numerator-farey-fraction y) ∙ + ap-add-ℕ + ( characteristic-property-adjacent-farey-fraction x y r) + ( commutative-mul-ℕ + ( denominator-farey-fraction y) + ( numerator-farey-fraction y)) ∙ + left-successor-law-add-ℕ + ( numerator-farey-fraction x *ℕ denominator-farey-fraction y) + ( numerator-farey-fraction y *ℕ denominator-farey-fraction y) ∙ + ap + ( succ-ℕ) + ( inv + ( right-distributive-mul-add-ℕ + ( numerator-farey-fraction x) + ( numerator-farey-fraction y) + ( denominator-farey-fraction y))) ``` -### The inclusion of Farey fractions into the integer fractions +## See also -```text -integer-fraction-farey-fraction : - farey-fraction → fraction-ℤ -integer-fraction-farey-fraction = ? -``` +- [Unbounded Farey fractions](elementary-number-theory.unbounded-farey-fractions.md) ## References diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 91c094166d..67b768bb97 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -134,9 +134,9 @@ irreflexive-le-ℕ (succ-ℕ n) = irreflexive-le-ℕ n ### If `x < y` then `x ≠ y` ```agda -neq-le-ℕ : {m n : ℕ} → m <-ℕ n → m ≠ n -neq-le-ℕ {zero-ℕ} {succ-ℕ n} H = is-nonzero-succ-ℕ n ∘ inv -neq-le-ℕ {succ-ℕ m} {succ-ℕ n} H p = neq-le-ℕ H (is-injective-succ-ℕ p) +neq-le-ℕ : (m n : ℕ) → m <-ℕ n → m ≠ n +neq-le-ℕ zero-ℕ (succ-ℕ n) H = is-nonzero-succ-ℕ n ∘ inv +neq-le-ℕ (succ-ℕ m) (succ-ℕ n) H p = neq-le-ℕ m n H (is-injective-succ-ℕ p) ``` ### The strict inequality on the natural numbers is antisymmetric diff --git a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md index b504d9cbcd..89facf988b 100644 --- a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md @@ -9,16 +9,22 @@ module
Imports ```agda +open import elementary-number-theory.decidable-types +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers +open import foundation.decidable-maps open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.fibers-of-maps +open import foundation.identity-types open import foundation.universe-levels open import order-theory.inflationary-maps-strictly-ordered-types + +open import univalent-combinatorics.dependent-pair-types ```
@@ -80,14 +86,20 @@ is-inflationary-strictly-inflationary-map-ℕ f n = ( is-strictly-inflationary-strictly-inflationary-map-ℕ f n) ``` -### Being a value of a strictly inflationary map is decidable +### Strictly inflationary maps are decidable ```agda is-decidable-fiber-strictly-inflationary-map-ℕ : (f : strictly-inflationary-map-ℕ) → - is-decidable-fam (fiber-strictly-inflationary-map-ℕ f) -is-decidable-fiber-strictly-inflationary-map-ℕ = {!!} + is-decidable-map (map-strictly-inflationary-map-ℕ f) +is-decidable-fiber-strictly-inflationary-map-ℕ (f , H) n = + is-decidable-iff + ( λ (x , K , p) → (x , p)) + ( λ (x , p) → (x , concatenate-le-eq-ℕ x (f x) n (H x) p , p)) + ( is-decidable-strictly-bounded-Σ-ℕ' n + ( λ x → f x = n) + ( λ x → has-decidable-equality-ℕ (f x) n)) ``` -### For any strictly inflationary map $f$ with $f(0) \leq b$, there is a maximal number $k$ such that $f(k) \leq b$ + diff --git a/src/elementary-number-theory/unbounded-farey-fractions.lagda.md b/src/elementary-number-theory/unbounded-farey-fractions.lagda.md new file mode 100644 index 0000000000..be7be51e73 --- /dev/null +++ b/src/elementary-number-theory/unbounded-farey-fractions.lagda.md @@ -0,0 +1,74 @@ +# Unbounded Farey fractions + +```agda +module elementary-number-theory.unbounded-farey-fractions where +``` + +
Imports + +```agda +open import elementary-number-theory.integers + +open import foundation.universe-levels +``` + +
+ +## Idea + +[Farey fractions](elementary-number-theory.farey-fractions.md) are a way of encoding all [rational numbers](elementary-number-theory.rational-numbers.md) between $0$ and $1$ inclusive. This idea can be generalized to encode all rational numbers. + +The type of {{#concept "unbounded Farey fractions" Agda=unbounded-farey-fraction}} is an inductive type that is mutually defined with an adjacency relation, given with the following constructors: + +```text + 𝒾 : ℤ → ℱ + 𝓂 : (x y : ℱ) → ℛ x y → ℱ + + 𝓈 : (a : ℤ) → ℛ (𝒾 a) (𝒾 (a + 1)) + 𝓇 : (x y : ℱ) (r : ℛ x y) → ℛ x (𝓂 x y r) + 𝓁 : (x y : ℱ) (r : ℛ x y) → ℛ (𝓂 x y r) y. +``` + +## Definitions + +### Unbounded Farey fractions + +```agda +mutual + + data + unbounded-farey-fraction : UU lzero + where + + farey-integer : + ℤ → unbounded-farey-fraction + + mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) → + adjacent-unbounded-farey-fraction x y → unbounded-farey-fraction + + data + adjacent-unbounded-farey-fraction : + unbounded-farey-fraction → unbounded-farey-fraction → UU lzero + where + + adjacent-farey-integer-succ : + (a : ℤ) → + adjacent-unbounded-farey-fraction + ( farey-integer a) + ( farey-integer (succ-ℤ a)) + + right-adjacent-mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) + (H : adjacent-unbounded-farey-fraction x y) → + adjacent-unbounded-farey-fraction + ( x) + ( mediant-unbounded-farey-fraction x y H) + + left-adjacent-mediant-unbounded-farey-fraction : + (x y : unbounded-farey-fraction) + (H : adjacent-unbounded-farey-fraction x y) → + adjacent-unbounded-farey-fraction + ( mediant-unbounded-farey-fraction x y H) + ( y) +``` diff --git a/src/univalent-combinatorics/standard-finite-types.lagda.md b/src/univalent-combinatorics/standard-finite-types.lagda.md index 88d54a1cf5..94b1231c5a 100644 --- a/src/univalent-combinatorics/standard-finite-types.lagda.md +++ b/src/univalent-combinatorics/standard-finite-types.lagda.md @@ -258,9 +258,19 @@ is-injective-nat-Fin : (k : ℕ) → is-injective (nat-Fin k) is-injective-nat-Fin (succ-ℕ k) {inl x} {inl y} p = ap inl (is-injective-nat-Fin k p) is-injective-nat-Fin (succ-ℕ k) {inl x} {inr star} p = - ex-falso (neq-le-ℕ (strict-upper-bound-nat-Fin k x) p) + ex-falso + ( neq-le-ℕ + ( nat-Fin k x) + ( nat-Fin (succ-ℕ k) (inr star)) + ( strict-upper-bound-nat-Fin k x) + ( p)) is-injective-nat-Fin (succ-ℕ k) {inr star} {inl y} p = - ex-falso (neq-le-ℕ (strict-upper-bound-nat-Fin k y) (inv p)) + ex-falso + ( neq-le-ℕ + ( nat-Fin (succ-ℕ k) (inl y)) + ( nat-Fin (succ-ℕ k) (inr star)) + ( strict-upper-bound-nat-Fin k y) + ( inv p)) is-injective-nat-Fin (succ-ℕ k) {inr star} {inr star} p = refl From ae95d3c2996a5f52aef2f8c5e29c54ea377d41c9 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 17 Jan 2025 16:34:06 -0500 Subject: [PATCH 53/72] Every positive natural number has a 2-adic decomposition --- .../2-adic-decomposition.lagda.md | 150 ++++++++++-------- .../divisibility-natural-numbers.lagda.md | 7 + .../exponentiation-natural-numbers.lagda.md | 81 +++++++--- .../farey-fractions.lagda.md | 25 ++- ...st-power-divisors-natural-numbers.lagda.md | 95 +++++++++++ .../rational-numbers.lagda.md | 5 + ...strict-inequality-natural-numbers.lagda.md | 25 +++ ...tionary-functions-natural-numbers.lagda.md | 19 ++- .../decidable-propositions.lagda.md | 37 +++-- .../finite-maps.lagda.md | 18 +++ 10 files changed, 355 insertions(+), 107 deletions(-) create mode 100644 src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 7638b6edc7..7b1b4f429b 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -12,6 +12,7 @@ open import elementary-number-theory.based-strong-induction-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.largest-power-divisors-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.parity-natural-numbers @@ -24,6 +25,7 @@ open import foundation.function-types open import foundation.functoriality-dependent-pair-types open import foundation.split-surjective-maps open import foundation.transport-along-identifications +open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types @@ -74,6 +76,10 @@ module _ index-odd-factor-2-adic-decomposition-ℕ : ℕ index-odd-factor-2-adic-decomposition-ℕ = pr1 (pr2 d) + odd-factor-2-adic-decomposition-ℕ : ℕ + odd-factor-2-adic-decomposition-ℕ = + odd-number-ℕ index-odd-factor-2-adic-decomposition-ℕ + eq-2-adic-decomposition-ℕ : 2-adic-composition-ℕ valuation-2-adic-decomposition-ℕ @@ -121,78 +127,92 @@ pr2 (pr2 (2-adic-decomposition-is-odd-ℕ n H)) = ### Every nonzero natural number has a $2$-adic decomposition -The proof is by [based strong induction](elementary-number-theory.based-strong-induction-natural-numbers.md). -We have already seen that every odd natural number has a $2$-adic expansion. -If $1 \leq x$ is an even number, then write $x = 2 \cdot y$. Given a $2$-adic decomposition $(k , l)$ of the number $y$, it follows that $(k + 1 , l)$ is a $2$-adic decomposition of $x$, because - -$$ - 2^{k+1}(2l + 1) = 2(2^k(2l + 1)) = 2*y = x. -$$ - -```text +```agda module _ - (x : ℕ) (H : 1 ≤-ℕ x) - (f : - based-□-≤-ℕ 1 (λ y → is-even-ℕ y + is-odd-ℕ y → 2-adic-decomposition-ℕ y) x) - (K@(q , p) : is-even-ℕ (succ-ℕ x)) + (n : ℕ) (H : 1 ≤-ℕ n) where + + valuation-2-adic-decomposition-nat-ℕ : + ℕ + valuation-2-adic-decomposition-nat-ℕ = + valuation-largest-power-divisor-ℕ 2 n star H + + div-exp-valuation-2-adic-decomposition-nat-ℕ : + div-ℕ (2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) n + div-exp-valuation-2-adic-decomposition-nat-ℕ = + div-exp-valuation-largest-power-divisor-ℕ 2 n star H - 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ : - 2-adic-decomposition-ℕ q - 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ = - f ( q) - ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) - ( upper-bound-quotient-is-even-succ-ℕ x K) - ( is-decidable-is-even-ℕ q) + is-upper-bound-valuation-2-adic-decomposition-nat-ℕ : + (k : ℕ) → div-ℕ (2 ^ℕ k) n → k ≤-ℕ valuation-2-adic-decomposition-nat-ℕ + is-upper-bound-valuation-2-adic-decomposition-nat-ℕ = + is-upper-bound-valuation-largest-power-divisor-ℕ 2 n star H - exponent-even-case-2-adic-decomposition-is-even-or-odd-ℕ : + odd-factor-2-adic-decomposition-nat-ℕ : ℕ - exponent-even-case-2-adic-decomposition-is-even-or-odd-ℕ = - succ-ℕ - ( valuation-2-adic-decomposition-ℕ - ( q) - ( 2-adic-decomposition-quotient-even-case-2-adic-decomposition-is-even-or-odd-ℕ)) - - even-case-2-adic-decomposition-is-even-or-odd-ℕ : - 2-adic-decomposition-ℕ (succ-ℕ x) - even-case-2-adic-decomposition-is-even-or-odd-ℕ = - ( succ-ℕ (pr1 d) , pr1 (pr2 d) , {!!}) - where - - d = - f ( q) - ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) - {!!} - {!!} - -2-adic-decomposition-is-even-or-odd-ℕ : - (n : ℕ) → 1 ≤-ℕ n → is-even-ℕ n + is-odd-ℕ n → - 2-adic-decomposition-ℕ n -2-adic-decomposition-is-even-or-odd-ℕ = - based-strong-ind-ℕ 1 - ( λ x → (is-even-ℕ x + is-odd-ℕ x) → 2-adic-decomposition-ℕ x) - ( rec-coproduct - ( λ H → ex-falso (is-odd-one-ℕ H)) - ( 2-adic-decomposition-is-odd-ℕ 1)) - ( λ x H f → - rec-coproduct - ( even-case-2-adic-decomposition-is-even-or-odd-ℕ x H f - {- - λ K@(q , p) → - let - - L : 2-adic-decomposition-ℕ q -d L = - f ( q) - ( leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) - {! upper-bound-quotient-div-ℕ 2 (succ-ℕ x) K!} - {!!} - in - {!f q (leq-one-quotient-div-ℕ 2 (succ-ℕ x) K (leq-zero-ℕ x)) ? ?!} - -}) - ( 2-adic-decomposition-is-odd-ℕ (succ-ℕ x))) + odd-factor-2-adic-decomposition-nat-ℕ = + quotient-div-ℕ + ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) + ( n) + ( div-exp-valuation-2-adic-decomposition-nat-ℕ) + + is-odd-odd-factor-2-adic-decomposition-nat-ℕ : + is-odd-ℕ odd-factor-2-adic-decomposition-nat-ℕ + is-odd-odd-factor-2-adic-decomposition-nat-ℕ K = + neg-succ-leq-ℕ + ( valuation-2-adic-decomposition-nat-ℕ) + ( is-upper-bound-valuation-2-adic-decomposition-nat-ℕ + ( succ-ℕ valuation-2-adic-decomposition-nat-ℕ) + ( tr + ( is-divisor-ℕ n) + ( inv (exp-succ-ℕ 2 valuation-2-adic-decomposition-nat-ℕ)) + ( div-div-quotient-div-ℕ 2 n + ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) + ( div-exp-valuation-2-adic-decomposition-nat-ℕ) + ( K)))) + + has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ : + has-odd-expansion-ℕ odd-factor-2-adic-decomposition-nat-ℕ + has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ = + has-odd-expansion-is-odd-ℕ + odd-factor-2-adic-decomposition-nat-ℕ + is-odd-odd-factor-2-adic-decomposition-nat-ℕ + + index-odd-factor-2-adic-decomposition-nat-ℕ : + ℕ + index-odd-factor-2-adic-decomposition-nat-ℕ = + pr1 has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ + + eq-index-odd-factor-2-adic-decomposition-nat-ℕ : + odd-number-ℕ index-odd-factor-2-adic-decomposition-nat-ℕ = + odd-factor-2-adic-decomposition-nat-ℕ + eq-index-odd-factor-2-adic-decomposition-nat-ℕ = + pr2 has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ + + eq-2-adic-decomposition-nat-ℕ : + 2-adic-composition-ℕ + ( valuation-2-adic-decomposition-nat-ℕ) + ( index-odd-factor-2-adic-decomposition-nat-ℕ) = + n + eq-2-adic-decomposition-nat-ℕ = + ( ap + ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ *ℕ_) + ( eq-index-odd-factor-2-adic-decomposition-nat-ℕ)) ∙ + ( eq-quotient-div-ℕ' + ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) + ( n) + ( div-exp-valuation-2-adic-decomposition-nat-ℕ)) + + 2-adic-decomposition-nat-ℕ : + 2-adic-decomposition-ℕ n + pr1 2-adic-decomposition-nat-ℕ = + valuation-2-adic-decomposition-nat-ℕ + pr1 (pr2 2-adic-decomposition-nat-ℕ) = + index-odd-factor-2-adic-decomposition-nat-ℕ + pr2 (pr2 2-adic-decomposition-nat-ℕ) = + eq-2-adic-decomposition-nat-ℕ ``` + ```text pair-expansion : ℕ → UU lzero pair-expansion n = diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index ef046f0ad2..15b02fca03 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -111,6 +111,13 @@ module _ pr2 (div-bounded-div-ℕ H) = eq-quotient-bounded-div-ℕ m n H ``` +### The predicate of being a divisor of a natural number + +```agda +is-divisor-ℕ : ℕ → ℕ → UU lzero +is-divisor-ℕ n x = div-ℕ x n +``` + ## Properties ### If `n` is divisible by a number `m` with proof of divisibility `(q , p)`, then `n` is divisible by the number `q`. diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index 2f56b58d34..5e04359f90 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -24,6 +24,7 @@ open import elementary-number-theory.strong-induction-natural-numbers open import foundation.action-on-identifications-functions open import foundation.coproduct-types +open import foundation.decidable-embeddings open import foundation.decidable-maps open import foundation.dependent-pair-types open import foundation.embeddings @@ -37,6 +38,8 @@ open import foundation.universe-levels open import order-theory.order-preserving-maps-posets open import order-theory.strict-order-preserving-maps + +open import univalent-combinatorics.finite-maps ```
@@ -347,14 +350,40 @@ preserves-strict-order-exponent-exp-ℕ m H (succ-ℕ n) (succ-ℕ k) K = ```agda is-strictly-inflationary-exp-ℕ : - (m : ℕ) → 1 <-ℕ m → is-strictly-inflationary-ℕ (m ^ℕ_) + (m : ℕ) → 1 <-ℕ m → is-strictly-inflationary-map-ℕ (m ^ℕ_) is-strictly-inflationary-exp-ℕ m H zero-ℕ = star -is-strictly-inflationary-exp-ℕ m H (succ-ℕ n) = {!!} +is-strictly-inflationary-exp-ℕ m H (succ-ℕ zero-ℕ) = H +is-strictly-inflationary-exp-ℕ m H (succ-ℕ (succ-ℕ n)) = + concatenate-leq-le-leq-ℕ + ( succ-ℕ (succ-ℕ n)) + ( succ-ℕ n *ℕ 2) + ( m ^ℕ succ-ℕ n *ℕ 2) + ( m ^ℕ succ-ℕ (succ-ℕ n)) + ( concatenate-leq-eq-ℕ + ( n +ℕ 2) + ( preserves-order-add-ℕ + { succ-ℕ n} + { succ-ℕ n *ℕ 1} + { 1} + { succ-ℕ n} + ( leq-eq-ℕ + ( succ-ℕ n) + ( succ-ℕ n *ℕ 1) + ( inv (right-unit-law-mul-ℕ (succ-ℕ n)))) + ( leq-zero-ℕ n)) + ( inv (right-successor-law-mul-ℕ (succ-ℕ n) 1))) + ( preserves-strict-order-right-mul-ℕ 2 + ( succ-ℕ n) + ( m ^ℕ succ-ℕ n) + ( is-nonzero-succ-ℕ 1) + ( is-strictly-inflationary-exp-ℕ m H (succ-ℕ n))) + ( concatenate-leq-eq-ℕ + ( m ^ℕ succ-ℕ n *ℕ 2) + ( preserves-order-right-mul-ℕ (m ^ℕ succ-ℕ n) 2 m (leq-succ-le-ℕ 1 m H)) + ( inv (exp-succ-ℕ m (succ-ℕ n)))) ``` -n + 1 < m^n + 1 ≤ m^{n+1} - -### The exponential function $n \mapsto m^n$ if a decidable function for any $m$ +### The exponential function $n \mapsto m^n$ if a decidable function ```agda leq-one-exp-zero-ℕ : @@ -386,21 +415,29 @@ is-decidable-map-exp-one-ℕ n = is-decidable-map-exp-ℕ : (m : ℕ) → is-decidable-map (exp-ℕ m) -is-decidable-map-exp-ℕ zero-ℕ zero-ℕ = - inl (1 , refl) -is-decidable-map-exp-ℕ (succ-ℕ m) zero-ℕ = - inr - ( λ (k , p) → - neq-le-ℕ 0 - ( succ-ℕ m ^ℕ k) - ( le-zero-exp-ℕ (succ-ℕ m) k star) - ( inv p)) -is-decidable-map-exp-ℕ zero-ℕ (succ-ℕ zero-ℕ) = - inl (0 , refl) -is-decidable-map-exp-ℕ zero-ℕ (succ-ℕ (succ-ℕ n)) = - inr - ( λ (k , p) → - concatenate-eq-leq-ℕ 1 (inv p) (leq-one-exp-zero-ℕ k)) -is-decidable-map-exp-ℕ (succ-ℕ m) (succ-ℕ n) = - {!!} +is-decidable-map-exp-ℕ zero-ℕ = + is-decidable-map-exp-zero-ℕ +is-decidable-map-exp-ℕ (succ-ℕ zero-ℕ) = + is-decidable-map-exp-one-ℕ +is-decidable-map-exp-ℕ (succ-ℕ (succ-ℕ m)) = + is-decidable-map-is-strictly-inflationary-map-ℕ + ( is-strictly-inflationary-exp-ℕ (succ-ℕ (succ-ℕ m)) star) +``` + +### The exponential function is a decidable embedding if its base is strictly greater than $1$ + +```agda +is-decidable-emb-exp-ℕ : + (m : ℕ) → 1 <-ℕ m → is-decidable-emb (exp-ℕ m) +pr1 (is-decidable-emb-exp-ℕ m H) = is-emb-exp-ℕ m H +pr2 (is-decidable-emb-exp-ℕ m H) = is-decidable-map-exp-ℕ m +``` + +### The exponential function is a finite map if its base is strictly greater than $1$ + +```agda +is-finite-map-exp-ℕ : + (m : ℕ) → 1 <-ℕ m → is-finite-map (exp-ℕ m) +is-finite-map-exp-ℕ m H = + is-finite-map-is-decidable-emb (is-decidable-emb-exp-ℕ m H) ``` diff --git a/src/elementary-number-theory/farey-fractions.lagda.md b/src/elementary-number-theory/farey-fractions.lagda.md index 600c6b5bda..4b754b7cd3 100644 --- a/src/elementary-number-theory/farey-fractions.lagda.md +++ b/src/elementary-number-theory/farey-fractions.lagda.md @@ -12,7 +12,10 @@ open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.empty-types open import foundation.identity-types +open import foundation.negation open import foundation.universe-levels ``` @@ -34,7 +37,7 @@ The Farey fractions ℱ can be inductively generated mutually with a binary rela 𝓁 : (x y : ℱ) (r : ℛ x y) → ℛ (𝓂 x y r) y ``` -The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions"}} of two adjacent Farey fractions. The elements $0$ and $1$ in the type of Farey fractions represent the Farey fractions $0/1$ and $1/1$. Given two adjacent Farey fractions representing $a/b$ and $c/d$, the mediant of $a/b$ and $c/d$ is the Farey fraction representing +The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions" Agda=mediant-farey-fraction}} of two adjacent Farey fractions. The elements $0$ and $1$ in the type of Farey fractions represent the Farey fractions $0/1$ and $1/1$. Given two adjacent Farey fractions representing $a/b$ and $c/d$, the mediant of $a/b$ and $c/d$ is the Farey fraction representing $$ \frac{a+c}{b+d}. @@ -160,6 +163,26 @@ characteristic-property-adjacent-farey-fraction ._ y ( denominator-farey-fraction y))) ``` +### There is no adjacency from any Farey fraction to $0$ + +```agda +not-adjacent-zero-farey-fraction : + (x : farey-fraction) → ¬ adjacent-farey-fraction x zero-farey-fraction +not-adjacent-zero-farey-fraction ._ + ( left-adjacent-mediant-farey-fraction x ._ r) = + not-adjacent-zero-farey-fraction x r +``` + +### There is no adjacency from $1$ to any Farey fraction + +```agda +not-adjacent-one-farey-fraction : + (x : farey-fraction) → ¬ adjacent-farey-fraction one-farey-fraction x +not-adjacent-one-farey-fraction + ( mediant-farey-fraction one-farey-fraction y s) r = + not-adjacent-one-farey-fraction y s +``` + ## See also - [Unbounded Farey fractions](elementary-number-theory.unbounded-farey-fractions.md) diff --git a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md new file mode 100644 index 0000000000..555c2eba46 --- /dev/null +++ b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md @@ -0,0 +1,95 @@ +# Largest power divisors of natural numbers + +```agda +module elementary-number-theory.largest-power-divisors-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.exponentiation-natural-numbers +open import elementary-number-theory.finite-maps-natural-numbers +open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.maximal-structured-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.strict-inequality-natural-numbers +open import elementary-number-theory.upper-bounds-natural-numbers + +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.universe-levels +``` + +
+ +## Idea + +The {{#concept "largest power divisor" Disambiguation="natural numbers" Agda=largest-power-divisor-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) $n$ with base $m$ is the [largest number](elementary-number-theory.maximal-structured-natural-numbers.md) of the form $m^k$ that [divides](elementary-number-theory.divisibility-natural-numbers.md) $n$. + +## Definitions + +### Power divisors of natural numbers + +```agda +is-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero +is-power-divisor-ℕ m n x = is-exponential-ℕ m x × div-ℕ x n +``` + +### The largest power divisor of a natural number + +```agda +module _ + (m n : ℕ) (H : 1 <-ℕ m) (K : 1 ≤-ℕ n) + where + + largest-power-divisor-ℕ : + maximal-element-ℕ + ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) + largest-power-divisor-ℕ = + maximal-structured-value-bound-input-property-is-finite-map-ℕ + ( is-divisor-ℕ n) + ( λ x → is-decidable-div-ℕ x n) + ( m ^ℕ_) + ( is-finite-map-exp-ℕ m H) + ( n) + ( 0) + ( K) + ( div-one-ℕ n) + + valuation-largest-power-divisor-ℕ : + ℕ + valuation-largest-power-divisor-ℕ = + nat-maximal-element-ℕ + ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) + ( largest-power-divisor-ℕ) + + upper-bound-exp-valuation-largest-power-divisor-ℕ : + m ^ℕ valuation-largest-power-divisor-ℕ ≤-ℕ n + upper-bound-exp-valuation-largest-power-divisor-ℕ = + pr1 + ( structure-maximal-element-ℕ + ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) + ( largest-power-divisor-ℕ)) + + div-exp-valuation-largest-power-divisor-ℕ : + div-ℕ (m ^ℕ valuation-largest-power-divisor-ℕ) n + div-exp-valuation-largest-power-divisor-ℕ = + pr2 + ( structure-maximal-element-ℕ + ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) + ( largest-power-divisor-ℕ)) + + is-upper-bound-valuation-largest-power-divisor-ℕ : + (k : ℕ) → div-ℕ (m ^ℕ k) n → k ≤-ℕ valuation-largest-power-divisor-ℕ + is-upper-bound-valuation-largest-power-divisor-ℕ k L = + is-upper-bound-maximal-element-ℕ + ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) + ( largest-power-divisor-ℕ) + ( k) + ( leq-div-ℕ (m ^ℕ k) n (is-nonzero-leq-one-ℕ n K) L , L) +``` + +## See also + +- [2-adic decomposition](elementary-number-theory.2-adic-decomposition.md) diff --git a/src/elementary-number-theory/rational-numbers.lagda.md b/src/elementary-number-theory/rational-numbers.lagda.md index f4b630199a..b72c2a284b 100644 --- a/src/elementary-number-theory/rational-numbers.lagda.md +++ b/src/elementary-number-theory/rational-numbers.lagda.md @@ -309,6 +309,11 @@ pr1 reflecting-map-sim-fraction = rational-fraction-ℤ pr2 reflecting-map-sim-fraction {x} {y} H = eq-ℚ-sim-fraction-ℤ x y H ``` +## See also + +- [Farey fractions](elementary-number-theory.farey-fractions.md) encode the type of rational numbers between $0$ and $1$ by a mutual recursive construction. +- [Unbounded Farey fractions](elementary-number-theory.unbounded-farey-fractions.md) encode the type of all rational numbers by a mutal recursive construction. + ## References {{#bibliography}} diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 67b768bb97..9654af34fd 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -263,6 +263,11 @@ concatenate-le-leq-ℕ : concatenate-le-leq-ℕ zero-ℕ (succ-ℕ y) (succ-ℕ z) H K = star concatenate-le-leq-ℕ (succ-ℕ x) (succ-ℕ y) (succ-ℕ z) H K = concatenate-le-leq-ℕ x y z H K + +concatenate-leq-le-leq-ℕ : + (w x y z : ℕ) → w ≤-ℕ x → x <-ℕ y → y ≤-ℕ z → w <-ℕ z +concatenate-leq-le-leq-ℕ w x y z H K L = + concatenate-leq-le-ℕ w x z H (concatenate-le-leq-ℕ x y z K L) ``` ### If `m < n` then `n ≰ m` @@ -557,6 +562,26 @@ preserves-strict-order-right-mul-ℕ k x y H = ( is-successor-is-nonzero-ℕ H) ``` +### Multiplying with a nonzero element is an inflationary function + +```agda +is-inflationary-left-mul-ℕ : + (k n : ℕ) → is-nonzero-ℕ k → n ≤-ℕ k *ℕ n +is-inflationary-left-mul-ℕ k n H = + concatenate-eq-leq-ℕ + ( k *ℕ n) + ( inv (left-unit-law-mul-ℕ n)) + ( preserves-order-left-mul-ℕ n 1 k (leq-one-is-nonzero-ℕ k H)) + +is-inflationary-right-mul-ℕ : + (k n : ℕ) → is-nonzero-ℕ k → n ≤-ℕ n *ℕ k +is-inflationary-right-mul-ℕ k n H = + concatenate-leq-eq-ℕ + ( n) + ( is-inflationary-left-mul-ℕ k n H) + ( commutative-mul-ℕ k n) +``` + ## See also - [The strictly bounded natural numbers](elementary-number-theory.strictly-bounded-natural-numbers.md) diff --git a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md index 89facf988b..421cb742c5 100644 --- a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md @@ -44,8 +44,8 @@ holds for every $x,y:\mathbb{N}$. If $f$ is a strictly inflationary function on ### The predicate of being a strictly inflationary map on the natural numbers ```agda -is-strictly-inflationary-ℕ : (ℕ → ℕ) → UU lzero -is-strictly-inflationary-ℕ = +is-strictly-inflationary-map-ℕ : (ℕ → ℕ) → UU lzero +is-strictly-inflationary-map-ℕ = is-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type ``` @@ -63,7 +63,7 @@ map-strictly-inflationary-map-ℕ = is-strictly-inflationary-strictly-inflationary-map-ℕ : (f : strictly-inflationary-map-ℕ) → - is-strictly-inflationary-ℕ (map-strictly-inflationary-map-ℕ f) + is-strictly-inflationary-map-ℕ (map-strictly-inflationary-map-ℕ f) is-strictly-inflationary-strictly-inflationary-map-ℕ = is-inflationary-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type @@ -89,16 +89,21 @@ is-inflationary-strictly-inflationary-map-ℕ f n = ### Strictly inflationary maps are decidable ```agda -is-decidable-fiber-strictly-inflationary-map-ℕ : - (f : strictly-inflationary-map-ℕ) → - is-decidable-map (map-strictly-inflationary-map-ℕ f) -is-decidable-fiber-strictly-inflationary-map-ℕ (f , H) n = +is-decidable-map-is-strictly-inflationary-map-ℕ : + {f : ℕ → ℕ} → is-strictly-inflationary-map-ℕ f → is-decidable-map f +is-decidable-map-is-strictly-inflationary-map-ℕ {f} H n = is-decidable-iff ( λ (x , K , p) → (x , p)) ( λ (x , p) → (x , concatenate-le-eq-ℕ x (f x) n (H x) p , p)) ( is-decidable-strictly-bounded-Σ-ℕ' n ( λ x → f x = n) ( λ x → has-decidable-equality-ℕ (f x) n)) + +is-decidable-map-strictly-inflationary-map-ℕ : + (f : strictly-inflationary-map-ℕ) → + is-decidable-map (map-strictly-inflationary-map-ℕ f) +is-decidable-map-strictly-inflationary-map-ℕ (f , H) n = + is-decidable-map-is-strictly-inflationary-map-ℕ H n ``` diff --git a/src/foundation/decidable-propositions.lagda.md b/src/foundation/decidable-propositions.lagda.md index c6c5bfb814..375603bd7c 100644 --- a/src/foundation/decidable-propositions.lagda.md +++ b/src/foundation/decidable-propositions.lagda.md @@ -218,18 +218,30 @@ abstract ### Decidable propositions have a count ```agda +count-is-decidable-prop : + {l : Level} {X : UU l} → is-decidable-prop X → count X +count-is-decidable-prop (H , inl x) = + count-is-contr (is-proof-irrelevant-is-prop H x) +count-is-decidable-prop (H , inr f) = + count-is-empty f + count-is-decidable-Prop : - {l : Level} (P : Prop l) → - is-decidable (type-Prop P) → count (type-Prop P) -count-is-decidable-Prop P (inl x) = - count-is-contr (is-proof-irrelevant-is-prop (is-prop-type-Prop P) x) -count-is-decidable-Prop P (inr x) = - count-is-empty x + {l : Level} (P : Prop l) → + is-decidable (type-Prop P) → count (type-Prop P) +count-is-decidable-Prop P d = + count-is-decidable-prop (is-prop-type-Prop P , d) ``` ### Decidable propositions are finite ```agda +abstract + is-finite-is-decidable-prop : + {l : Level} {X : UU l} → + is-decidable-prop X → is-finite X + is-finite-is-decidable-prop H = + is-finite-count (count-is-decidable-prop H) + abstract is-finite-is-decidable-Prop : {l : Level} (P : Prop l) → @@ -237,12 +249,13 @@ abstract is-finite-is-decidable-Prop P x = is-finite-count (count-is-decidable-Prop P x) -is-finite-type-Decidable-Prop : - {l : Level} (P : Decidable-Prop l) → is-finite (type-Decidable-Prop P) -is-finite-type-Decidable-Prop P = - is-finite-is-decidable-Prop - ( prop-Decidable-Prop P) - ( is-decidable-Decidable-Prop P) +abstract + is-finite-type-Decidable-Prop : + {l : Level} (P : Decidable-Prop l) → is-finite (type-Decidable-Prop P) + is-finite-type-Decidable-Prop P = + is-finite-is-decidable-Prop + ( prop-Decidable-Prop P) + ( is-decidable-Decidable-Prop P) ``` ### The type of decidable propositions of any universe level is finite diff --git a/src/univalent-combinatorics/finite-maps.lagda.md b/src/univalent-combinatorics/finite-maps.lagda.md index 1b77f62b1e..1af4bb61e5 100644 --- a/src/univalent-combinatorics/finite-maps.lagda.md +++ b/src/univalent-combinatorics/finite-maps.lagda.md @@ -9,6 +9,9 @@ module univalent-combinatorics.finite-maps where ```agda open import elementary-number-theory.natural-numbers +open import foundation.decidable-embeddings +open import foundation.decidable-propositions +open import foundation.dependent-pair-types open import foundation.fibers-of-maps open import foundation.propositions open import foundation.universe-levels @@ -43,3 +46,18 @@ module _ is-prop-is-finite-map : is-prop is-finite-map is-prop-is-finite-map = is-prop-type-Prop is-finite-prop-map ``` + +## Properties + +### Decidable embeddings are finite maps + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-finite-map-is-decidable-emb : + is-decidable-emb f → is-finite-map f + is-finite-map-is-decidable-emb H x = + is-finite-is-decidable-prop (is-decidable-prop-map-is-decidable-emb H x) +``` From a23ea646b48007ca4280c3b51cbb571ee852baa2 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 19 Jan 2025 09:53:54 -0500 Subject: [PATCH 54/72] work --- .../2-adic-decomposition.lagda.md | 16 ++++ .../divisibility-natural-numbers.lagda.md | 23 +++++ .../exponentiation-natural-numbers.lagda.md | 93 +++++++++++++++++++ ...st-power-divisors-natural-numbers.lagda.md | 14 +++ .../sums-of-natural-numbers.lagda.md | 9 +- .../powers-of-elements-semirings.lagda.md | 3 +- 6 files changed, 154 insertions(+), 4 deletions(-) diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 7b1b4f429b..2c0929ddb2 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -23,6 +23,7 @@ open import foundation.dependent-pair-types open import foundation.equality-cartesian-product-types open import foundation.function-types open import foundation.functoriality-dependent-pair-types +open import foundation.propositions open import foundation.split-surjective-maps open import foundation.transport-along-identifications open import foundation.unit-type @@ -212,6 +213,21 @@ module _ eq-2-adic-decomposition-nat-ℕ ``` +### If $2^km$ is a $2$-adic decomposition of a number $n$, then $2^k$ is the largest power divisor of $n$ base $2$ + +```agda +largest-power-divisor-2-adic-decomposition-ℕ : + (n : ℕ) → 2-adic-decomposition-ℕ n → {!!} +largest-power-divisor-2-adic-decomposition-ℕ = {!!} +``` + +### The type of 2-adic decompositions of any natural number is a proposition + +```agda +all-elements-equal-2-adic-decomposition-ℕ : + (n : ℕ) → all-elements-equal (2-adic-decomposition-ℕ n) +all-elements-equal-2-adic-decomposition-ℕ n (k , m , p) (k' , m' , p') = {!!} +``` ```text pair-expansion : ℕ → UU lzero diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 15b02fca03..53261264ce 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -714,6 +714,29 @@ le-quotient-div-ℕ d x f H g = ( upper-bound-quotient-div-ℕ d x H)) ``` +### If `m` is nonzero, then `a | b` if and only if `ma | mb` + +```agda +reflects-div-left-mul-ℕ : + (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (m *ℕ a) (m *ℕ b) → div-ℕ a b +pr1 (reflects-div-left-mul-ℕ m a b H K) = quotient-div-ℕ (m *ℕ a) (m *ℕ b) K +pr2 (reflects-div-left-mul-ℕ m a b H K) = + is-injective-left-mul-ℕ m H + ( ( left-swap-mul-ℕ m + ( quotient-div-ℕ (m *ℕ a) (m *ℕ b) K) + ( a)) ∙ + ( eq-quotient-div-ℕ (m *ℕ a) (m *ℕ b) K)) + +reflects-div-right-mul-ℕ : + (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (a *ℕ m) (b *ℕ m) → div-ℕ a b +reflects-div-right-mul-ℕ m a b H K = + reflects-div-left-mul-ℕ m a b H + ( concatenate-eq-div-eq-ℕ + ( commutative-mul-ℕ m a) + ( K) + ( commutative-mul-ℕ b m)) +``` + ## See also - [Bounded divisibility of natural numbers](elementary-number-theory.bounded-divisibility-natural-numbers.md) diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index 5e04359f90..b9f9dfa82b 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -12,6 +12,7 @@ open import commutative-algebra.powers-of-elements-commutative-semirings open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.bounded-divisibility-natural-numbers open import elementary-number-theory.commutative-semiring-of-natural-numbers +open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers @@ -95,6 +96,15 @@ _^ℕ_ = exp-ℕ ```agda is-exponential-ℕ : ℕ → ℕ → UU lzero is-exponential-ℕ m n = fiber (exp-ℕ m) n + +valuation-is-exponential-ℕ : + (m n : ℕ) → is-exponential-ℕ m n → ℕ +valuation-is-exponential-ℕ m n = pr1 + +eq-valuation-is-exponential-ℕ : + (m n : ℕ) (H : is-exponential-ℕ m n) → + exp-ℕ m (valuation-is-exponential-ℕ m n H) = n +eq-valuation-is-exponential-ℕ m n = pr2 ``` ## Properties @@ -226,6 +236,11 @@ div-exp-succ-ℕ : pr1 (div-exp-succ-ℕ m n) = m ^ℕ n pr2 (div-exp-succ-ℕ m n) = inv (exp-succ-ℕ m n) +div-exp-succ-ℕ' : + (m n : ℕ) → div-ℕ (m ^ℕ n) (m ^ℕ succ-ℕ n) +pr1 (div-exp-succ-ℕ' m n) = m +pr2 (div-exp-succ-ℕ' m n) = inv (exp-succ-ℕ' m n) + div-exp-is-successor-ℕ : (m n : ℕ) → is-successor-ℕ n → div-ℕ m (m ^ℕ n) div-exp-is-successor-ℕ m .(succ-ℕ k) (k , refl) = @@ -441,3 +456,81 @@ is-finite-map-exp-ℕ : is-finite-map-exp-ℕ m H = is-finite-map-is-decidable-emb (is-decidable-emb-exp-ℕ m H) ``` + +### Exponentiation with base $m>1$ reflects divisibility into inequality + +```agda +leq-exponent-div-exp-ℕ : + (m n k : ℕ) → 1 <-ℕ m → div-ℕ (m ^ℕ k) (m ^ℕ n) → k ≤-ℕ n +leq-exponent-div-exp-ℕ m zero-ℕ zero-ℕ H K = refl-leq-ℕ 0 +leq-exponent-div-exp-ℕ m zero-ℕ (succ-ℕ k) H K = + neq-le-ℕ 1 m H + ( inv (is-one-div-one-ℕ (m ^ℕ succ-ℕ k) K) ∙ + exp-succ-ℕ m k ∙ + ap + ( λ x → m ^ℕ x *ℕ m) + ( is-zero-leq-zero-ℕ k + ( leq-exponent-div-exp-ℕ m 0 k H + ( transitive-div-ℕ + ( m ^ℕ k) + ( m ^ℕ succ-ℕ k) + ( 1) + ( K) + ( div-exp-succ-ℕ' m k)))) ∙ + left-unit-law-mul-ℕ m) +leq-exponent-div-exp-ℕ m (succ-ℕ n) zero-ℕ H K = leq-zero-ℕ (succ-ℕ n) +leq-exponent-div-exp-ℕ m (succ-ℕ n) (succ-ℕ k) H K = + leq-exponent-div-exp-ℕ m n k H + ( reflects-div-right-mul-ℕ m + ( m ^ℕ k) + ( m ^ℕ n) + ( is-nonzero-le-ℕ 1 m H) + ( concatenate-eq-div-eq-ℕ (inv (exp-succ-ℕ m k)) K (exp-succ-ℕ m n))) +``` + +### If $m^k \mid m^n$ and $m$ is nonzero, then its quotient is $m^{d(n,k)}$ + +In the case where $m=0$ we have $0^n \mid 0^n$. For any nonzero number $n$ this would give $0 \mid 0$. The quotient of $0$ divided by $0$ is the unique number $k \leq 0$ such that $k\cdot 0=0$, i.e., it is $0$. However, since $d(n,n)=0$ we have $m^{d(n,n)}=1$, which isn't the quotient under our definitions. + + +```agda +valuation-is-exponent-div-exp-exp-ℕ : + (m n k : ℕ) → is-nonzero-ℕ m → div-ℕ (m ^ℕ k) (m ^ℕ n) → ℕ +valuation-is-exponent-div-exp-exp-ℕ m n k H K = + dist-ℕ n k + +eq-valuation-is-exponent-div-exp-exp-ℕ : + (m n k : ℕ) (H : is-nonzero-ℕ m) (K : div-ℕ (m ^ℕ k) (m ^ℕ n)) → + exp-ℕ m (valuation-is-exponent-div-exp-exp-ℕ m n k H K) = + quotient-div-ℕ (m ^ℕ k) (m ^ℕ n) K +eq-valuation-is-exponent-div-exp-exp-ℕ zero-ℕ n k H K = + ex-falso (H refl) +eq-valuation-is-exponent-div-exp-exp-ℕ (succ-ℕ zero-ℕ) n k H K = + annihilation-law-exp-ℕ (dist-ℕ n k) ∙ + compute-quotient-div-ℕ + ( inv (annihilation-law-exp-ℕ k)) + ( inv (annihilation-law-exp-ℕ n)) + ( refl-div-ℕ 1) + ( K) +eq-valuation-is-exponent-div-exp-exp-ℕ (succ-ℕ (succ-ℕ m)) n k H K = + is-injective-left-mul-ℕ + ( succ-ℕ (succ-ℕ m) ^ℕ k) + ( is-nonzero-exp-ℕ (succ-ℕ (succ-ℕ m)) k H) + ( inv (left-distributive-exp-add-ℕ k (dist-ℕ n k)) ∙ + ap (succ-ℕ (succ-ℕ m) ^ℕ_) (ap (k +ℕ_) (symmetric-dist-ℕ n k) ∙ + is-difference-dist-ℕ k n + ( leq-exponent-div-exp-ℕ (succ-ℕ (succ-ℕ m)) n k star K)) ∙ + inv + ( eq-quotient-div-ℕ' + ( succ-ℕ (succ-ℕ m) ^ℕ k) + ( succ-ℕ (succ-ℕ m) ^ℕ n) + ( K))) + +is-exponent-div-exp-exp-ℕ : + (m n k : ℕ) → is-nonzero-ℕ m → (K : div-ℕ (m ^ℕ k) (m ^ℕ n)) → + is-exponential-ℕ m (quotient-div-ℕ (m ^ℕ k) (m ^ℕ n) K) +pr1 (is-exponent-div-exp-exp-ℕ m n k H K) = + valuation-is-exponent-div-exp-exp-ℕ m n k H K +pr2 (is-exponent-div-exp-exp-ℕ m n k H K) = + eq-valuation-is-exponent-div-exp-exp-ℕ m n k H K +``` diff --git a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md index 555c2eba46..a557d8e499 100644 --- a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md @@ -18,6 +18,7 @@ open import elementary-number-theory.upper-bounds-natural-numbers open import foundation.cartesian-product-types open import foundation.dependent-pair-types +open import foundation.logical-equivalences open import foundation.universe-levels ``` @@ -36,6 +37,14 @@ is-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero is-power-divisor-ℕ m n x = is-exponential-ℕ m x × div-ℕ x n ``` +### The predicate of being the largest power divisor of a natural number + +```agda +is-largest-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero +is-largest-power-divisor-ℕ m n x = + (y : ℕ) → is-power-divisor-ℕ m n y ↔ div-ℕ y x +``` + ### The largest power divisor of a natural number ```agda @@ -88,6 +97,11 @@ module _ ( largest-power-divisor-ℕ) ( k) ( leq-div-ℕ (m ^ℕ k) n (is-nonzero-leq-one-ℕ n K) L , L) + + is-largest-power-divisor-valuation-largest-power-divisor-ℕ : + is-largest-power-divisor-ℕ m n valuation-largest-power-divisor-ℕ + is-largest-power-divisor-valuation-largest-power-divisor-ℕ = + {!is-least-upper-bound-is-upper-bound-ℕ!} ``` ## See also diff --git a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md index 341fd6a89f..3790c80f92 100644 --- a/src/elementary-number-theory/sums-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/sums-of-natural-numbers.lagda.md @@ -98,7 +98,8 @@ This function defines the sum of a family of natural numbers indexed by natural numbers up to, and including, a given upper bound. ```agda -bounded-sum-ℕ : (N : ℕ) → ((i : ℕ) → leq-ℕ i N → ℕ) → ℕ +bounded-sum-ℕ : + (N : ℕ) → ((i : ℕ) → leq-ℕ i N → ℕ) → ℕ bounded-sum-ℕ zero-ℕ f = f 0 (refl-leq-ℕ 0) bounded-sum-ℕ (succ-ℕ N) f = @@ -113,8 +114,10 @@ This function defines the sum of a family of natural numbers indexed by natural numbers up to, but not including, a given upper bound. ```agda -strictly-bounded-sum-ℕ : (u : ℕ) → ((x : ℕ) → le-ℕ x u → ℕ) → ℕ -strictly-bounded-sum-ℕ zero-ℕ f = zero-ℕ +strictly-bounded-sum-ℕ : + (u : ℕ) → ((x : ℕ) → le-ℕ x u → ℕ) → ℕ +strictly-bounded-sum-ℕ zero-ℕ f = + zero-ℕ strictly-bounded-sum-ℕ (succ-ℕ u) f = add-ℕ ( strictly-bounded-sum-ℕ u (λ x H → f x (preserves-le-succ-ℕ x u H))) diff --git a/src/ring-theory/powers-of-elements-semirings.lagda.md b/src/ring-theory/powers-of-elements-semirings.lagda.md index 43501de244..c7798b6794 100644 --- a/src/ring-theory/powers-of-elements-semirings.lagda.md +++ b/src/ring-theory/powers-of-elements-semirings.lagda.md @@ -60,7 +60,8 @@ module _ (n : ℕ) → power-Semiring R (succ-ℕ n) (zero-Semiring R) = zero-Semiring R power-zero-succ-Semiring zero-ℕ = refl power-zero-succ-Semiring (succ-ℕ n) = - right-zero-law-mul-Semiring R (power-Semiring R (succ-ℕ n) (zero-Semiring R)) + right-zero-law-mul-Semiring R + ( power-Semiring R (succ-ℕ n) (zero-Semiring R)) ``` ### `xⁿ⁺¹ = xⁿx` From f64232c491d5ad435262cdcfd4c064b8d1a3b549 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 20 Jan 2025 19:11:43 -0500 Subject: [PATCH 55/72] literature, exercises, and historical references --- references.bib | 54 +++ .../addition-natural-numbers.lagda.md | 14 +- .../cofibonacci.lagda.md | 16 +- .../fibonacci-sequence.lagda.md | 456 ++++++++++++------ .../parity-natural-numbers.lagda.md | 17 + .../pisano-periods.lagda.md | 10 +- src/foundation-core/function-types.lagda.md | 5 +- src/foundation-core/identity-types.lagda.md | 2 +- src/literature/oeis.lagda.md | 2 +- 9 files changed, 409 insertions(+), 167 deletions(-) diff --git a/references.bib b/references.bib index 64520bd954..acfbebb92c 100644 --- a/references.bib +++ b/references.bib @@ -355,6 +355,15 @@ @article{Esc21 keywords = {55U40 03B15,Mathematics - Algebraic Geometry,Mathematics - Logic} } +@book{Euler1748, + author = {Leonhard Euler}, + title = {Introductio in analysin infinitorum}, + year = {1748}, + publisher = {Typographia Academiae Imperialis Petropolitanae}, + url = {https://archive.org/details/introductioanaly00eule_0}, + note = {Accessed: 2025-01-20} +} + @book{FBL73, author = {Fraenkel, Abraham A. and Bar-Hillel, Yehoshua and Levy, Azriel}, @@ -378,6 +387,14 @@ @online{Felixwellen/DCHoTT-Agda howpublished = {{{GitHub}} repository} } +@misc{Fibonacci1202, + author = {Fibonacci}, + title = {Liber Abaci}, + year = {1202}, + note = {A scanned copy of Fibonacci's *Liber Abaci* is available through the Linda Hall Library Digital Collections.}, + url = {https://old.maa.org/press/periodicals/convergence/mathematical-treasure-fibonacci-s-liber-abaci} +} + @online{GGMS24, title = {The {{Category}} of {{Iterative Sets}} in {{Homotopy Type Theory}} and {{Univalent Foundations}}}, author = {Gratzer, Daniel and Gylterud, Håkon and Mörtberg, Anders and Stenholm, Elisabeth}, @@ -445,6 +462,14 @@ @inproceedings{KvR19 eprintclass = {cs, math} } +@book{Leibniz1693, + author = {Gottfried Wilhelm Leibniz}, + title = {De geometria analysis}, + year = {1693}, + publisher = {Leibniz's own press, Vienna}, + note = {This text is often cited as one of Leibniz's early works on calculus.} +} + @book{Leveque12volI, title = {Topics in Number Theory, Volume I}, author = {LeVeque, W.J.}, @@ -463,6 +488,15 @@ @book{Leveque12volII publisher = {Dover Publications} } +@book{lucas1891, + title = {Théorie des Nombres}, + author = {Lucas, Édouard}, + year = {1891}, + publisher = {Gauthier-Villars}, + address = {Paris}, + url = {https://archive.org/details/thoriedesnombr01lucauoft} +} + @incollection{Makkai98, author = {Makkai, M.}, title = {Towards a categorical foundation of mathematics}, @@ -643,6 +677,15 @@ @phdthesis{Qui16 langid = {english} } +@book{Recorde1557, + author = {Robert Recorde}, + title = {The Whetstone of Witte}, + year = {1557}, + publisher = {Printed by John Daye, London}, + note = {This work contains the first recorded use of the equals sign (=).}, + url = {https://archive.org/embed/TheWhetstoneOfWitte} +} + @book{Rie17, title = {Category {{Theory}} in {{Context}}}, author = {Riehl, Emily}, @@ -950,3 +993,14 @@ @online{Warn24 pubstate = {preprint}, keywords = {Mathematics - Algebraic Topology,Mathematics - Category Theory} } + +@book{Widmann1489, + author = {Johannes Widmann}, + title = {Behende und hüpsche Rechenung auff allen Kauffmanschafft}, + year = {1489}, + publisher = {Conrad Kachelofen}, + address = {Leipzig}, + note = {First use of "+" and "−" symbols in printed mathematics}, + language = {German}, + url = {https://www.loc.gov/item/49038907} +} \ No newline at end of file diff --git a/src/elementary-number-theory/addition-natural-numbers.lagda.md b/src/elementary-number-theory/addition-natural-numbers.lagda.md index 1be8622c65..7cf13fa8d8 100644 --- a/src/elementary-number-theory/addition-natural-numbers.lagda.md +++ b/src/elementary-number-theory/addition-natural-numbers.lagda.md @@ -26,7 +26,15 @@ open import foundation.sets
-## Definition +## Idea + +The {{#concept "addition" Disambiguation="natural numbers" Agda=_+ℕ_}} operation on the [natural numbers](elementary-number-theory.natural-numbers.md) is a binary operation $m,n\mapsto m+n$ on the natural numbers given by $n$ times iteratively taking the successor of the number $m$. + +The notation $+$ for addition and $-$ for subtraction appeared first in print in 1489 in _Behende und hüpsche Rechenung auff allen Kauffmanschafft_ by [Johannes Widmann](https://en.wikipedia.org/wiki/Johannes_Widmann), for example on page 327 of {{#cite Widmann1489}}. + +## Definitions + +### Addition on the natural numbers ```agda add-ℕ : ℕ → ℕ → ℕ @@ -227,3 +235,7 @@ abstract - The commutative monoid of the natural numbers with addition is defined in [`monoid-of-natural-numbers-with-addition`](elementary-number-theory.monoid-of-natural-numbers-with-addition.md). + +## References + +{{#bibliography}} diff --git a/src/elementary-number-theory/cofibonacci.lagda.md b/src/elementary-number-theory/cofibonacci.lagda.md index 16ac9c058c..9d2112effe 100644 --- a/src/elementary-number-theory/cofibonacci.lagda.md +++ b/src/elementary-number-theory/cofibonacci.lagda.md @@ -34,7 +34,7 @@ The {{#concept "cofibonacci sequence" Agda=cofibonacci}} is the unique function `G : ℕ → ℕ` satisfying the property that ```text - div-ℕ (G m) n ↔ div-ℕ m (Fibonacci-ℕ n). + div-ℕ (G m) n ↔ div-ℕ m (fibonacci-ℕ n). ``` The sequence starts @@ -51,7 +51,7 @@ The sequence starts ```agda is-multiple-of-cofibonacci : (m x : ℕ) → UU lzero is-multiple-of-cofibonacci m x = - is-nonzero-ℕ m → is-nonzero-ℕ x × div-ℕ m (Fibonacci-ℕ x) + is-nonzero-ℕ m → is-nonzero-ℕ x × div-ℕ m (fibonacci-ℕ x) is-decidable-is-multiple-of-cofibonacci : (m x : ℕ) → is-decidable (is-multiple-of-cofibonacci m x) @@ -60,7 +60,7 @@ is-decidable-is-multiple-of-cofibonacci m x = ( is-decidable-zero-is-nonzero-ℕ m) ( is-decidable-product ( is-decidable-zero-is-nonzero-ℕ x) - ( is-decidable-div-ℕ m (Fibonacci-ℕ x))) + ( is-decidable-div-ℕ m (fibonacci-ℕ x))) cofibonacci-multiple : (m : ℕ) → Σ ℕ (is-multiple-of-cofibonacci m) cofibonacci-multiple zero-ℕ = pair zero-ℕ (λ f → (ex-falso (f refl))) @@ -112,12 +112,12 @@ is-zero-cofibonacci-zero-ℕ = ```agda forward-is-left-adjoint-cofibonacci : - (m n : ℕ) → div-ℕ (cofibonacci m) n → div-ℕ m (Fibonacci-ℕ n) + (m n : ℕ) → div-ℕ (cofibonacci m) n → div-ℕ m (fibonacci-ℕ n) forward-is-left-adjoint-cofibonacci zero-ℕ n H = tr ( div-ℕ zero-ℕ) ( ap - ( Fibonacci-ℕ) + ( fibonacci-ℕ) ( inv ( is-zero-div-zero-ℕ n ( tr (λ t → div-ℕ t n) is-zero-cofibonacci-zero-ℕ H)))) @@ -125,7 +125,7 @@ forward-is-left-adjoint-cofibonacci zero-ℕ n H = forward-is-left-adjoint-cofibonacci (succ-ℕ m) zero-ℕ H = div-zero-ℕ (succ-ℕ m) forward-is-left-adjoint-cofibonacci (succ-ℕ m) (succ-ℕ n) H = - div-Fibonacci-div-ℕ + div-fibonacci-div-ℕ ( succ-ℕ m) ( cofibonacci (succ-ℕ m)) ( succ-ℕ n) @@ -137,11 +137,11 @@ forward-is-left-adjoint-cofibonacci (succ-ℕ m) (succ-ℕ n) H = {- converse-is-left-adjoint-cofibonacci : - (m n : ℕ) → div-ℕ m (Fibonacci-ℕ n) → div-ℕ (cofibonacci m) n + (m n : ℕ) → div-ℕ m (fibonacci-ℕ n) → div-ℕ (cofibonacci m) n converse-is-left-adjoint-cofibonacci m n H = {!!} is-left-adjoint-cofibonacci : - (m n : ℕ) → div-ℕ (cofibonacci m) n ↔ div-ℕ m (Fibonacci-ℕ n) + (m n : ℕ) → div-ℕ (cofibonacci m) n ↔ div-ℕ m (fibonacci-ℕ n) is-left-adjoint-cofibonacci m n = {!!} -} ``` diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index c0e10b8df6..808a83b6f9 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -8,18 +8,23 @@ module elementary-number-theory.fibonacci-sequence where ```agda open import elementary-number-theory.addition-natural-numbers +open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.greatest-common-divisor-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.integers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers +open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.relatively-prime-natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import elementary-number-theory.squares-natural-numbers +open import elementary-number-theory.sums-of-natural-numbers open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types +open import foundation.function-types open import foundation.identity-types open import foundation.transport-along-identifications open import foundation.unit-type @@ -30,7 +35,7 @@ open import foundation.unit-type ## Idea The -{{#concept "Fibonacci sequence" WD="Fibonacci sequence" WDID=Q23835349 Agda=Fibonacci-ℕ}} +{{#concept "Fibonacci sequence" WD="Fibonacci sequence" WDID=Q23835349 Agda=fibonacci-ℕ OEIS=A000045}} is a recursively defined [sequence](foundation.sequences.md) of [natural numbers](elementary-number-theory.natural-numbers.md) given by the equations @@ -48,18 +53,27 @@ Fibonacci numbers are Fₙ 0 1 1 2 3 5 8 13 21 34 ``` +The Fibonacci sequence was introduced by [Leonardo Fibonacci](https://en.wikipedia.org/wiki/Fibonacci) in 1202 in his book _Liber Abaci_ {{#cite Fibonacci1202}}. + ## Definitions -### The standard definition using pattern matching +### The Fibonacci sequence + +```agda +fibonacci-ℕ : ℕ → ℕ +fibonacci-ℕ 0 = 0 +fibonacci-ℕ (succ-ℕ 0) = 1 +fibonacci-ℕ (succ-ℕ (succ-ℕ n)) = fibonacci-ℕ (succ-ℕ n) +ℕ fibonacci-ℕ n +``` + +### The Fibonacci sequence in the integers ```agda -Fibonacci-ℕ : ℕ → ℕ -Fibonacci-ℕ 0 = 0 -Fibonacci-ℕ (succ-ℕ 0) = 1 -Fibonacci-ℕ (succ-ℕ (succ-ℕ n)) = (Fibonacci-ℕ (succ-ℕ n)) +ℕ (Fibonacci-ℕ n) +fibonacci-ℤ : ℕ → ℤ +fibonacci-ℤ = int-ℕ ∘ fibonacci-ℕ ``` -### A definition using the induction principle of `ℕ` +### The Fibonacci sequence defined directly with the induction principle of `ℕ` The above definition of the Fibonacci sequence uses Agda's powerful pattern matching definitions. Below, we will give a definition of the Fibonacci sequence @@ -77,30 +91,40 @@ Such a function is easy to give with induction, using the map $ℕ² → ℕ²$ takes a pair `(m,n)` to the pair `(n,n+m)`. Starting the iteration with `(0,1)` we obtain the Fibonacci sequence by taking the first projection. -However, we haven't defined cartesian products or booleans yet. Therefore we -mimic the above idea, using $ℕ → ℕ$ instead of $ℕ²$. +Furthermore, it is possible to define the fibonacci sequence without using cartesian products, mimicing the above idea using $\mathbb{N} \to \matbb{N}$ instead of $\mathbb{N}^2$. ```agda -shift-one : ℕ → (ℕ → ℕ) → (ℕ → ℕ) -shift-one n f = ind-ℕ n (λ x y → f x) - -shift-two : ℕ → ℕ → (ℕ → ℕ) → (ℕ → ℕ) -shift-two m n f = shift-one m (shift-one n f) - -Fibo-zero-ℕ : ℕ → ℕ -Fibo-zero-ℕ = shift-two 0 1 (λ x → 0) - -Fibo-succ-ℕ : (ℕ → ℕ) → (ℕ → ℕ) -Fibo-succ-ℕ f = shift-two (f 1) ((f 1) +ℕ (f 0)) (λ x → 0) - -Fibo-function : ℕ → ℕ → ℕ -Fibo-function = +shift-one : + ℕ → (ℕ → ℕ) → (ℕ → ℕ) +shift-one n f = + ind-ℕ n (λ x y → f x) + +shift-two : + ℕ → ℕ → (ℕ → ℕ) → (ℕ → ℕ) +shift-two m n f = + shift-one m (shift-one n f) + +fibo-zero-ℕ : + ℕ → ℕ +fibo-zero-ℕ = + shift-two 0 1 (λ x → 0) + +fibo-succ-ℕ : + (ℕ → ℕ) → (ℕ → ℕ) +fibo-succ-ℕ f = + shift-two (f 1) ((f 1) +ℕ (f 0)) (λ x → 0) + +fibo-function : + ℕ → ℕ → ℕ +fibo-function = ind-ℕ - ( Fibo-zero-ℕ) - ( λ n → Fibo-succ-ℕ) + ( fibo-zero-ℕ) + ( λ n → fibo-succ-ℕ) -Fibo : ℕ → ℕ -Fibo k = Fibo-function k 0 +fibo-ℕ : + ℕ → ℕ +fibo-ℕ k = + fibo-function k 0 ``` ## Properties @@ -108,67 +132,68 @@ Fibo k = Fibo-function k 0 ### `F(m+n+1) = F(m+1)F(n+1) + F(m)F(n)` ```agda -Fibonacci-add-ℕ : +fibonacci-add-ℕ : (m n : ℕ) → - Fibonacci-ℕ (m +ℕ (succ-ℕ n)) = - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ (succ-ℕ n))) +ℕ - ( (Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ n)) -Fibonacci-add-ℕ m zero-ℕ = + fibonacci-ℕ (m +ℕ succ-ℕ n) = + fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ (succ-ℕ n) +ℕ + fibonacci-ℕ m *ℕ fibonacci-ℕ n +fibonacci-add-ℕ m zero-ℕ = ap-add-ℕ - ( inv (right-unit-law-mul-ℕ (Fibonacci-ℕ (succ-ℕ m)))) - ( inv (right-zero-law-mul-ℕ (Fibonacci-ℕ m))) -Fibonacci-add-ℕ m (succ-ℕ n) = - ( ap Fibonacci-ℕ (inv (left-successor-law-add-ℕ m (succ-ℕ n)))) ∙ - ( Fibonacci-add-ℕ (succ-ℕ m) n) ∙ + ( inv (right-unit-law-mul-ℕ (fibonacci-ℕ (succ-ℕ m)))) + ( inv (right-zero-law-mul-ℕ (fibonacci-ℕ m))) +fibonacci-add-ℕ m (succ-ℕ n) = + ( ap fibonacci-ℕ (inv (left-successor-law-add-ℕ m (succ-ℕ n)))) ∙ + ( fibonacci-add-ℕ (succ-ℕ m) n) ∙ ( ap - ( _+ℕ ((Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ n))) + ( _+ℕ fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ n) ( right-distributive-mul-add-ℕ - ( Fibonacci-ℕ (succ-ℕ m)) - ( Fibonacci-ℕ m) - ( Fibonacci-ℕ (succ-ℕ n)))) ∙ + ( fibonacci-ℕ (succ-ℕ m)) + ( fibonacci-ℕ m) + ( fibonacci-ℕ (succ-ℕ n)))) ∙ ( associative-add-ℕ - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ (succ-ℕ n))) - ( (Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ (succ-ℕ n))) - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ n))) ∙ + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ m *ℕ fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ n)) ∙ ( ap - ( ((Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ (succ-ℕ n))) +ℕ_) + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ (succ-ℕ n) +ℕ_) ( commutative-add-ℕ - ( (Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ (succ-ℕ n))) - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ n)))) ∙ + ( fibonacci-ℕ m *ℕ fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ n))) ∙ ( inv ( associative-add-ℕ - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ (succ-ℕ n))) - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ n)) - ( (Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ (succ-ℕ n))))) ∙ + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ n) + ( fibonacci-ℕ m *ℕ fibonacci-ℕ (succ-ℕ n)))) ∙ ( ap - ( _+ℕ ((Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ (succ-ℕ n)))) + ( _+ℕ fibonacci-ℕ m *ℕ fibonacci-ℕ (succ-ℕ n)) ( inv ( left-distributive-mul-add-ℕ - ( Fibonacci-ℕ (succ-ℕ m)) - ( Fibonacci-ℕ (succ-ℕ n)) - ( Fibonacci-ℕ n)))) + ( fibonacci-ℕ (succ-ℕ m)) + ( fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ n)))) ``` ### Consecutive Fibonacci numbers are relatively prime ```agda is-one-div-fibonacci-succ-div-fibonacci-ℕ : - (d n : ℕ) → div-ℕ d (Fibonacci-ℕ n) → div-ℕ d (Fibonacci-ℕ (succ-ℕ n)) → + (d n : ℕ) → div-ℕ d (fibonacci-ℕ n) → div-ℕ d (fibonacci-ℕ (succ-ℕ n)) → is-one-ℕ d -is-one-div-fibonacci-succ-div-fibonacci-ℕ d zero-ℕ H K = is-one-div-one-ℕ d K +is-one-div-fibonacci-succ-div-fibonacci-ℕ d zero-ℕ H K = + is-one-div-one-ℕ d K is-one-div-fibonacci-succ-div-fibonacci-ℕ d (succ-ℕ n) H K = is-one-div-fibonacci-succ-div-fibonacci-ℕ d n - ( div-right-summand-ℕ d (Fibonacci-ℕ (succ-ℕ n)) (Fibonacci-ℕ n) H K) + ( div-right-summand-ℕ d (fibonacci-ℕ (succ-ℕ n)) (fibonacci-ℕ n) H K) ( H) relatively-prime-fibonacci-fibonacci-succ-ℕ : - (n : ℕ) → is-relatively-prime-ℕ (Fibonacci-ℕ n) (Fibonacci-ℕ (succ-ℕ n)) + (n : ℕ) → is-relatively-prime-ℕ (fibonacci-ℕ n) (fibonacci-ℕ (succ-ℕ n)) relatively-prime-fibonacci-fibonacci-succ-ℕ n = is-one-div-fibonacci-succ-div-fibonacci-ℕ - ( gcd-ℕ (Fibonacci-ℕ n) (Fibonacci-ℕ (succ-ℕ n))) + ( gcd-ℕ (fibonacci-ℕ n) (fibonacci-ℕ (succ-ℕ n))) ( n) - ( div-left-factor-gcd-ℕ (Fibonacci-ℕ n) (Fibonacci-ℕ (succ-ℕ n))) - ( div-right-factor-gcd-ℕ (Fibonacci-ℕ n) (Fibonacci-ℕ (succ-ℕ n))) + ( div-left-factor-gcd-ℕ (fibonacci-ℕ n) (fibonacci-ℕ (succ-ℕ n))) + ( div-right-factor-gcd-ℕ (fibonacci-ℕ n) (fibonacci-ℕ (succ-ℕ n))) ``` ### A 3-for-2 property of divisibility of Fibonacci numbers @@ -176,66 +201,196 @@ relatively-prime-fibonacci-fibonacci-succ-ℕ n = We prove that if an two of the following three properties hold, then so does the third: -1. `d | Fibonacci-ℕ m` -2. `d | Fibonacci-ℕ n` -3. `d | Fibonacci-ℕ (m +ℕ n)`. +1. `d | fibonacci-ℕ m` +2. `d | fibonacci-ℕ n` +3. `d | fibonacci-ℕ (m +ℕ n)`. ```agda -div-Fibonacci-add-ℕ : - (d m n : ℕ) → div-ℕ d (Fibonacci-ℕ m) → div-ℕ d (Fibonacci-ℕ n) → - div-ℕ d (Fibonacci-ℕ (m +ℕ n)) -div-Fibonacci-add-ℕ d m zero-ℕ H1 H2 = H1 -div-Fibonacci-add-ℕ d m (succ-ℕ n) H1 H2 = +div-fibonacci-add-ℕ : + (d m n : ℕ) → div-ℕ d (fibonacci-ℕ m) → div-ℕ d (fibonacci-ℕ n) → + div-ℕ d (fibonacci-ℕ (m +ℕ n)) +div-fibonacci-add-ℕ d m zero-ℕ H1 H2 = + H1 +div-fibonacci-add-ℕ d m (succ-ℕ n) H1 H2 = tr ( div-ℕ d) - ( inv (Fibonacci-add-ℕ m n)) + ( inv (fibonacci-add-ℕ m n)) ( div-add-ℕ d - ( (Fibonacci-ℕ (succ-ℕ m)) *ℕ (Fibonacci-ℕ (succ-ℕ n))) - ( (Fibonacci-ℕ m) *ℕ (Fibonacci-ℕ n)) - ( div-mul-ℕ (Fibonacci-ℕ (succ-ℕ m)) d (Fibonacci-ℕ (succ-ℕ n)) H2) + ( fibonacci-ℕ (succ-ℕ m) *ℕ fibonacci-ℕ (succ-ℕ n)) + ( fibonacci-ℕ m *ℕ fibonacci-ℕ n) + ( div-mul-ℕ (fibonacci-ℕ (succ-ℕ m)) d (fibonacci-ℕ (succ-ℕ n)) H2) ( tr ( div-ℕ d) - ( commutative-mul-ℕ (Fibonacci-ℕ n) (Fibonacci-ℕ m)) - ( div-mul-ℕ (Fibonacci-ℕ n) d (Fibonacci-ℕ m) H1))) + ( commutative-mul-ℕ (fibonacci-ℕ n) (fibonacci-ℕ m)) + ( div-mul-ℕ (fibonacci-ℕ n) d (fibonacci-ℕ m) H1))) ``` ### If `m | n`, then `d | F(m)` implies `d | F(n)` ```agda -div-Fibonacci-div-ℕ : - (d m n : ℕ) → div-ℕ m n → div-ℕ d (Fibonacci-ℕ m) → div-ℕ d (Fibonacci-ℕ n) -div-Fibonacci-div-ℕ d m .zero-ℕ (zero-ℕ , refl) H = div-zero-ℕ d -div-Fibonacci-div-ℕ d zero-ℕ .(k *ℕ zero-ℕ) (succ-ℕ k , refl) H = +div-fibonacci-div-ℕ : + (d m n : ℕ) → div-ℕ m n → div-ℕ d (fibonacci-ℕ m) → div-ℕ d (fibonacci-ℕ n) +div-fibonacci-div-ℕ d m .zero-ℕ (zero-ℕ , refl) H = + div-zero-ℕ d +div-fibonacci-div-ℕ d zero-ℕ .(k *ℕ zero-ℕ) (succ-ℕ k , refl) H = tr ( div-ℕ d) - ( ap Fibonacci-ℕ (inv (right-zero-law-mul-ℕ (succ-ℕ k)))) + ( ap fibonacci-ℕ (inv (right-zero-law-mul-ℕ (succ-ℕ k)))) ( div-zero-ℕ d) -div-Fibonacci-div-ℕ d (succ-ℕ m) ._ (succ-ℕ k , refl) H = - div-Fibonacci-add-ℕ d - ( k *ℕ (succ-ℕ m)) +div-fibonacci-div-ℕ d (succ-ℕ m) ._ (succ-ℕ k , refl) H = + div-fibonacci-add-ℕ d + ( k *ℕ succ-ℕ m) ( succ-ℕ m) - ( div-Fibonacci-div-ℕ d - ( succ-ℕ m) - ( k *ℕ (succ-ℕ m)) - ( pair k refl) - ( H)) + ( div-fibonacci-div-ℕ d (succ-ℕ m) (k *ℕ succ-ℕ m) (k , refl) H) ( H) ``` -### Fibonacci-ℕ is an order preserving map on ℕ ordered by divisibility +### The function `fibonacci-ℕ` is an order preserving map on the natural numbers ordered by divisibility + +```agda +preserves-div-fibonacci-ℕ : + (m n : ℕ) → div-ℕ m n → div-ℕ (fibonacci-ℕ m) (fibonacci-ℕ n) +preserves-div-fibonacci-ℕ m n H = + div-fibonacci-div-ℕ (fibonacci-ℕ m) m n H (refl-div-ℕ (fibonacci-ℕ m)) +``` + +### The sum of the first $n + 1$ Fibonacci numbers + +We show that + +$$ + \sum_{k=0}^n F_k = F_{n+2}-1 +$$ + +```agda +sum-of-fibonacci-ℕ : + ℕ → ℕ +sum-of-fibonacci-ℕ n = + bounded-sum-ℕ n (λ k _ → fibonacci-ℕ k) + +compute-sum-of-fibonacci-ℕ' : + (n : ℕ) → succ-ℕ (sum-of-fibonacci-ℕ n) = fibonacci-ℕ (n +ℕ 2) +compute-sum-of-fibonacci-ℕ' zero-ℕ = + refl +compute-sum-of-fibonacci-ℕ' (succ-ℕ n) = + ( inv + ( left-successor-law-add-ℕ + ( bounded-sum-ℕ n (λ k _ → fibonacci-ℕ k)) + ( fibonacci-ℕ (succ-ℕ n)))) ∙ + ( ap (_+ℕ fibonacci-ℕ (succ-ℕ n)) (compute-sum-of-fibonacci-ℕ' n)) + +compute-sum-of-fibonacci-ℕ : + (n : ℕ) → sum-of-fibonacci-ℕ n = dist-ℕ (fibonacci-ℕ (n +ℕ 2)) 1 +compute-sum-of-fibonacci-ℕ n = + ( rewrite-left-add-dist-ℕ + ( bounded-sum-ℕ n (λ k _ → fibonacci-ℕ k)) + ( 1) + ( fibonacci-ℕ (n +ℕ 2)) + ( compute-sum-of-fibonacci-ℕ' n)) ∙ + ( symmetric-dist-ℕ 1 (fibonacci-ℕ (n +ℕ 2))) + +le-sum-of-fibonacci-ℕ : + (n : ℕ) → sum-of-fibonacci-ℕ n <-ℕ fibonacci-ℕ (n +ℕ 2) +le-sum-of-fibonacci-ℕ n = + concatenate-le-eq-ℕ + ( sum-of-fibonacci-ℕ n) + ( succ-ℕ (sum-of-fibonacci-ℕ n)) + ( fibonacci-ℕ (n +ℕ 2)) + ( succ-le-ℕ (sum-of-fibonacci-ℕ n)) + ( compute-sum-of-fibonacci-ℕ' n) +``` + +### The sum of the first $n$ oddly indexed Fibonacci numbers + +We show that + +$$ + \sum_{k=0}^{n-1} F_{2k+1} = F_{2n}. +$$ ```agda -preserves-div-Fibonacci-ℕ : - (m n : ℕ) → div-ℕ m n → div-ℕ (Fibonacci-ℕ m) (Fibonacci-ℕ n) -preserves-div-Fibonacci-ℕ m n H = - div-Fibonacci-div-ℕ (Fibonacci-ℕ m) m n H (refl-div-ℕ (Fibonacci-ℕ m)) +strictly-bounded-sum-fibonacci-odd-number-ℕ : + ℕ → ℕ +strictly-bounded-sum-fibonacci-odd-number-ℕ n = + strictly-bounded-sum-ℕ n (λ k _ → fibonacci-ℕ (odd-number-ℕ k)) + +compute-strictly-bounded-sum-fibonacci-odd-number-ℕ : + (n : ℕ) → + strictly-bounded-sum-fibonacci-odd-number-ℕ n = fibonacci-ℕ (2 *ℕ n) +compute-strictly-bounded-sum-fibonacci-odd-number-ℕ zero-ℕ = + refl +compute-strictly-bounded-sum-fibonacci-odd-number-ℕ (succ-ℕ n) = + ( ap + ( _+ℕ fibonacci-ℕ (odd-number-ℕ n)) + ( compute-strictly-bounded-sum-fibonacci-odd-number-ℕ n)) ∙ + ( commutative-add-ℕ + ( fibonacci-ℕ (2 *ℕ n)) + ( fibonacci-ℕ (odd-number-ℕ n))) ∙ + ( inv (ap fibonacci-ℕ (left-distributive-mul-add-ℕ 2 n 1))) ``` -### LeVeque's strict bound on the growth of the Fibonacci numbers +### The sum of the first $n + 1$ evenly indexed Fibonacci numbers -The $n$th Fibonacci number is always strictly less than $(\frac{7}{4})^n$. This -claim appears on pages 7 and 8 in section 1.2 of {{#cite Leveque12volI}} as an -instructive example of a proof by induction. The upper bound works for any +We show that + +$$ + \sum_{k=0}^{n} F_{2k} = F_{2n+1} - 1. +$$ + +```agda +bounded-sum-fibonacci-even-ℕ : + ℕ → ℕ +bounded-sum-fibonacci-even-ℕ n = + bounded-sum-ℕ n (λ k _ → fibonacci-ℕ (even-number-ℕ k)) + +compute-bounded-sum-fibonacci-even-ℕ' : + (n : ℕ) → + succ-ℕ (bounded-sum-fibonacci-even-ℕ n) = fibonacci-ℕ (odd-number-ℕ n) +compute-bounded-sum-fibonacci-even-ℕ' zero-ℕ = + refl +compute-bounded-sum-fibonacci-even-ℕ' (succ-ℕ n) = + ( right-swap-add-ℕ + ( bounded-sum-fibonacci-even-ℕ n) + ( fibonacci-ℕ (even-number-ℕ (succ-ℕ n))) + ( 1)) ∙ + ( ap-add-ℕ + ( compute-bounded-sum-fibonacci-even-ℕ' n) + ( ap fibonacci-ℕ (even-number-succ-ℕ n))) ∙ + ( commutative-add-ℕ + ( fibonacci-ℕ (odd-number-ℕ n)) + ( fibonacci-ℕ (succ-ℕ (odd-number-ℕ n)))) ∙ + ( inv (ap fibonacci-ℕ (odd-number-succ-ℕ n))) + +compute-bounded-sum-fibonacci-even-ℕ : + (n : ℕ) → + bounded-sum-fibonacci-even-ℕ n = dist-ℕ (fibonacci-ℕ (odd-number-ℕ n)) 1 +compute-bounded-sum-fibonacci-even-ℕ n = + ( rewrite-left-add-dist-ℕ + ( bounded-sum-fibonacci-even-ℕ n) + ( 1) + ( fibonacci-ℕ (odd-number-ℕ n)) + ( compute-bounded-sum-fibonacci-even-ℕ' n)) ∙ + ( symmetric-dist-ℕ 1 (fibonacci-ℕ (odd-number-ℕ n))) + +le-bounded-sum-fibonacci-even-ℕ : + (n : ℕ) → bounded-sum-fibonacci-even-ℕ n <-ℕ fibonacci-ℕ (odd-number-ℕ n) +le-bounded-sum-fibonacci-even-ℕ n = + concatenate-le-eq-ℕ + ( bounded-sum-fibonacci-even-ℕ n) + ( succ-ℕ (bounded-sum-fibonacci-even-ℕ n)) + ( fibonacci-ℕ (odd-number-ℕ n)) + ( succ-le-ℕ (bounded-sum-fibonacci-even-ℕ n)) + ( compute-bounded-sum-fibonacci-even-ℕ' n) +``` + +### A strict bound on the growth of the Fibonacci numbers + +The $n$th Fibonacci number is always strictly less than $\phi^n$, where $\phi$ is the golden ratio. This strict inequality is due to Édouard Lucas, who proved it in the late 19th century {{#cite lucas1891}}. We will prove a rational version of this claim by showing that $F_n<(\frac{b}{a})^n$ for any rational number $\frac{b}{a}>\phi$ + +A variation of this claim also appears on pages 7 and 8 in section 1.2 of +{{#cite Leveque12volI}} as an instructive example of a proof by induction, where +he shows that $F_n<(\frac{7}{4})^n$. This +upper bound works for any fraction $\frac{b}{a}$ where $a(b+a) Date: Mon, 20 Jan 2025 19:11:49 -0500 Subject: [PATCH 56/72] literature, exercises, and historical references --- src/literature/andrews-number-theory.lagda.md | 90 +++++++++++++++++++ 1 file changed, 90 insertions(+) create mode 100644 src/literature/andrews-number-theory.lagda.md diff --git a/src/literature/andrews-number-theory.lagda.md b/src/literature/andrews-number-theory.lagda.md new file mode 100644 index 0000000000..486e3ae4eb --- /dev/null +++ b/src/literature/andrews-number-theory.lagda.md @@ -0,0 +1,90 @@ +# Andrews, G. E., Number Theory + +This file collects references to formalization of constructions and theorems +from {{#cite Andrews94}}. + +```agda +module literature.andrews-number-theory where +``` + +## Part 1 Multiplicativity–Divisibility + +### Chapter 1 Basis Representation + +#### Section 1–1 Basis Representation + +Theorem 1–1 establishes the identity for [triangular numbers](elementary-number-theory.triangular-numbers.md) + +```agda +open import elementary-number-theory.triangular-numbers using + ( compute-triangular-number-ℕ) +``` + +Theorem 1–2 is the formula for the geometric series stated for real numbers. This would be best formalized as an identity of polynomials. However, we don't have polynomials yet. + +Corollary 1–1 states that the [exponential function](elementary-number-theory.exponentiation-natural-numbers.md) is [strictly inflationary](elementary-number-theory.strictly-inflationary-functions-natural-numbers.md). + +```agda +open import elementary-number-theory.exponentiation-natural-numbers using + ( is-strictly-inflationary-exp-ℕ) +``` + +Exercise 1–1 asks to compute the $n$th [square pyramidal number](elementary-number-theory.square-pyramidal-numbers.md) + +```agda +open import elementary-number-theory.square-pyramidal-numbers using + ( compute-square-pyramidal-number-ℕ) +``` + +Exercise 1–2 asks to prove [Nicomachus's theorem](elementary-number-theory.nicomachuss-theorem.md) + +```agda +open import elementary-number-theory.nicomachuss-theorem using + ( nicomachuss-theorem-ℕ) +``` + +Exercise 1–3 is a polynomial identity, which we have to skip for now. + +Exercise 1–4 computes the sum of the first $n$ [pronic numbers](elementary-number-theory.pronic-numbers.md) + +```agda +open import elementary-number-theory.pronic-numbers using + ( compute-sum-of-pronic-numbers-ℕ) +``` + +Exercise 1–5 computes the sum of the first $n$ [odd numbers](elementary-number-theory.parity-natural-numbers.md) + +```agda +open import elementary-number-theory.squares-natural-numbers using + ( compute-sum-of-odd-numbers-ℕ) +``` + +Exercise 1–6 asks to compute the sum of the first $n$ multiplicative inverses of the pronic numbers. The infrastructure exists in the library, but we haven't done so yet. + +Exercise 1–7 computes the sum of the first $n$ Fibonacci numbers + +```agda +open import elementary-number-theory.fibonacci-sequence using + ( compute-sum-of-fibonacci-ℕ) +``` + +Exercise 1–8 computes the sum of the first $n$ oddly indexed Fibonacci numbers + +```agda +open import elementary-number-theory.fibonacci-sequence using + ( compute-strictly-bounded-sum-fibonacci-odd-number-ℕ) +``` + +Exercise 1–9 computes the sum of the first $n$ evenly indexed Fibonacci numbers + +```agda +open import elementary-number-theory.fibonacci-sequence using + ( compute-bounded-sum-fibonacci-even-ℕ) +``` + +Exercise 1–10 asks to prove +[Cassini's identity](https://en.wikipedia.org/wiki/Cassini_and_Catalan_identities), named after the Italian mathematican and astronomer Giovanni Domenico Cassini, who proved the result in 1680. Note that this identity has two further generalizations, one due to Eugène Catalan and one due to Steven Vajda. + +```agda + +``` From b66615ebfa036c02677b38f11275c10775fdf846 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 20 Jan 2025 19:19:47 -0500 Subject: [PATCH 57/72] formatting --- .../2-adic-decomposition.lagda.md | 19 ++++-- .../addition-natural-numbers.lagda.md | 12 +++- ...nded-divisibility-natural-numbers.lagda.md | 35 ++++++++--- .../bounded-natural-numbers.lagda.md | 7 ++- .../congruence-natural-numbers.lagda.md | 23 ++++++-- .../decidable-maps-natural-numbers.lagda.md | 7 ++- ...ecidable-subtypes-natural-numbers.lagda.md | 16 +++-- .../distance-natural-numbers.lagda.md | 15 +++-- .../divisibility-natural-numbers.lagda.md | 25 ++++++-- ...uclidean-division-natural-numbers.lagda.md | 34 ++++++++--- .../exponentiation-natural-numbers.lagda.md | 20 +++++-- .../farey-fractions.lagda.md | 27 ++++++--- .../fibonacci-sequence.lagda.md | 25 +++++--- .../finite-maps-natural-numbers.lagda.md | 58 ++++++++++++++----- .../inequality-integers.lagda.md | 5 +- ...tionary-functions-natural-numbers.lagda.md | 9 ++- ...st-power-divisors-natural-numbers.lagda.md | 10 +++- .../lower-bounds-natural-numbers.lagda.md | 31 +++++++--- ...aximal-structured-natural-numbers.lagda.md | 13 ++++- .../maximum-natural-numbers.lagda.md | 9 ++- .../mersenne-exponents.lagda.md | 5 +- ...inimal-structured-natural-numbers.lagda.md | 10 +++- .../natural-numbers.lagda.md | 13 ++++- .../number-of-divisors.lagda.md | 12 +++- .../parity-natural-numbers.lagda.md | 6 +- ...l-numbers-ordered-by-divisibility.lagda.md | 18 ++++-- .../prime-numbers.lagda.md | 8 ++- .../proper-divisors-natural-numbers.lagda.md | 5 +- .../rational-numbers.lagda.md | 6 +- .../squares-integers.lagda.md | 4 +- .../squares-natural-numbers.lagda.md | 11 +++- ...strict-inequality-natural-numbers.lagda.md | 2 +- .../strictly-bounded-natural-numbers.lagda.md | 8 ++- ...tionary-functions-natural-numbers.lagda.md | 11 ++-- .../unbounded-farey-fractions.lagda.md | 12 +++- .../upper-bounds-natural-numbers.lagda.md | 34 ++++++++--- ...rdering-principle-natural-numbers.lagda.md | 2 +- ...well-ordering-principles-integers.lagda.md | 45 ++++++++++---- 38 files changed, 454 insertions(+), 158 deletions(-) diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 2c0929ddb2..752add93ec 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -41,15 +41,24 @@ open import foundation-core.injective-maps ## Idea -The {{#concept "2-adic decomposition" Agda=2-adic-decomposition-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) $n$ is a factorization of $n$ into a [power](elementary-number-theory.exponentiation-natural-numbers.md) of $2$ and an [odd](elementary-number-theory.parity-natural-numbers.md) natural number. +The {{#concept "2-adic decomposition" Agda=2-adic-decomposition-ℕ}} of a +[natural number](elementary-number-theory.natural-numbers.md) $n$ is a +factorization of $n$ into a +[power](elementary-number-theory.exponentiation-natural-numbers.md) of $2$ and +an [odd](elementary-number-theory.parity-natural-numbers.md) natural number. -The $2$-adic decomposition of the natural numbers can be used to construct an [equivalence](foundation-core.equivalences.md) $\mathbb{N}\times\mathbb{N} \simeq \mathbb{N}$ by mapping +The $2$-adic decomposition of the natural numbers can be used to construct an +[equivalence](foundation-core.equivalences.md) +$\mathbb{N}\times\mathbb{N} \simeq \mathbb{N}$ by mapping $$ (m , n) \mapsto 2^m(2n + 1) - 1. $$ -The exponent $k$ such that the 2-adic decomposition of $n$ is $2^k \cdot m=n$ is called the {{#concept "2-adic valuation" Disambiguation="natural numbers" Agda=valuation-2-adic-decomposition-ℕ}} of $n$. +The exponent $k$ such that the 2-adic decomposition of $n$ is $2^k \cdot m=n$ is +called the +{{#concept "2-adic valuation" Disambiguation="natural numbers" Agda=valuation-2-adic-decomposition-ℕ}} +of $n$. ## Definitions @@ -132,12 +141,12 @@ pr2 (pr2 (2-adic-decomposition-is-odd-ℕ n H)) = module _ (n : ℕ) (H : 1 ≤-ℕ n) where - + valuation-2-adic-decomposition-nat-ℕ : ℕ valuation-2-adic-decomposition-nat-ℕ = valuation-largest-power-divisor-ℕ 2 n star H - + div-exp-valuation-2-adic-decomposition-nat-ℕ : div-ℕ (2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) n div-exp-valuation-2-adic-decomposition-nat-ℕ = diff --git a/src/elementary-number-theory/addition-natural-numbers.lagda.md b/src/elementary-number-theory/addition-natural-numbers.lagda.md index 7cf13fa8d8..02904af394 100644 --- a/src/elementary-number-theory/addition-natural-numbers.lagda.md +++ b/src/elementary-number-theory/addition-natural-numbers.lagda.md @@ -28,9 +28,15 @@ open import foundation.sets ## Idea -The {{#concept "addition" Disambiguation="natural numbers" Agda=_+ℕ_}} operation on the [natural numbers](elementary-number-theory.natural-numbers.md) is a binary operation $m,n\mapsto m+n$ on the natural numbers given by $n$ times iteratively taking the successor of the number $m$. - -The notation $+$ for addition and $-$ for subtraction appeared first in print in 1489 in _Behende und hüpsche Rechenung auff allen Kauffmanschafft_ by [Johannes Widmann](https://en.wikipedia.org/wiki/Johannes_Widmann), for example on page 327 of {{#cite Widmann1489}}. +The {{#concept "addition" Disambiguation="natural numbers" Agda=_+ℕ_}} operation +on the [natural numbers](elementary-number-theory.natural-numbers.md) is a +binary operation $m,n\mapsto m+n$ on the natural numbers given by $n$ times +iteratively taking the successor of the number $m$. + +The notation $+$ for addition and $-$ for subtraction appeared first in print in +1489 in _Behende und hüpsche Rechenung auff allen Kauffmanschafft_ by +[Johannes Widmann](https://en.wikipedia.org/wiki/Johannes_Widmann), for example +on page 327 of {{#cite Widmann1489}}. ## Definitions diff --git a/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md index 6af203f470..ec00a48f98 100644 --- a/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bounded-divisibility-natural-numbers.lagda.md @@ -34,9 +34,25 @@ open import foundation.universe-levels ## Idea -Consider two [natural numbers](elementary-number-theory.natural-numbers.md) `m` and `n`. The {{#concept "bounded divisibility relation" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} is a [binary relation](foundation.binary-relations.md) on the type of [natural numbers](elementary-number-theory.natural-numbers.md), where we say that a number `n` is bounded divisible by `m` if there is an integer `q ≤ n` such that `q * m = n`. - -The bounded divisibility relation is a slight strengthening of the [divisibility relation](elementary-number-theory.divisibility-natural-numbers.md) by ensuring that the quotient is bounded from above by `n`. This ensures that the bounded divisibility relation is valued in [propositions](foundation-core.propositions.md) for all `m` and `n`, unlike the divisibility relation. Since the bounded divisibility relation is [logically equivalent](foundation.logical-equivalences.md) to the divisibility relation, but it is always valued in the propositions, we use the bounded divisibility relation in the definition of the [poset](order-theory.posets.md) of the [natural numbers ordered by division](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). +Consider two [natural numbers](elementary-number-theory.natural-numbers.md) `m` +and `n`. The +{{#concept "bounded divisibility relation" Disambiguation="natural numbers" Agda=bounded-div-ℕ}} +is a [binary relation](foundation.binary-relations.md) on the type of +[natural numbers](elementary-number-theory.natural-numbers.md), where we say +that a number `n` is bounded divisible by `m` if there is an integer `q ≤ n` +such that `q * m = n`. + +The bounded divisibility relation is a slight strengthening of the +[divisibility relation](elementary-number-theory.divisibility-natural-numbers.md) +by ensuring that the quotient is bounded from above by `n`. This ensures that +the bounded divisibility relation is valued in +[propositions](foundation-core.propositions.md) for all `m` and `n`, unlike the +divisibility relation. Since the bounded divisibility relation is +[logically equivalent](foundation.logical-equivalences.md) to the divisibility +relation, but it is always valued in the propositions, we use the bounded +divisibility relation in the definition of the [poset](order-theory.posets.md) +of the +[natural numbers ordered by division](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md). ## Definitions @@ -81,7 +97,9 @@ is-nonzero-quotient-bounded-div-ℕ m n N H refl = ### If a nonzero number `n` is bounded divisible by `m`, then `m` is bounded from above by `n` -**Proof.** Suppose that `q ≤ n` is such that `q * m = n`. Since `n` is assumed to be nonzero, it follows that `q` is nonzero. Therefore it follows that `m ≤ q * m = n`. +**Proof.** Suppose that `q ≤ n` is such that `q * m = n`. Since `n` is assumed +to be nonzero, it follows that `q` is nonzero. Therefore it follows that +`m ≤ q * m = n`. ```agda upper-bound-divisor-bounded-div-ℕ : @@ -140,7 +158,8 @@ concatenate-eq-bounded-div-eq-ℕ refl p refl = p ### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal -Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. +Since the quotient is defined in terms of explicit proofs of divisibility, we +assume arbitrary proofs of dibisibility on both sides. ```agda eq-quotient-bounded-div-eq-divisor-ℕ : @@ -160,7 +179,9 @@ eq-quotient-bounded-div-eq-divisor-ℕ c d n N p H I = ### If two natural numbers are equal and one is divisible by a number `d`, then the other is divisible by `d` and their quotients are the same -The first part of the claim was proven above. Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. +The first part of the claim was proven above. Since the quotient is defined in +terms of explicit proofs of divisibility, we assume arbitrary proofs of +dibisibility on both sides. ```agda eq-quotient-bounded-div-eq-is-nonzero-divisor-ℕ : @@ -366,7 +387,7 @@ upper-bound-quotient-transitive-bounded-div-ℕ (succ-ℕ m) n k K H = ( k) ( is-nonzero-succ-ℕ m) ( eq-quotient-transitive-bounded-div-ℕ (succ-ℕ m) n k K H) - + transitive-bounded-div-ℕ : is-transitive bounded-div-ℕ pr1 (transitive-bounded-div-ℕ m n k K H) = quotient-transitive-bounded-div-ℕ m n k K H diff --git a/src/elementary-number-theory/bounded-natural-numbers.lagda.md b/src/elementary-number-theory/bounded-natural-numbers.lagda.md index 99ee9966fa..830699c005 100644 --- a/src/elementary-number-theory/bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bounded-natural-numbers.lagda.md @@ -31,13 +31,16 @@ open import univalent-combinatorics.standard-finite-types ## Idea -The type of {{#concept "bounded natural numbers" Agda=bounded-ℕ}} with upper bound $n$ is the type +The type of {{#concept "bounded natural numbers" Agda=bounded-ℕ}} with upper +bound $n$ is the type $$ \mathbb{N}_{\leq n} := \{k : ℕ \mid k \leq n\}. $$ -The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) to the [standard finite type](univalent-combinatorics.standard-finite-types.md) $\mathsf{Fin}(n+1)$. +The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) +to the [standard finite type](univalent-combinatorics.standard-finite-types.md) +$\mathsf{Fin}(n+1)$. ## Definition diff --git a/src/elementary-number-theory/congruence-natural-numbers.lagda.md b/src/elementary-number-theory/congruence-natural-numbers.lagda.md index adde88a456..e3774d9c6f 100644 --- a/src/elementary-number-theory/congruence-natural-numbers.lagda.md +++ b/src/elementary-number-theory/congruence-natural-numbers.lagda.md @@ -29,15 +29,30 @@ open import univalent-combinatorics.standard-finite-types ## Idea -Two [natural numbers](elementary-number-theory.natural-numbers.md) `x` and `y` are said to be {{#concept "congruent" Disambiguation="natural numbers" Agda=cong-ℕ WDID=Q3773677 WD="congruence of integers"}} modulo `k` if their [distance](elementary-number-theory.distance-natural-numbers.md) `dist-ℕ x y` is [divisible](elementary-number-theory.divisibility-natural-numbers.md) by `k`, i.e., if +Two [natural numbers](elementary-number-theory.natural-numbers.md) `x` and `y` +are said to be +{{#concept "congruent" Disambiguation="natural numbers" Agda=cong-ℕ WDID=Q3773677 WD="congruence of integers"}} +modulo `k` if their +[distance](elementary-number-theory.distance-natural-numbers.md) `dist-ℕ x y` is +[divisible](elementary-number-theory.divisibility-natural-numbers.md) by `k`, +i.e., if ```text k | dist-ℕ x y. ``` -For each natural number `k`, the congruence relation modulo `k` defines an [equivalence relation](foundation.equivalence-relations.md). Furthermore, the congruence relations respect [addition](elementary-number-theory.addition-natural-numbers.md) and [multiplication](elementary-number-theory.multiplication-natural-numbers.md). - -[Quotienting](foundation.set-quotients.md) by the congruence relation leads to [modular arithmetic](elementary-number-theory.modular-arithmetic.md). Properties of the congruence relation with respect to the [standard finite types](univalent-combinatorics.standard-finite-types.md) are formalized in the file [`modular-arithmetic-standard-finite-types`](elementary-number-theory.modular-arithmetic-standard-finite-types.md). +For each natural number `k`, the congruence relation modulo `k` defines an +[equivalence relation](foundation.equivalence-relations.md). Furthermore, the +congruence relations respect +[addition](elementary-number-theory.addition-natural-numbers.md) and +[multiplication](elementary-number-theory.multiplication-natural-numbers.md). + +[Quotienting](foundation.set-quotients.md) by the congruence relation leads to +[modular arithmetic](elementary-number-theory.modular-arithmetic.md). Properties +of the congruence relation with respect to the +[standard finite types](univalent-combinatorics.standard-finite-types.md) are +formalized in the file +[`modular-arithmetic-standard-finite-types`](elementary-number-theory.modular-arithmetic-standard-finite-types.md). ## Properties diff --git a/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md b/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md index 1c0513909c..c7a0cd8d47 100644 --- a/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md +++ b/src/elementary-number-theory/decidable-maps-natural-numbers.lagda.md @@ -14,4 +14,9 @@ module elementary-number-theory.decidable-maps-natural-numbers where ## Idea -Consider a map $f : A \to \mathbb{N}$ into the [natural numbers](elementary-number-theory.natural-numbers.md). Then $f$ is said to be a {{#concept "decidable map" Disambiguation="natural numbers"}} if its [fibers](foundation-core.fibers-of-maps.md) are [decidable](foundation.decidable-types.md), i.e., if it is a [decidable map](foundation.decidable-maps.md). +Consider a map $f : A \to \mathbb{N}$ into the +[natural numbers](elementary-number-theory.natural-numbers.md). Then $f$ is said +to be a {{#concept "decidable map" Disambiguation="natural numbers"}} if its +[fibers](foundation-core.fibers-of-maps.md) are +[decidable](foundation.decidable-types.md), i.e., if it is a +[decidable map](foundation.decidable-maps.md). diff --git a/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md index 25c268a519..d092cf2e5f 100644 --- a/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md +++ b/src/elementary-number-theory/decidable-subtypes-natural-numbers.lagda.md @@ -39,9 +39,18 @@ open import univalent-combinatorics.decidable-subtypes ## Idea -In this file we study [decidable subtypes](foundation.decidable-subtypes.md) of the [natural numbers](elementary-number-theory.natural-numbers.md). The type of all decidable subtypes of the natural numbers is the *decidable powerset of the natural numbers*. - -As a direct consequence of the [well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) of the natural numbers, which is formulated for decidable structures over the natural numbers, it follows that every [inhabited](foundation.inhabited-subtypes.md) decidable subtype of the natural numbers has a least element. We also show that every finite decidable subtype has a largest element. +In this file we study [decidable subtypes](foundation.decidable-subtypes.md) of +the [natural numbers](elementary-number-theory.natural-numbers.md). The type of +all decidable subtypes of the natural numbers is the _decidable powerset of the +natural numbers_. + +As a direct consequence of the +[well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) +of the natural numbers, which is formulated for decidable structures over the +natural numbers, it follows that every +[inhabited](foundation.inhabited-subtypes.md) decidable subtype of the natural +numbers has a least element. We also show that every finite decidable subtype +has a largest element. ## Properties @@ -155,4 +164,3 @@ module _ ( P ∘ pr1) ( is-finite-bounded-ℕ m)) ``` - diff --git a/src/elementary-number-theory/distance-natural-numbers.lagda.md b/src/elementary-number-theory/distance-natural-numbers.lagda.md index 07288f0299..b3abae4b99 100644 --- a/src/elementary-number-theory/distance-natural-numbers.lagda.md +++ b/src/elementary-number-theory/distance-natural-numbers.lagda.md @@ -31,10 +31,17 @@ open import foundation.universe-levels The {{#concept "distance function" Disambiguation="natural numbers" Agda=dist-ℕ}} between [natural numbers](elementary-number-theory.natural-numbers.md) measures -how far two natural numbers are apart. If the [inequality](elementary-number-theory.inequality-natural-numbers.md) $x \leq y$ holds, then the distance between $x$ and $y$ is the unique natural number $d$ equipped with an [identification](foundation-core.identity-types.md) $x + d = y$. - -**Note.** In the agda-unimath library, we often -prefer to work with `dist-ℕ` over the subtraction operation, which is either partially defined or it returns nonsensical values. Not only is the distance function sensibly defined for any pair of natural numbers, but it also satisfies the more pleasant and predictable properties of a [metric](metric-spaces.metric-spaces.md). +how far two natural numbers are apart. If the +[inequality](elementary-number-theory.inequality-natural-numbers.md) $x \leq y$ +holds, then the distance between $x$ and $y$ is the unique natural number $d$ +equipped with an [identification](foundation-core.identity-types.md) +$x + d = y$. + +**Note.** In the agda-unimath library, we often prefer to work with `dist-ℕ` +over the subtraction operation, which is either partially defined or it returns +nonsensical values. Not only is the distance function sensibly defined for any +pair of natural numbers, but it also satisfies the more pleasant and predictable +properties of a [metric](metric-spaces.metric-spaces.md). ## Definition diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 53261264ce..c06fc6223a 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -67,7 +67,8 @@ then the type `div-ℕ m n` is always a The divisibility relation is [logically equivalent](foundation.logical-equivalences.md), though not [equivalent](foundation-core.equivalences.md), to the -[bounded divisibility relation](elementary-number-theory.bounded-divisibility-natural-numbers.md), which is defined by +[bounded divisibility relation](elementary-number-theory.bounded-divisibility-natural-numbers.md), +which is defined by ```text bounded-div-ℕ m n := Σ (k : ℕ), (k ≤ n) × (k *ℕ m = n). @@ -78,15 +79,24 @@ The discrepancy between divisibility and bounded divisibility is manifested at [contractible](foundation-core.contractible-types.md). For all other values we have `div-ℕ m n ≃ bounded-div-ℕ m n`. -When a natural number `n is divisible by a natural number `m`, with an element `H : div-ℕ m n`, then we define the {{#concept "quotient" Disambiguation="divisibility of natural numbers" Agda=quotient-div-ℕ}} to be the unique number `q ≤ n` such that `q * m = n`. +When a natural number +`n is divisible by a natural number `m`, with an element `H : div-ℕ m +n`, then we define the {{#concept "quotient" Disambiguation="divisibility of natural numbers" Agda=quotient-div-ℕ}} to be the unique number `q +≤ n`such that`q \* m = n`. ## Definitions ### The divisibility relation on the natural numbers -We introduce the divisibility relation on the natural numbers, and some infrastructure. +We introduce the divisibility relation on the natural numbers, and some +infrastructure. -**Note:** Perhaps a more natural name for `pr1-div-ℕ`, the number `q` such that `q * d = n`, would be `quotient-div-ℕ`. However, since the divisibility relation is not always a proposition, the quotient could have some undesirable properties. Later in this file, we will define `quotient-div-ℕ` to the the quotient of the bounded divisibility relation, which is logically equivalent to the divisibility relation. +**Note:** Perhaps a more natural name for `pr1-div-ℕ`, the number `q` such that +`q * d = n`, would be `quotient-div-ℕ`. However, since the divisibility +relation is not always a proposition, the quotient could have some undesirable +properties. Later in this file, we will define `quotient-div-ℕ` to the the +quotient of the bounded divisibility relation, which is logically equivalent to +the divisibility relation. ```agda module _ @@ -245,7 +255,8 @@ concatenate-eq-div-eq-ℕ refl p refl = p ### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal -Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. +Since the quotient is defined in terms of explicit proofs of divisibility, we +assume arbitrary proofs of dibisibility on both sides. ```agda eq-quotient-div-eq-divisor-ℕ : @@ -265,7 +276,9 @@ eq-quotient-div-eq-divisor-ℕ c d n N p H I = ### If two natural numbers are equal and one is divisible by a number `d`, then the other is divisible by `d` and their quotients are the same -The first part of the claim was proven above. Since the quotient is defined in terms of explicit proofs of divisibility, we assume arbitrary proofs of dibisibility on both sides. +The first part of the claim was proven above. Since the quotient is defined in +terms of explicit proofs of divisibility, we assume arbitrary proofs of +dibisibility on both sides. ```agda eq-quotient-div-eq-is-nonzero-divisor-ℕ : diff --git a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md index 2478117a58..c8227757b4 100644 --- a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md +++ b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md @@ -29,17 +29,34 @@ open import univalent-combinatorics.standard-finite-types ## Idea -{{#concept "Euclidean division" Agda=euclidean-division-ℕ}} is a binary operation on the [natural numbers](elementary-number-theory.natural-numbers.md) that returns the division with remainder. In other words, the Euclidean division of -`n` by `d` gives the unique pair of natural numbers `q` and `r < d` such that the identification `qd + r = n` holds. - -There are several ways of specifying Euclidean division. Since we have already defined the [congruence relations](elementary-number-theory.congruence-natural-numbers.md) independently, we can define the Euclidean division of a natural number `n` by `d` as the data consisting of: -- A natural number `r`, called the {{#concept "remainder" Disambiguation="Euclidean division of natural numbers" Agda=remainder-euclidean-division-ℕ}} +{{#concept "Euclidean division" Agda=euclidean-division-ℕ}} is a binary +operation on the [natural numbers](elementary-number-theory.natural-numbers.md) +that returns the division with remainder. In other words, the Euclidean division +of `n` by `d` gives the unique pair of natural numbers `q` and `r < d` such that +the identification `qd + r = n` holds. + +There are several ways of specifying Euclidean division. Since we have already +defined the +[congruence relations](elementary-number-theory.congruence-natural-numbers.md) +independently, we can define the Euclidean division of a natural number `n` by +`d` as the data consisting of: + +- A natural number `r`, called the + {{#concept "remainder" Disambiguation="Euclidean division of natural numbers" Agda=remainder-euclidean-division-ℕ}} - A proof that `r` is congruent to `n` modulo `d` - A proof that if `d ≠ 0`, then `r < d`. -Given that `r` is congruent to `n` modulo `d` we obtain a number `q` such that `d * q = dist-ℕ r x`, where `dist-ℕ` is the [distance function](elementary-number-theory.distance-natural-numbers.md). This number `q` is the {{#concept "quotient" Disambiguation="Euclidean divistion of naturalnumbers" Agda=quotient-euclidean-division-ℕ}} after Euclidean division. +Given that `r` is congruent to `n` modulo `d` we obtain a number `q` such that +`d * q = dist-ℕ r x`, where `dist-ℕ` is the +[distance function](elementary-number-theory.distance-natural-numbers.md). This +number `q` is the +{{#concept "quotient" Disambiguation="Euclidean divistion of naturalnumbers" Agda=quotient-euclidean-division-ℕ}} +after Euclidean division. -Note that if `d = 0`, then the unique natural number that is congruent to `n` modulo `d` is the number `n` itself. The type of this data is therefore always [contractible](foundation-core.contractible-types.md), even in the case `d = 0`. +Note that if `d = 0`, then the unique natural number that is congruent to `n` +modulo `d` is the number `n` itself. The type of this data is therefore always +[contractible](foundation-core.contractible-types.md), even in the case +`d = 0`. ## Definitions @@ -150,7 +167,8 @@ quotient-euclidean-division-ℕ' (succ-ℕ k) n = ### If `d` divides a number `n`, then its quotient by euclidean division is its quotient by division -Suppose `q * d = n`. Then the congruence class `r` of `n` modulo `d` is `0`, so the distance between `r` and `n` is `n`. +Suppose `q * d = n`. Then the congruence class `r` of `n` modulo `d` is `0`, so +the distance between `r` and `n` is `n`. ```text compute-euclidean-division-div-ℕ : diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index b9f9dfa82b..a7aaba635d 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -49,7 +49,8 @@ open import univalent-combinatorics.finite-maps The {{#concept "exponential" Agda=exp-ℕ WD="exponentiation" WDID=Q33456}} $m^n$ of two [natural numbers](elementary-number-theory.natural-numbers.md) is the -number obtained by multiplying $m$ with itself $n$ times. We use the following terminology in our formalization and naming scheme: +number obtained by multiplying $m$ with itself $n$ times. We use the following +terminology in our formalization and naming scheme: - The number $m^n$ is the **$n$th power** of $m$. - The number $m^n$ is an **exponential**. @@ -57,9 +58,12 @@ number obtained by multiplying $m$ with itself $n$ times. We use the following t - The number $m$ in $m^n$ is the **base** of the exponential. - The operation $m,n \mapsto m^n$ is called **exponentiation**. - The function $n \mapsto m^n$ is the **exponentiation function** with base $m$. -- The function $m \mapsto m^n$ is the **$n$th power function**. Specific instances are [squaring](elementary-number-theory.squares-natural-numbers.md) and [cubing](elementary-number-theory.cubes-natural-numbers.md). +- The function $m \mapsto m^n$ is the **$n$th power function**. Specific + instances are [squaring](elementary-number-theory.squares-natural-numbers.md) + and [cubing](elementary-number-theory.cubes-natural-numbers.md). -The natural numbers satisfy Tarski's high school arithmetic laws for exponentiation. +The natural numbers satisfy Tarski's high school arithmetic laws for +exponentiation. ## Definition @@ -81,7 +85,9 @@ is-power-ℕ m n = fiber (power-ℕ n) m ### Exponentiation of natural numbers -The function `exp-ℕ : ℕ → ℕ → ℕ` defines the exponentiation map $m,n \mapsto m^n$. Note that this is just the power function, with the order of its arguments swapped. +The function `exp-ℕ : ℕ → ℕ → ℕ` defines the exponentiation map +$m,n \mapsto m^n$. Note that this is just the power function, with the order of +its arguments swapped. ```agda exp-ℕ : ℕ → ℕ → ℕ @@ -490,8 +496,10 @@ leq-exponent-div-exp-ℕ m (succ-ℕ n) (succ-ℕ k) H K = ### If $m^k \mid m^n$ and $m$ is nonzero, then its quotient is $m^{d(n,k)}$ -In the case where $m=0$ we have $0^n \mid 0^n$. For any nonzero number $n$ this would give $0 \mid 0$. The quotient of $0$ divided by $0$ is the unique number $k \leq 0$ such that $k\cdot 0=0$, i.e., it is $0$. However, since $d(n,n)=0$ we have $m^{d(n,n)}=1$, which isn't the quotient under our definitions. - +In the case where $m=0$ we have $0^n \mid 0^n$. For any nonzero number $n$ this +would give $0 \mid 0$. The quotient of $0$ divided by $0$ is the unique number +$k \leq 0$ such that $k\cdot 0=0$, i.e., it is $0$. However, since $d(n,n)=0$ we +have $m^{d(n,n)}=1$, which isn't the quotient under our definitions. ```agda valuation-is-exponent-div-exp-exp-ℕ : diff --git a/src/elementary-number-theory/farey-fractions.lagda.md b/src/elementary-number-theory/farey-fractions.lagda.md index 4b754b7cd3..fd09ece4d8 100644 --- a/src/elementary-number-theory/farey-fractions.lagda.md +++ b/src/elementary-number-theory/farey-fractions.lagda.md @@ -23,9 +23,14 @@ open import foundation.universe-levels ## Idea -A {{#concept "Farey fraction" Agda=farey-fraction}} is a [reduced](elementary-number-theory.reduced-integer-fractions.md) [integer fraction](elementary-number-theory.integer-fractions.md) between $0$ and $1$ inclusive. More specifically, a Farey fraction of order $n$ is a reduced integer fraction between $0$ and $1$ whose denominator does not exceed $n$. +A {{#concept "Farey fraction" Agda=farey-fraction}} is a +[reduced](elementary-number-theory.reduced-integer-fractions.md) +[integer fraction](elementary-number-theory.integer-fractions.md) between $0$ +and $1$ inclusive. More specifically, a Farey fraction of order $n$ is a reduced +integer fraction between $0$ and $1$ whose denominator does not exceed $n$. -The Farey fractions ℱ can be inductively generated mutually with a binary relation $R$ with the following constructors: +The Farey fractions ℱ can be inductively generated mutually with a binary +relation $R$ with the following constructors: ```text 0 : ℱ @@ -37,15 +42,22 @@ The Farey fractions ℱ can be inductively generated mutually with a binary rela 𝓁 : (x y : ℱ) (r : ℛ x y) → ℛ (𝓂 x y r) y ``` -The operation $m$ returns the {{#concept "mediant" Disambiguation="Farey fractions" Agda=mediant-farey-fraction}} of two adjacent Farey fractions. The elements $0$ and $1$ in the type of Farey fractions represent the Farey fractions $0/1$ and $1/1$. Given two adjacent Farey fractions representing $a/b$ and $c/d$, the mediant of $a/b$ and $c/d$ is the Farey fraction representing +The operation $m$ returns the +{{#concept "mediant" Disambiguation="Farey fractions" Agda=mediant-farey-fraction}} +of two adjacent Farey fractions. The elements $0$ and $1$ in the type of Farey +fractions represent the Farey fractions $0/1$ and $1/1$. Given two adjacent +Farey fractions representing $a/b$ and $c/d$, the mediant of $a/b$ and $c/d$ is +the Farey fraction representing $$ \frac{a+c}{b+d}. $$ -The mediant of any two adjacent Farey fractions representing reduced fractions $a/b$ and $c/d$ represents again a reduced fraction. +The mediant of any two adjacent Farey fractions representing reduced fractions +$a/b$ and $c/d$ represents again a reduced fraction. -Farey fractions appear in Chapter 3 of {{#cite HW08}}, but they are covered in more detail in Chapter 6 of {{#cite NZM}}. +Farey fractions appear in Chapter 3 of {{#cite HW08}}, but they are covered in +more detail in Chapter 6 of {{#cite NZM}}. ## Definitions @@ -112,7 +124,8 @@ $$ bc - ad = 1. $$ -This is also known as the **characteristic property of adjacent Farey fractions**. +This is also known as the **characteristic property of adjacent Farey +fractions**. ```agda characteristic-property-adjacent-farey-fraction : @@ -169,7 +182,7 @@ characteristic-property-adjacent-farey-fraction ._ y not-adjacent-zero-farey-fraction : (x : farey-fraction) → ¬ adjacent-farey-fraction x zero-farey-fraction not-adjacent-zero-farey-fraction ._ - ( left-adjacent-mediant-farey-fraction x ._ r) = + ( left-adjacent-mediant-farey-fraction x ._ r) = not-adjacent-zero-farey-fraction x r ``` diff --git a/src/elementary-number-theory/fibonacci-sequence.lagda.md b/src/elementary-number-theory/fibonacci-sequence.lagda.md index 808a83b6f9..7f014d7c57 100644 --- a/src/elementary-number-theory/fibonacci-sequence.lagda.md +++ b/src/elementary-number-theory/fibonacci-sequence.lagda.md @@ -53,7 +53,9 @@ Fibonacci numbers are Fₙ 0 1 1 2 3 5 8 13 21 34 ``` -The Fibonacci sequence was introduced by [Leonardo Fibonacci](https://en.wikipedia.org/wiki/Fibonacci) in 1202 in his book _Liber Abaci_ {{#cite Fibonacci1202}}. +The Fibonacci sequence was introduced by +[Leonardo Fibonacci](https://en.wikipedia.org/wiki/Fibonacci) in 1202 in his +book _Liber Abaci_ {{#cite Fibonacci1202}}. ## Definitions @@ -91,7 +93,9 @@ Such a function is easy to give with induction, using the map $ℕ² → ℕ²$ takes a pair `(m,n)` to the pair `(n,n+m)`. Starting the iteration with `(0,1)` we obtain the Fibonacci sequence by taking the first projection. -Furthermore, it is possible to define the fibonacci sequence without using cartesian products, mimicing the above idea using $\mathbb{N} \to \matbb{N}$ instead of $\mathbb{N}^2$. +Furthermore, it is possible to define the fibonacci sequence without using +cartesian products, mimicing the above idea using $\mathbb{N} \to \matbb{N}$ +instead of $\mathbb{N}^2$. ```agda shift-one : @@ -385,21 +389,24 @@ le-bounded-sum-fibonacci-even-ℕ n = ### A strict bound on the growth of the Fibonacci numbers -The $n$th Fibonacci number is always strictly less than $\phi^n$, where $\phi$ is the golden ratio. This strict inequality is due to Édouard Lucas, who proved it in the late 19th century {{#cite lucas1891}}. We will prove a rational version of this claim by showing that $F_n<(\frac{b}{a})^n$ for any rational number $\frac{b}{a}>\phi$ +The $n$th Fibonacci number is always strictly less than $\phi^n$, where $\phi$ +is the golden ratio. This strict inequality is due to Édouard Lucas, who proved +it in the late 19th century {{#cite lucas1891}}. We will prove a rational +version of this claim by showing that $F_n<(\frac{b}{a})^n$ for any rational +number $\frac{b}{a}>\phi$ A variation of this claim also appears on pages 7 and 8 in section 1.2 of {{#cite Leveque12volI}} as an instructive example of a proof by induction, where -he shows that $F_n<(\frac{7}{4})^n$. This -upper bound works for any -fraction $\frac{b}{a}$ where $a(b+a) -Consider a type family $P$ over $\mathbb{N}$. A {{#concept "maximal structured natural number" Agda=maximal-element-ℕ}} in $P$ is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is an [upper bound](elementary-number-theory.upper-bounds-natural-numbers.md) for $P$. - -Forthermore, a {{#concept "bounded maximal structured natural number" Agda=bounded-maximal-element-ℕ}} in $P$ with bound $b$ is a natural number $n \leq b$ equipped with an element $p : P(n)$ such that $n$ is an upper bound of for the type family $x \mapsto (x \leq b) × P(x)$. +Consider a type family $P$ over $\mathbb{N}$. A +{{#concept "maximal structured natural number" Agda=maximal-element-ℕ}} in $P$ +is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is an +[upper bound](elementary-number-theory.upper-bounds-natural-numbers.md) for $P$. + +Forthermore, a +{{#concept "bounded maximal structured natural number" Agda=bounded-maximal-element-ℕ}} +in $P$ with bound $b$ is a natural number $n \leq b$ equipped with an element +$p : P(n)$ such that $n$ is an upper bound of for the type family +$x \mapsto (x \leq b) × P(x)$. ## Definitions diff --git a/src/elementary-number-theory/maximum-natural-numbers.lagda.md b/src/elementary-number-theory/maximum-natural-numbers.lagda.md index 22e54f6d82..0f3b2507b3 100644 --- a/src/elementary-number-theory/maximum-natural-numbers.lagda.md +++ b/src/elementary-number-theory/maximum-natural-numbers.lagda.md @@ -36,7 +36,10 @@ open import univalent-combinatorics.standard-finite-types ## Idea -The {{#concept "maximum" Disambiguation="natural numbers" Agda=max-ℕ}} of two [natural numbers](elementary-number-theory.natural-numbers.md) is a binary operation that returns the largest of two natural numbers. It is defined recursively by +The {{#concept "maximum" Disambiguation="natural numbers" Agda=max-ℕ}} of two +[natural numbers](elementary-number-theory.natural-numbers.md) is a binary +operation that returns the largest of two natural numbers. It is defined +recursively by ```text max-ℕ 0 n = n @@ -44,7 +47,9 @@ The {{#concept "maximum" Disambiguation="natural numbers" Agda=max-ℕ}} of two max-ℕ (m + 1) (n + 1) = (max m n) + 1. ``` -Furthermore, we define the maximum of any [standard finite](univalent-combinatorics.standard-finite-types.md) family of natural numbers. +Furthermore, we define the maximum of any +[standard finite](univalent-combinatorics.standard-finite-types.md) family of +natural numbers. ## Definition diff --git a/src/elementary-number-theory/mersenne-exponents.lagda.md b/src/elementary-number-theory/mersenne-exponents.lagda.md index dfd2114046..40c7af86fc 100644 --- a/src/elementary-number-theory/mersenne-exponents.lagda.md +++ b/src/elementary-number-theory/mersenne-exponents.lagda.md @@ -19,7 +19,10 @@ open import foundation.universe-levels ## Idea -A {{#concept "Mersenne exponent" Agda=mersenne-exponent-ℕ OEIS=A000043}} is a [prime number](elementary-number-theory.prime-numbers.md) $p$ such that the [Mersenne number](elementary-number-theory.mersenne-numbers.md) $M_p$ is a [Mersenne prime](elementary-number-theory.mersenne-primes.md). +A {{#concept "Mersenne exponent" Agda=mersenne-exponent-ℕ OEIS=A000043}} is a +[prime number](elementary-number-theory.prime-numbers.md) $p$ such that the +[Mersenne number](elementary-number-theory.mersenne-numbers.md) $M_p$ is a +[Mersenne prime](elementary-number-theory.mersenne-primes.md). ## Definitions diff --git a/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md index a08f2c74fa..857a021b37 100644 --- a/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md +++ b/src/elementary-number-theory/minimal-structured-natural-numbers.lagda.md @@ -27,7 +27,10 @@ open import foundation.universe-levels ## Idea -Consider a type family $P$ over $\mathbb{N}$. A {{#concept "minimal structured natural number" Agda=minimal-element-ℕ}} in $P$ is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is a [lower bound](elementary-number-theory.lower-bounds-natural-numbers.md) for $P$. +Consider a type family $P$ over $\mathbb{N}$. A +{{#concept "minimal structured natural number" Agda=minimal-element-ℕ}} in $P$ +is a natural number $n$ equipped with an element $p : P(n)$ such that $n$ is a +[lower bound](elementary-number-theory.lower-bounds-natural-numbers.md) for $P$. ## Definitions @@ -100,11 +103,12 @@ module _ ### Shifting minimal elements -If $P(0)$ is a decidable type, then we can obtain a minimal structured element of $P$ from a minimal structured element of $n\mapsto P(n+1)$. +If $P(0)$ is a decidable type, then we can obtain a minimal structured element +of $P$ from a minimal structured element of $n\mapsto P(n+1)$. ```agda module _ - {l : Level} (P : ℕ → UU l) + {l : Level} (P : ℕ → UU l) where shift-minimal-element-ℕ : diff --git a/src/elementary-number-theory/natural-numbers.lagda.md b/src/elementary-number-theory/natural-numbers.lagda.md index bf7fe0271b..af0b132b43 100644 --- a/src/elementary-number-theory/natural-numbers.lagda.md +++ b/src/elementary-number-theory/natural-numbers.lagda.md @@ -22,9 +22,16 @@ open import foundation-core.negation ## Idea The {{#concept "natural numbers" WD="natural number" WDID=Q21199 Agda=ℕ}} is an -inductive type generated by the {{#concept "zero" Disambiguation="natural number" Agda=zero-ℕ WDID=Q204 WD="zero"}} element and the {{#concept "successor function" Disambiguation="natural numbers" Agda=succ-ℕ WDID=Q7632653 WD="successor function"}}. The -{{#concept "induction principle" Disambiguation="natural numbers" Agda=ind-ℕ WDID=Q178377 WD="mathematical induction"}} for the natural numbers works to construct sections of type -families over the natural numbers. Another name for the induction principle for the natural numbers is {{#concept "mathematical induction" Agda=ind-ℕ WDID=Q178377 WD="mathematical induction"}}. +inductive type generated by the +{{#concept "zero" Disambiguation="natural number" Agda=zero-ℕ WDID=Q204 WD="zero"}} +element and the +{{#concept "successor function" Disambiguation="natural numbers" Agda=succ-ℕ WDID=Q7632653 WD="successor function"}}. +The +{{#concept "induction principle" Disambiguation="natural numbers" Agda=ind-ℕ WDID=Q178377 WD="mathematical induction"}} +for the natural numbers works to construct sections of type families over the +natural numbers. Another name for the induction principle for the natural +numbers is +{{#concept "mathematical induction" Agda=ind-ℕ WDID=Q178377 WD="mathematical induction"}}. ## Definitions diff --git a/src/elementary-number-theory/number-of-divisors.lagda.md b/src/elementary-number-theory/number-of-divisors.lagda.md index 8ac0801357..59b02e14df 100644 --- a/src/elementary-number-theory/number-of-divisors.lagda.md +++ b/src/elementary-number-theory/number-of-divisors.lagda.md @@ -215,7 +215,8 @@ pr2 (has-fixed-point-quotient-divisor-is-square-ℕ n H) = ### The type of fixed points of the involution on the type of divisors is a proposition -This essentially claims that only the square root of a square can be the fixed point of the involution on the type of divisors. +This essentially claims that only the square root of a square can be the fixed +point of the involution on the type of divisors. ```agda all-elements-equal-fixed-point-quotient-divisor-ℕ : @@ -239,4 +240,11 @@ is-prop-fixed-point-quotient-divisor-ℕ n = ## Remarks -In the properties above we have shown that the quotient operation on the type of divisors of any natural number $n$ is an involution with at most one fixed point, and it has a fixed point if and only if $n$ is square. This implies that the number of divisors is odd if and only if $n$ is a square. However, in the agda-unimath library we don't have the appropriate infrastructure yet for counting the elements of types with involutions. The conclusion that the number of divisors of $n$ is odd if and only if $n$ is a square will be formalized in the future. +In the properties above we have shown that the quotient operation on the type of +divisors of any natural number $n$ is an involution with at most one fixed +point, and it has a fixed point if and only if $n$ is square. This implies that +the number of divisors is odd if and only if $n$ is a square. However, in the +agda-unimath library we don't have the appropriate infrastructure yet for +counting the elements of types with involutions. The conclusion that the number +of divisors of $n$ is odd if and only if $n$ is a square will be formalized in +the future. diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 663055785a..96454b4e97 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -370,7 +370,8 @@ is-odd-is-1-mod-2-ℕ (succ-ℕ n) H K = ### If a successor number `n + 1` is even, then its quotient after division by `2` is at most `n` -**Proof.** Suppose that `q * 2 = n + 1` for some natural number `q`. Then `q` is a successor, say `q = q' + 1`. It follows that +**Proof.** Suppose that `q * 2 = n + 1` for some natural number `q`. Then `q` +is a successor, say `q = q' + 1`. It follows that ```text q + 1 ≤ q + q' + 1 = q + q = q * 2 = n + 1. @@ -519,6 +520,7 @@ is-nonzero-is-odd-ℕ .zero-ℕ H refl = H is-even-zero-ℕ Further laws of parity are proven in other files, e.g.: - [Parity of integers](elementary-number-theory.parity-integers.md) -- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md). Here we also show that the sum of the first $n$ odd numbers is $n^2$. +- [Squares of natural numbers](elementary-number-theory.squares-natural-numbers.md). + Here we also show that the sum of the first $n$ odd numbers is $n^2$. - The fact that the pronic numbers $n(n+1)$ are even is proven in [Pronic numbers](elementary-number-theory.pronic-numbers.md). diff --git a/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md b/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md index e0fd9653a5..9f6af0c203 100644 --- a/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md +++ b/src/elementary-number-theory/poset-of-natural-numbers-ordered-by-divisibility.lagda.md @@ -33,10 +33,20 @@ open import order-theory.preorders The **poset of natural numbers ordered by divisibility** consists of the [natural numbers](elementary-number-theory.natural-numbers.md) and its ordering -is defined by [bounded divisibility](elementary-number-theory.bounded-divisibility-natural-numbers.md), i.e., the type `m ≤ n` is defined to be the type of natural numbers `q ≤ n` such that `q * m = n`. - -Recall that bounded divisibility is [logically equivalent](foundation.logical-equivalences.md) to the more standard [divisibility relation](elementar-number-theory.divisibility-natural-numbers.md). However, the divisibility relation `m | n` is only valued in the [propositions](foundation.propositions.md) when both `m` and `n` -are [nonzero](elementary-number-theory.nonzero-natural-numbers.md). On the other hand, bounded divisibility is always valued in propositions. By using bounded divisibility we avoid the need for [propoositional truncation](foundation.propositional-truncations.md). +is defined by +[bounded divisibility](elementary-number-theory.bounded-divisibility-natural-numbers.md), +i.e., the type `m ≤ n` is defined to be the type of natural numbers `q ≤ n` such +that `q * m = n`. + +Recall that bounded divisibility is +[logically equivalent](foundation.logical-equivalences.md) to the more standard +[divisibility relation](elementar-number-theory.divisibility-natural-numbers.md). +However, the divisibility relation `m | n` is only valued in the +[propositions](foundation.propositions.md) when both `m` and `n` are +[nonzero](elementary-number-theory.nonzero-natural-numbers.md). On the other +hand, bounded divisibility is always valued in propositions. By using bounded +divisibility we avoid the need for +[propoositional truncation](foundation.propositional-truncations.md). ## Definition diff --git a/src/elementary-number-theory/prime-numbers.lagda.md b/src/elementary-number-theory/prime-numbers.lagda.md index c553df8fbb..d045439b73 100644 --- a/src/elementary-number-theory/prime-numbers.lagda.md +++ b/src/elementary-number-theory/prime-numbers.lagda.md @@ -39,7 +39,10 @@ open import foundation.universe-levels ## Idea -A {{#concept "prime number" Agda=is-prime-ℕ WDID=Q49008 WD="prime number"}} is a [natural number](elementary-number-theory.natural-numbers.md) of which 1 is the only [proper divisor](elementary-number-theory.proper-divisors-natural-numbers.md). +A {{#concept "prime number" Agda=is-prime-ℕ WDID=Q49008 WD="prime number"}} is a +[natural number](elementary-number-theory.natural-numbers.md) of which 1 is the +only +[proper divisor](elementary-number-theory.proper-divisors-natural-numbers.md). ## Definition @@ -265,7 +268,8 @@ strict-upper-bound-quotient-div-is-prime-ℕ p x N P H = ### If a prime number `p` divides a number `x + 1`, then `(x + 1)/p ≤ x` -Note that this upper bound is slightly sharper than the usual upper bound `x + 1` we get for arbitrary quotients of divisible natural numbers. +Note that this upper bound is slightly sharper than the usual upper bound +`x + 1` we get for arbitrary quotients of divisible natural numbers. ```agda upper-bound-quotient-div-is-prime-ℕ : diff --git a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md index 833bae1a15..508de8c27f 100644 --- a/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/proper-divisors-natural-numbers.lagda.md @@ -32,7 +32,10 @@ open import foundation.universe-levels ## Idea -A {{#concept "proper divisor" Disambituation="natural numbers" Agda=is-proper-divisor-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a natural number `d ≠ n` that +A +{{#concept "proper divisor" Disambituation="natural numbers" Agda=is-proper-divisor-ℕ}} +of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a +natural number `d ≠ n` that [divides](elementary-number-theory.divisbility-natural-numbers.md) `n`. ## Definitions diff --git a/src/elementary-number-theory/rational-numbers.lagda.md b/src/elementary-number-theory/rational-numbers.lagda.md index b72c2a284b..fbb9ce84e7 100644 --- a/src/elementary-number-theory/rational-numbers.lagda.md +++ b/src/elementary-number-theory/rational-numbers.lagda.md @@ -311,8 +311,10 @@ pr2 reflecting-map-sim-fraction {x} {y} H = eq-ℚ-sim-fraction-ℤ x y H ## See also -- [Farey fractions](elementary-number-theory.farey-fractions.md) encode the type of rational numbers between $0$ and $1$ by a mutual recursive construction. -- [Unbounded Farey fractions](elementary-number-theory.unbounded-farey-fractions.md) encode the type of all rational numbers by a mutal recursive construction. +- [Farey fractions](elementary-number-theory.farey-fractions.md) encode the type + of rational numbers between $0$ and $1$ by a mutual recursive construction. +- [Unbounded Farey fractions](elementary-number-theory.unbounded-farey-fractions.md) + encode the type of all rational numbers by a mutal recursive construction. ## References diff --git a/src/elementary-number-theory/squares-integers.lagda.md b/src/elementary-number-theory/squares-integers.lagda.md index cde52a27cf..04d300199e 100644 --- a/src/elementary-number-theory/squares-integers.lagda.md +++ b/src/elementary-number-theory/squares-integers.lagda.md @@ -39,8 +39,8 @@ open import foundation-core.transport-along-identifications ## Idea -The {{#concept "square" Disambiguation="integer" Agda=square-ℤ}} `a²` of -an [integer](elementary-number-theory.integers.md) `a` is the +The {{#concept "square" Disambiguation="integer" Agda=square-ℤ}} `a²` of an +[integer](elementary-number-theory.integers.md) `a` is the [product](elementary-number-theory.multiplication-natural-numbers.md) ```text diff --git a/src/elementary-number-theory/squares-natural-numbers.lagda.md b/src/elementary-number-theory/squares-natural-numbers.lagda.md index d9d98e9d3c..098953bf8d 100644 --- a/src/elementary-number-theory/squares-natural-numbers.lagda.md +++ b/src/elementary-number-theory/squares-natural-numbers.lagda.md @@ -350,7 +350,8 @@ is-decidable-is-square-ℕ n = ### Any number divides its own square -In other words, the squaring function is inflationary with respect to divisibility. +In other words, the squaring function is inflationary with respect to +divisibility. ```agda is-inflationary-bounded-div-square-ℕ : @@ -588,9 +589,13 @@ square-add-ℕ m n = ### The formula for the distance between squares -The formula for the distance between squares is more commonly known as the formula for the difference of squares. However, since we prefer using the distance operation on the natural numbers over the partial difference operation, we will state and prove the analogous formula using the distance function. +The formula for the distance between squares is more commonly known as the +formula for the difference of squares. However, since we prefer using the +distance operation on the natural numbers over the partial difference operation, +we will state and prove the analogous formula using the distance function. -The formula for the difference of squares of integers is formalized in its usual form. +The formula for the difference of squares of integers is formalized in its usual +form. ```agda distance-of-squares-ℕ' : diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 9654af34fd..bc50599917 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -220,7 +220,7 @@ subtraction-le-ℕ zero-ℕ m p = subtraction-le-ℕ (succ-ℕ n) (succ-ℕ m) p = ( pr1 P , pr1 (pr2 P) , ap succ-ℕ (pr2 (pr2 P))) where - + P : Σ ℕ (λ l' → (is-nonzero-ℕ l') × (l' +ℕ n = m)) P = subtraction-le-ℕ n m p diff --git a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md index 24af1afb01..c8f78bc0ef 100644 --- a/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-bounded-natural-numbers.lagda.md @@ -29,13 +29,17 @@ open import univalent-combinatorics.standard-finite-types ## Idea -The type of {{#concept "strictly bounded natural numbers" Agda=strictly-bounded-ℕ}} with strict upper bound $n$ is the type +The type of +{{#concept "strictly bounded natural numbers" Agda=strictly-bounded-ℕ}} with +strict upper bound $n$ is the type $$ \mathbb{N}_{< n} := \{k : ℕ \mid k < n\}. $$ -The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) to the [standard finite type](univalent-combinatorics.standard-finite-types.md) $\mathsf{Fin}(n+1)$. +The type $\mathbb{N}_{\leq n}$ is [equivalent](foundation-core.equivalences.md) +to the [standard finite type](univalent-combinatorics.standard-finite-types.md) +$\mathsf{Fin}(n+1)$. ## Definition diff --git a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md index 421cb742c5..729df17246 100644 --- a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md @@ -31,13 +31,17 @@ open import univalent-combinatorics.dependent-pair-types ## Idea -A function $f : \mathbb{N} \to \mathbb{N}$ is a {{#concept "strictly inflationary function" Disambiguation="natural numbers" Agda=is-strictly-inflationary-ℕ}} if the implication +A function $f : \mathbb{N} \to \mathbb{N}$ is a +{{#concept "strictly inflationary function" Disambiguation="natural numbers" Agda=is-strictly-inflationary-ℕ}} +if the implication $$ (x < y) \to (f(x) < f(y)) $$ -holds for every $x,y:\mathbb{N}$. If $f$ is a strictly inflationary function on the natural numbers such that $f(0) \leq b$ for some natural number $b$, then there is a maximal natural number $k$ such that $f(k) \leq b$. +holds for every $x,y:\mathbb{N}$. If $f$ is a strictly inflationary function on +the natural numbers such that $f(0) \leq b$ for some natural number $b$, then +there is a maximal natural number $k$ such that $f(k) \leq b$. ## Definitions @@ -105,6 +109,3 @@ is-decidable-map-strictly-inflationary-map-ℕ : is-decidable-map-strictly-inflationary-map-ℕ (f , H) n = is-decidable-map-is-strictly-inflationary-map-ℕ H n ``` - - - diff --git a/src/elementary-number-theory/unbounded-farey-fractions.lagda.md b/src/elementary-number-theory/unbounded-farey-fractions.lagda.md index be7be51e73..c4f4c98d97 100644 --- a/src/elementary-number-theory/unbounded-farey-fractions.lagda.md +++ b/src/elementary-number-theory/unbounded-farey-fractions.lagda.md @@ -16,9 +16,15 @@ open import foundation.universe-levels ## Idea -[Farey fractions](elementary-number-theory.farey-fractions.md) are a way of encoding all [rational numbers](elementary-number-theory.rational-numbers.md) between $0$ and $1$ inclusive. This idea can be generalized to encode all rational numbers. - -The type of {{#concept "unbounded Farey fractions" Agda=unbounded-farey-fraction}} is an inductive type that is mutually defined with an adjacency relation, given with the following constructors: +[Farey fractions](elementary-number-theory.farey-fractions.md) are a way of +encoding all [rational numbers](elementary-number-theory.rational-numbers.md) +between $0$ and $1$ inclusive. This idea can be generalized to encode all +rational numbers. + +The type of +{{#concept "unbounded Farey fractions" Agda=unbounded-farey-fraction}} is an +inductive type that is mutually defined with an adjacency relation, given with +the following constructors: ```text 𝒾 : ℤ → ℱ diff --git a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md index 330b0a6084..1d8cc032d7 100644 --- a/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md +++ b/src/elementary-number-theory/upper-bounds-natural-numbers.lagda.md @@ -31,14 +31,29 @@ open import foundation.universe-levels ## Idea -Consider a type family $P$ over the [natural numbers](elementary-number-theory.natural-numbers.md). A {{#concept "structured natural number"}} is simply a natural number $n$ equipped with an element $P(n)$. In this file we consider various upper bounds for structured natural numbers, and relations between them. This file builds the prerequisite infrastructure for the [well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) of the natural numbers, and its direct consequences. - -- A natural number $n$ is said to be an {{#concept "upper bound" Disambiguation="structured natural numbers" Agda=is-upper-bound-ℕ}} if there is a -function from $P(x)$ to the type $x \leq n$ for all $x : \mathbb{N}$. -- A natural number $n$ is said to be a {{#concept "strict upper bound" Disambiguation="structured natural numbers" Agda=is-strict-upper-bound-ℕ}} if there is a -function from $P(x)$ to the type $x < n$ for all $x : \mathbb{N}$. -- A natural number $n$ is said to be a {{#concept "least upper bound" Disambiguation="structured natural numbers" Agda=is-least-upper-bound-ℕ}} if any natural number $x$ is an upper bound if and only if $n \leq x$. -- A natural number $n$ is said to be a {{#concept "least strict upper bound" Disambiguation="structured natural numbers" Agda=is-least-strict-upper-bound-ℕ}} if any natural number $x$ is a strict upper bound if and only if $n \leq x$. +Consider a type family $P$ over the +[natural numbers](elementary-number-theory.natural-numbers.md). A +{{#concept "structured natural number"}} is simply a natural number $n$ equipped +with an element $P(n)$. In this file we consider various upper bounds for +structured natural numbers, and relations between them. This file builds the +prerequisite infrastructure for the +[well-ordering principle](elementary-number-theory.well-ordering-principle-natural-numbers.md) +of the natural numbers, and its direct consequences. + +- A natural number $n$ is said to be an + {{#concept "upper bound" Disambiguation="structured natural numbers" Agda=is-upper-bound-ℕ}} + if there is a function from $P(x)$ to the type $x \leq n$ for all + $x : \mathbb{N}$. +- A natural number $n$ is said to be a + {{#concept "strict upper bound" Disambiguation="structured natural numbers" Agda=is-strict-upper-bound-ℕ}} + if there is a function from $P(x)$ to the type $x < n$ for all + $x : \mathbb{N}$. +- A natural number $n$ is said to be a + {{#concept "least upper bound" Disambiguation="structured natural numbers" Agda=is-least-upper-bound-ℕ}} + if any natural number $x$ is an upper bound if and only if $n \leq x$. +- A natural number $n$ is said to be a + {{#concept "least strict upper bound" Disambiguation="structured natural numbers" Agda=is-least-strict-upper-bound-ℕ}} + if any natural number $x$ is a strict upper bound if and only if $n \leq x$. ## Definitions @@ -295,7 +310,8 @@ is-upper-bound-leq-is-upper-bound-ℕ P b H n K x p = ### Being an upper bound of a decidable type family is decidable, given an upper bound -The type family `is-upper-bound-ℕ P` is an example of a type family `Q` over the natural numbers satisfying +The type family `is-upper-bound-ℕ P` is an example of a type family `Q` over the +natural numbers satisfying ```text Q b → Π (n : ℕ), is-decidable (Q n) diff --git a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md index 87c948ddba..2a0a08a875 100644 --- a/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principle-natural-numbers.lagda.md @@ -75,7 +75,7 @@ well-ordering-principle-ℕ' P d (inr f) (succ-ℕ n) p = module _ {l : Level} (P : ℕ → UU l) (d : is-decidable-fam P) {n : ℕ} (p : P n) where - + well-ordering-principle-ℕ : minimal-element-ℕ P well-ordering-principle-ℕ = diff --git a/src/elementary-number-theory/well-ordering-principles-integers.lagda.md b/src/elementary-number-theory/well-ordering-principles-integers.lagda.md index 6c8b215669..bb24df943f 100644 --- a/src/elementary-number-theory/well-ordering-principles-integers.lagda.md +++ b/src/elementary-number-theory/well-ordering-principles-integers.lagda.md @@ -33,12 +33,27 @@ open import foundation.universe-levels ## Idea -The [poset of integers](elementary-number-theory.inequality-integers.md) does not satisfy the [well-ordering principle of the natural numbers](elementary-number-theory.well-ordering-principle-natural-numbers.md). However, there are several useful approximations of the well-ordering principle for the [integers](elementary-number-theory.integers.md) that are worth recording: - -- For every family `A` of [decidable types](foundation.decidable-types.md) over `ℤ` that -is bounded from below and comes equipped with an integer `a` and an element of type `A a`, we can find a least integer `a₀` equipped with an element of type `A a₀`. Here, a type family `A` over the integers is bounded from below if it comes equipped with a [lower bound](order-theory.lower-bounds-posets.md) `b` such that `b ≤ x` for all `x : ℤ` equipped with an element of type `A x`. -- For every family `A` of [decidable types](foundation.decidable-types.md) over `ℤ` that -is bounded from above and comes equipped with an integer `a` and an element of type `A a`, we can find a largest integer `a₀` equipped with an element of type `A a₀`. Here, a type family `A` over the integers is bounded from above if it comes equipped with an [upper bound](order-theory.lower-bounds-posets.md) `b` such that `x ≤ b` for all `x : ℤ` equipped with an element of type `A x`. +The [poset of integers](elementary-number-theory.inequality-integers.md) does +not satisfy the +[well-ordering principle of the natural numbers](elementary-number-theory.well-ordering-principle-natural-numbers.md). +However, there are several useful approximations of the well-ordering principle +for the [integers](elementary-number-theory.integers.md) that are worth +recording: + +- For every family `A` of [decidable types](foundation.decidable-types.md) over + `ℤ` that is bounded from below and comes equipped with an integer `a` and an + element of type `A a`, we can find a least integer `a₀` equipped with an + element of type `A a₀`. Here, a type family `A` over the integers is bounded + from below if it comes equipped with a + [lower bound](order-theory.lower-bounds-posets.md) `b` such that `b ≤ x` for + all `x : ℤ` equipped with an element of type `A x`. +- For every family `A` of [decidable types](foundation.decidable-types.md) over + `ℤ` that is bounded from above and comes equipped with an integer `a` and an + element of type `A a`, we can find a largest integer `a₀` equipped with an + element of type `A a₀`. Here, a type family `A` over the integers is bounded + from above if it comes equipped with an + [upper bound](order-theory.lower-bounds-posets.md) `b` such that `x ≤ b` for + all `x : ℤ` equipped with an element of type `A x`. ## Definitions @@ -119,7 +134,12 @@ module _ ### The well-ordering principle for the nonnegative integers -The {{#concept "well-ordering principle for the nonnegative integers" Agda=well-ordering-principle-nonnegative-ℤ}} states that for any family `A` of decidable types over `ℤ` such that the implication `A x → is-nonnegative-ℤ x` holds for any `x : ℤ`, if we are given an element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an element of type `A a₀`. +The +{{#concept "well-ordering principle for the nonnegative integers" Agda=well-ordering-principle-nonnegative-ℤ}} +states that for any family `A` of decidable types over `ℤ` such that the +implication `A x → is-nonnegative-ℤ x` holds for any `x : ℤ`, if we are given an +element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an +element of type `A a₀`. ```agda module _ @@ -136,10 +156,10 @@ module _ pr2 (pr2 (well-ordering-principle-nonnegative-ℤ' H (n , x , K))) b y = concatenate-leq-eq-ℤ (int-ℕ n) (leq-int-ℕ n m (K m (tr A (inv α) y))) α where - + m : ℕ m = nat-nonnegative-ℤ (b , H b y) - + α : int-ℕ m = b α = ap pr1 (is-section-nat-nonnegative-ℤ (b , H b y)) @@ -162,7 +182,12 @@ module _ ### The well-ordering principle for the positive integers -The {{#concept "well-ordering principle for the positive integers" Agda=well-ordering-principle-positive-ℤ}} states that for any family `A` of decidable types over `ℤ` such that the implication `A x → is-positive-ℤ x` holds for any `x : ℤ`, if we are given an element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an element of type `A a₀`. +The +{{#concept "well-ordering principle for the positive integers" Agda=well-ordering-principle-positive-ℤ}} +states that for any family `A` of decidable types over `ℤ` such that the +implication `A x → is-positive-ℤ x` holds for any `x : ℤ`, if we are given an +element `(a , a) : Σ ℤ A` then there is a minimal integer `a₀` equipped with an +element of type `A a₀`. ```agda module _ From 06f229d65fb374cbbc8f4287ffa8d818cdd31c15 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 26 Jan 2025 20:52:45 -0500 Subject: [PATCH 58/72] refactoring divisibility --- references.bib | 16 +- .../divisibility-natural-numbers.lagda.md | 444 +++++++++--------- ...uclidean-division-natural-numbers.lagda.md | 15 +- ...st-common-divisor-natural-numbers.lagda.md | 40 +- 4 files changed, 253 insertions(+), 262 deletions(-) diff --git a/references.bib b/references.bib index acfbebb92c..da25accd47 100644 --- a/references.bib +++ b/references.bib @@ -1,4 +1,4 @@ -@online{1000+theorems, +G@online{1000+theorems, title = {1000+ theorems}, author = {Freek Wiedijk}, url = {https://1000-plus.github.io/} @@ -410,13 +410,13 @@ @online{GGMS24 } @book{HW08, - title = {{An Introduction to the Theory of Numbers (6th edition)}}, - author = {Hardy, G. H. and Wright, Edward M. and Heath-Brown, D. R. and Silverman, Joseph H.}, - abstract = {{The sixth edition of the classic undergraduate text in elementary number theory includes a new chapter on elliptic curves and their role in the proof of Fermat's Last Theorem, a foreword by Andrew Wiles and extensively revised and updated end-of-chapter notes.}}, - isbn = {0199219869}, - keywords = {congruences, primitive roots, residue systems, instructional exposition}, - publisher = {Oxford University Press}, - year = 2008 + title = {{An Introduction to the Theory of Numbers (6th edition)}}, + author = {Hardy, G. H. and Wright, Edward M. and Heath-Brown, D. R. and Silverman, Joseph H.}, + isbn = {0199219869}, + keywords = {congruences, primitive roots, residue systems, instructional exposition}, + publisher = {Oxford University Press}, + year = {2008}, + abstract = {{The sixth edition of the classic undergraduate text in elementary number theory includes a new chapter on elliptic curves and their role in the proof of Fermat's Last Theorem, a foreword by Andrew Wiles and extensively revised and updated end-of-chapter notes.}}, } @article{KECA17, diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index c06fc6223a..2a3dc3dfff 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -26,6 +26,7 @@ open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.equivalences +open import foundation.function-types open import foundation.identity-types open import foundation.logical-equivalences open import foundation.negated-equality @@ -119,6 +120,18 @@ module _ div-bounded-div-ℕ : bounded-div-ℕ m n → div-ℕ pr1 (div-bounded-div-ℕ H) = quotient-bounded-div-ℕ m n H pr2 (div-bounded-div-ℕ H) = eq-quotient-bounded-div-ℕ m n H + +concatenate-eq-div-ℕ : + {x y z : ℕ} → x = y → div-ℕ y z → div-ℕ x z +concatenate-eq-div-ℕ refl p = p + +concatenate-div-eq-ℕ : + {x y z : ℕ} → div-ℕ x y → y = z → div-ℕ x z +concatenate-div-eq-ℕ p refl = p + +concatenate-eq-div-eq-ℕ : + {x y z w : ℕ} → x = y → div-ℕ y z → z = w → div-ℕ x w +concatenate-eq-div-eq-ℕ refl p refl = p ``` ### The predicate of being a divisor of a natural number @@ -130,7 +143,9 @@ is-divisor-ℕ n x = div-ℕ x n ## Properties -### If `n` is divisible by a number `m` with proof of divisibility `(q , p)`, then `n` is divisible by the number `q`. +### Divisibility and bounded divisibility are logically equivalent + +#### If `n` is divisible by a number `m` with proof of divisibility `(q , p)`, then `n` is divisible by the number `q`. ```agda div-pr1-div-ℕ : @@ -139,7 +154,7 @@ pr1 (div-pr1-div-ℕ m n (u , p)) = m pr2 (div-pr1-div-ℕ m n (u , p)) = commutative-mul-ℕ m u ∙ p ``` -### If `x` is nonzero and `d | x`, then `d ≤ x` +#### If `x` is nonzero and `d | x`, then `d ≤ x` ```agda leq-div-succ-ℕ : (d x : ℕ) → div-ℕ d (succ-ℕ x) → leq-ℕ d (succ-ℕ x) @@ -154,28 +169,9 @@ leq-pr1-div-ℕ : (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → leq-ℕ (pr1-div-ℕ d x H) x leq-pr1-div-ℕ d x f H = leq-div-ℕ (pr1-div-ℕ d x H) x f (div-pr1-div-ℕ d x H) - -leq-pr1-div-ℕ' : - (d x : ℕ) → is-nonzero-ℕ d → (H : div-ℕ d x) → leq-ℕ (pr1-div-ℕ d x H) x -leq-pr1-div-ℕ' d zero-ℕ f (zero-ℕ , p) = star -leq-pr1-div-ℕ' d zero-ℕ f (succ-ℕ n , p) = - f (is-zero-right-is-zero-add-ℕ _ d p) -leq-pr1-div-ℕ' d (succ-ℕ x) f H = - leq-pr1-div-ℕ d (succ-ℕ x) (is-nonzero-succ-ℕ x) H -``` - -### If `x` is nonzero, if `d | x` and `d ≠ x`, then `d < x` - -```agda -le-div-succ-ℕ : - (d x : ℕ) → div-ℕ d (succ-ℕ x) → d ≠ succ-ℕ x → le-ℕ d (succ-ℕ x) -le-div-succ-ℕ d x H f = le-leq-neq-ℕ (leq-div-succ-ℕ d x H) f - -le-div-ℕ : (d x : ℕ) → is-nonzero-ℕ x → div-ℕ d x → d ≠ x → le-ℕ d x -le-div-ℕ d x H K f = le-leq-neq-ℕ (leq-div-ℕ d x H K) f ``` -### Divisibility is logically equivalent to bounded divisibility +#### The logical equivalence of divisibility and bounded divisibility ```agda bounded-div-div-ℕ : (m n : ℕ) → div-ℕ m n → bounded-div-ℕ m n @@ -191,7 +187,7 @@ pr2 (logical-equivalence-bounded-div-div-ℕ m n) = bounded-div-div-ℕ m n ``` -### The quotient of a natural number by a number it is divisible by +#### The quotient of a natural number by a number it is divisible by ```agda module _ @@ -237,22 +233,6 @@ compute-quotient-div-ℕ q p H K = ( bounded-div-div-ℕ _ _ K) ``` -### Concatenating equality and divisibility - -```agda -concatenate-eq-div-ℕ : - {x y z : ℕ} → x = y → div-ℕ y z → div-ℕ x z -concatenate-eq-div-ℕ refl p = p - -concatenate-div-eq-ℕ : - {x y z : ℕ} → div-ℕ x y → y = z → div-ℕ x z -concatenate-div-eq-ℕ p refl = p - -concatenate-eq-div-eq-ℕ : - {x y z w : ℕ} → x = y → div-ℕ y z → z = w → div-ℕ x w -concatenate-eq-div-eq-ℕ refl p refl = p -``` - ### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal Since the quotient is defined in terms of explicit proofs of divisibility, we @@ -260,8 +240,7 @@ assume arbitrary proofs of dibisibility on both sides. ```agda eq-quotient-div-eq-divisor-ℕ : - (c d n : ℕ) → is-nonzero-ℕ c → c = d → - (H : div-ℕ c n) → (I : div-ℕ d n) → + (c d n : ℕ) → is-nonzero-ℕ c → c = d → (H : div-ℕ c n) (I : div-ℕ d n) → pr1-div-ℕ c n H = pr1-div-ℕ d n I eq-quotient-div-eq-divisor-ℕ c d n N p H I = is-injective-left-mul-ℕ c N @@ -301,36 +280,6 @@ eq-quotient-div-eq-is-nonzero-ℕ {succ-ℕ d} N refl H K = ( is-nonzero-succ-ℕ d) refl H K ``` -### Divisibility is decidable - -```agda -is-decidable-div-ℕ : - (m n : ℕ) → is-decidable (div-ℕ m n) -is-decidable-div-ℕ m n = - is-decidable-iff - ( div-bounded-div-ℕ m n) - ( bounded-div-div-ℕ m n) - ( is-decidable-bounded-div-ℕ m n) -``` - -### Divisibility is a property except at `(0,0)` - -```agda -is-prop-div-ℕ : - (k x : ℕ) → is-nonzero-ℕ k + is-nonzero-ℕ x → is-prop (div-ℕ k x) -is-prop-div-ℕ k x (inl H) = is-prop-map-is-emb (is-emb-right-mul-ℕ k H) x -is-prop-div-ℕ zero-ℕ x (inr H) = - is-prop-is-proof-irrelevant - ( λ (q , p) → ex-falso (H (inv p ∙ right-zero-law-mul-ℕ q))) -is-prop-div-ℕ (succ-ℕ k) x (inr H) = - is-prop-map-is-emb (is-emb-right-mul-ℕ (succ-ℕ k) (is-nonzero-succ-ℕ k)) x - -div-ℕ-Decidable-Prop : (d x : ℕ) → is-nonzero-ℕ d → Decidable-Prop lzero -pr1 (div-ℕ-Decidable-Prop d x H) = div-ℕ d x -pr1 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-prop-div-ℕ d x (inl H) -pr2 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-decidable-div-ℕ d x -``` - ### The divisibility relation is a partial order on the natural numbers The [poset](order-theory.posets.md) of natural numbers ordered by divisibility @@ -404,7 +353,7 @@ div-is-zero-ℕ : div-is-zero-ℕ k .zero-ℕ refl = div-zero-ℕ k ``` -### `0 | x` implies `x = 0` and `x | 1` implies `x = 1` +### `0 | x` implies `x = 0` ```agda is-zero-div-zero-ℕ : (x : ℕ) → div-ℕ zero-ℕ x → is-zero-ℕ x @@ -412,6 +361,21 @@ is-zero-div-zero-ℕ x H = antisymmetric-div-ℕ x zero-ℕ (div-zero-ℕ x) H is-zero-is-zero-div-ℕ : (x y : ℕ) → div-ℕ x y → is-zero-ℕ x → is-zero-ℕ y is-zero-is-zero-div-ℕ .zero-ℕ y d refl = is-zero-div-zero-ℕ y d + +is-zero-quotient-div-is-zero-dividend-ℕ : + (x y : ℕ) (H : div-ℕ x y) → is-zero-ℕ y → is-zero-ℕ (quotient-div-ℕ x y H) +is-zero-quotient-div-is-zero-dividend-ℕ x ._ H refl = + is-zero-leq-zero-ℕ + ( quotient-div-ℕ x 0 H) + ( upper-bound-quotient-div-ℕ x 0 H) + +is-zero-quotient-div-is-zero-divisor-ℕ : + (x y : ℕ) (H : div-ℕ x y) → is-zero-ℕ x → is-zero-ℕ (quotient-div-ℕ x y H) +is-zero-quotient-div-is-zero-divisor-ℕ x y H p = + is-zero-quotient-div-is-zero-dividend-ℕ x y H + ( ( inv (eq-quotient-div-ℕ x y H)) ∙ + ( ap (quotient-div-ℕ x y H *ℕ_) p) ∙ + ( right-zero-law-mul-ℕ (quotient-div-ℕ x y H))) ``` ### Any divisor of a nonzero number is nonzero @@ -422,6 +386,16 @@ is-nonzero-div-ℕ : is-nonzero-div-ℕ .zero-ℕ x H K refl = K (is-zero-div-zero-ℕ x H) ``` +### If `d` divides a nonzero number `x`, then the quotient `x/d` is also nonzero + +```agda +is-nonzero-quotient-div-ℕ : + {d x : ℕ} (H : div-ℕ d x) → + is-nonzero-ℕ x → is-nonzero-ℕ (quotient-div-ℕ d x H) +is-nonzero-quotient-div-ℕ {d} {x} H K = + is-nonzero-quotient-bounded-div-ℕ d x K (bounded-div-div-ℕ d x H) +``` + ### Any divisor of a number at least `1` is at least `1` ```agda @@ -431,6 +405,24 @@ leq-one-div-ℕ d x H L = leq-one-is-nonzero-ℕ d (is-nonzero-div-ℕ d x H (is-nonzero-leq-one-ℕ x L)) ``` +### If `d` divides a number `1 ≤ x`, then `1 ≤ x/d` + +```agda +leq-one-quotient-div-ℕ : + (d x : ℕ) (H : div-ℕ d x) → leq-ℕ 1 x → leq-ℕ 1 (quotient-div-ℕ d x H) +leq-one-quotient-div-ℕ d x H K = + leq-one-div-ℕ + ( quotient-div-ℕ d x H) + ( x) + ( div-quotient-div-ℕ d x H) + ( K) + +leq-one-quotient-div-is-nonzero-ℕ : + (d x : ℕ) (H : div-ℕ d x) → is-nonzero-ℕ x → leq-ℕ 1 (quotient-div-ℕ d x H) +leq-one-quotient-div-is-nonzero-ℕ d x H N = + leq-one-quotient-div-ℕ d x H (leq-one-is-nonzero-ℕ x N) +``` + ### If `x < d` and `d | x`, then `x` must be `0` ```agda @@ -445,13 +437,36 @@ is-zero-div-ℕ d (succ-ℕ x) H (succ-ℕ k , p) = ( leq-add-ℕ' d (k *ℕ d)) p)) ``` +### `a/a = 1` + +```agda +is-idempotent-quotient-div-ℕ : + (a : ℕ) → is-nonzero-ℕ a → (H : div-ℕ a a) → is-one-ℕ (quotient-div-ℕ a a H) +is-idempotent-quotient-div-ℕ zero-ℕ nz (u , p) = ex-falso (nz refl) +is-idempotent-quotient-div-ℕ (succ-ℕ a) nz (u , p) = + is-one-is-left-unit-mul-ℕ u a p +``` + +### If `x` is nonzero, if `d | x` and `d ≠ x`, then `d < x` + +```agda +le-div-succ-ℕ : + (d x : ℕ) → div-ℕ d (succ-ℕ x) → d ≠ succ-ℕ x → le-ℕ d (succ-ℕ x) +le-div-succ-ℕ d x H f = le-leq-neq-ℕ (leq-div-succ-ℕ d x H) f + +le-div-ℕ : (d x : ℕ) → is-nonzero-ℕ x → div-ℕ d x → d ≠ x → le-ℕ d x +le-div-ℕ d x H K f = le-leq-neq-ℕ (leq-div-ℕ d x H K) f +``` + ### If `x` divides `y` then `x` divides any multiple of `y` ```agda div-mul-ℕ : (k x y : ℕ) → div-ℕ x y → div-ℕ x (k *ℕ y) -div-mul-ℕ k x y H = - transitive-div-ℕ x y (k *ℕ y) (k , refl) H +pr1 (div-mul-ℕ k x y H) = k *ℕ quotient-div-ℕ x y H +pr2 (div-mul-ℕ k x y H) = + associative-mul-ℕ k (quotient-div-ℕ x y H) x ∙ + ap (k *ℕ_) (eq-quotient-div-ℕ x y H) div-mul-ℕ' : (k x y : ℕ) → div-ℕ x y → div-ℕ x (y *ℕ k) @@ -459,6 +474,89 @@ div-mul-ℕ' k x y H = tr (div-ℕ x) (commutative-mul-ℕ k y) (div-mul-ℕ k x y H) ``` +### Divisibility is decidable + +```agda +is-decidable-div-ℕ : + (m n : ℕ) → is-decidable (div-ℕ m n) +is-decidable-div-ℕ m n = + is-decidable-iff + ( div-bounded-div-ℕ m n) + ( bounded-div-div-ℕ m n) + ( is-decidable-bounded-div-ℕ m n) +``` + +### Divisibility is a property except at `(0,0)` + +```agda +is-prop-div-ℕ : + (k x : ℕ) → is-nonzero-ℕ k + is-nonzero-ℕ x → is-prop (div-ℕ k x) +is-prop-div-ℕ k x (inl H) = is-prop-map-is-emb (is-emb-right-mul-ℕ k H) x +is-prop-div-ℕ zero-ℕ x (inr H) = + is-prop-is-proof-irrelevant + ( λ (q , p) → ex-falso (H (inv p ∙ right-zero-law-mul-ℕ q))) +is-prop-div-ℕ (succ-ℕ k) x (inr H) = + is-prop-map-is-emb (is-emb-right-mul-ℕ (succ-ℕ k) (is-nonzero-succ-ℕ k)) x + +div-ℕ-Decidable-Prop : (d x : ℕ) → is-nonzero-ℕ d → Decidable-Prop lzero +pr1 (div-ℕ-Decidable-Prop d x H) = div-ℕ d x +pr1 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-prop-div-ℕ d x (inl H) +pr2 (pr2 (div-ℕ-Decidable-Prop d x H)) = is-decidable-div-ℕ d x +``` + +### If `b` divides `a` and `c` divides `b`, then `a/b · b/c = a/c` + +We use our convention that $0/0 = 0$ to avoid the hypothesis that $c$ is nonzero. + +```agda +simplify-mul-quotient-div-ℕ : + (a b c : ℕ) → + (H : div-ℕ b a) (K : div-ℕ c b) (L : div-ℕ c a) → + quotient-div-ℕ b a H *ℕ quotient-div-ℕ c b K = quotient-div-ℕ c a L +simplify-mul-quotient-div-ℕ a b zero-ℕ H K L = + ( ap + ( quotient-div-ℕ b a H *ℕ_) + ( is-zero-quotient-div-is-zero-divisor-ℕ 0 b K refl)) ∙ + ( right-zero-law-mul-ℕ (quotient-div-ℕ b a H)) ∙ + ( inv (is-zero-quotient-div-is-zero-divisor-ℕ 0 a L refl)) +simplify-mul-quotient-div-ℕ a b (succ-ℕ c) H K L = + is-injective-right-mul-ℕ (succ-ℕ c) (is-nonzero-succ-ℕ c) + ( equational-reasoning + (a/b *ℕ b/c+1) *ℕ succ-ℕ c + = a/b *ℕ (b/c+1 *ℕ succ-ℕ c) + by associative-mul-ℕ a/b b/c+1 (succ-ℕ c) + = a/b *ℕ b + by ap (a/b *ℕ_) (eq-quotient-div-ℕ (succ-ℕ c) b K) + = a + by eq-quotient-div-ℕ b a H + = a/c+1 *ℕ succ-ℕ c + by inv (eq-quotient-div-ℕ (succ-ℕ c) a L)) + where + a/b : ℕ + a/b = quotient-div-ℕ b a H + b/c+1 : ℕ + b/c+1 = quotient-div-ℕ (succ-ℕ c) b K + a/c+1 : ℕ + a/c+1 = quotient-div-ℕ (succ-ℕ c) a L +``` + +### Suppose `b | a` and `c | b`. If `d` divides `b/c` then `d` divides `a/c` + +```agda +div-quotient-div-div-quotient-div-ℕ : + {a b c d : ℕ} (H : div-ℕ b a) (K : div-ℕ c b) (L : div-ℕ c a) → + div-ℕ d (quotient-div-ℕ c b K) → div-ℕ d (quotient-div-ℕ c a L) +div-quotient-div-div-quotient-div-ℕ {a} {b} {c} {d} H K L M = + tr + ( div-ℕ d) + ( simplify-mul-quotient-div-ℕ a b c H K L) + ( div-mul-ℕ + ( quotient-div-ℕ b a H) + ( d) + ( quotient-div-ℕ c b K) + ( M)) +``` + ### A 3-for-2 property of division with respect to addition ```agda @@ -529,28 +627,43 @@ preserves-div-right-mul-ℕ k x y H = ### Multiplication by a nonzero number reflects divisibility ```agda -reflects-div-mul-ℕ : - (k x y : ℕ) → is-nonzero-ℕ k → div-ℕ (k *ℕ x) (k *ℕ y) → div-ℕ x y -pr1 (reflects-div-mul-ℕ k x y H (q , p)) = q -pr2 (reflects-div-mul-ℕ k x y H (q , p)) = - is-injective-left-mul-ℕ k H - ( ( inv (associative-mul-ℕ k q x)) ∙ - ( ( ap (_*ℕ x) (commutative-mul-ℕ k q)) ∙ - ( ( associative-mul-ℕ q k x) ∙ - ( p)))) +reflects-div-left-mul-ℕ : + (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (m *ℕ a) (m *ℕ b) → div-ℕ a b +pr1 (reflects-div-left-mul-ℕ m a b H K) = quotient-div-ℕ (m *ℕ a) (m *ℕ b) K +pr2 (reflects-div-left-mul-ℕ m a b H K) = + is-injective-left-mul-ℕ m H + ( ( left-swap-mul-ℕ m + ( quotient-div-ℕ (m *ℕ a) (m *ℕ b) K) + ( a)) ∙ + ( eq-quotient-div-ℕ (m *ℕ a) (m *ℕ b) K)) + +reflects-div-right-mul-ℕ : + (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (a *ℕ m) (b *ℕ m) → div-ℕ a b +reflects-div-right-mul-ℕ m a b H K = + reflects-div-left-mul-ℕ m a b H + ( concatenate-eq-div-eq-ℕ + ( commutative-mul-ℕ m a) + ( K) + ( commutative-mul-ℕ b m)) ``` -### If a nonzero number `d` divides `y`, then `dx` divides `y` if and only if `x` divides the quotient `y/d` +### If `d` divides `y`, then `dx` divides `y` if and only if `x` divides the quotient `y/d` + +This logical equivalence holds for all integers, using the convention that `0/0 = 0`. ```agda div-quotient-div-div-ℕ : - (x y d : ℕ) (H : div-ℕ d y) → is-nonzero-ℕ d → + (x y d : ℕ) (H : div-ℕ d y) → div-ℕ (d *ℕ x) y → div-ℕ x (quotient-div-ℕ d y H) -div-quotient-div-div-ℕ x y d H f K = - reflects-div-mul-ℕ d x - ( quotient-div-ℕ d y H) - ( f) - ( tr (div-ℕ (d *ℕ x)) (inv (eq-quotient-div-ℕ' d y H)) K) +div-quotient-div-div-ℕ x y zero-ℕ H K = + div-is-zero-ℕ x + ( quotient-div-ℕ 0 y H) + ( is-zero-quotient-div-is-zero-divisor-ℕ 0 y H refl) +div-quotient-div-div-ℕ x y (succ-ℕ d) H K = + reflects-div-left-mul-ℕ (succ-ℕ d) x + ( quotient-div-ℕ (succ-ℕ d) y H) + ( is-nonzero-succ-ℕ d) + ( tr (div-ℕ (succ-ℕ d *ℕ x)) (inv (eq-quotient-div-ℕ' (succ-ℕ d) y H)) K) div-div-quotient-div-ℕ : (x y d : ℕ) (H : div-ℕ d y) → @@ -560,120 +673,22 @@ div-div-quotient-div-ℕ x y d H K = ( div-ℕ (d *ℕ x)) ( eq-quotient-div-ℕ' d y H) ( preserves-div-left-mul-ℕ d x (quotient-div-ℕ d y H) K) -``` - -### If `d` divides a nonzero number `x`, then the quotient `x/d` is also nonzero - -```agda -is-nonzero-quotient-div-ℕ : - {d x : ℕ} (H : div-ℕ d x) → - is-nonzero-ℕ x → is-nonzero-ℕ (quotient-div-ℕ d x H) -is-nonzero-quotient-div-ℕ {d} {x} H K = - is-nonzero-quotient-bounded-div-ℕ d x K (bounded-div-div-ℕ d x H) -``` - -### If `d` divides a number `1 ≤ x`, then `1 ≤ x/d` - -```agda -leq-one-quotient-div-ℕ : - (d x : ℕ) (H : div-ℕ d x) → leq-ℕ 1 x → leq-ℕ 1 (quotient-div-ℕ d x H) -leq-one-quotient-div-ℕ d x H K = - leq-one-div-ℕ - ( quotient-div-ℕ d x H) - ( x) - ( div-quotient-div-ℕ d x H) - ( K) - -leq-one-quotient-div-is-nonzero-ℕ : - (d x : ℕ) (H : div-ℕ d x) → is-nonzero-ℕ x → leq-ℕ 1 (quotient-div-ℕ d x H) -leq-one-quotient-div-is-nonzero-ℕ d x H N = - leq-one-quotient-div-ℕ d x H (leq-one-is-nonzero-ℕ x N) -``` - -### `a/a = 1` - -```agda -is-idempotent-quotient-div-ℕ : - (a : ℕ) → is-nonzero-ℕ a → (H : div-ℕ a a) → is-one-ℕ (quotient-div-ℕ a a H) -is-idempotent-quotient-div-ℕ zero-ℕ nz (u , p) = ex-falso (nz refl) -is-idempotent-quotient-div-ℕ (succ-ℕ a) nz (u , p) = - is-one-is-left-unit-mul-ℕ u a p -``` - -### If `b` divides `a` and `c` divides `b` and `c` is nonzero, then `a/b · b/c = a/c` - -```agda -simplify-mul-quotient-div-ℕ : - {a b c : ℕ} → is-nonzero-ℕ c → - (H : div-ℕ b a) (K : div-ℕ c b) (L : div-ℕ c a) → - quotient-div-ℕ b a H *ℕ quotient-div-ℕ c b K = quotient-div-ℕ c a L -simplify-mul-quotient-div-ℕ {a} {b} {c} nz H K L = - is-injective-right-mul-ℕ c nz - ( equational-reasoning - (a/b *ℕ b/c) *ℕ c - = a/b *ℕ (b/c *ℕ c) - by associative-mul-ℕ a/b b/c c - = a/b *ℕ b - by ap (a/b *ℕ_) (eq-quotient-div-ℕ c b K) - = a - by eq-quotient-div-ℕ b a H - = a/c *ℕ c - by inv (eq-quotient-div-ℕ c a L)) - where - a/b : ℕ - a/b = quotient-div-ℕ b a H - b/c : ℕ - b/c = quotient-div-ℕ c b K - a/c : ℕ - a/c = quotient-div-ℕ c a L -``` - -### If `d | a` and `d` is nonzero, then `x | a/d` if and only if `xd | a` -```agda simplify-div-quotient-div-ℕ : - {a d x : ℕ} → is-nonzero-ℕ d → (H : div-ℕ d a) → - (div-ℕ x (quotient-div-ℕ d a H)) ↔ (div-ℕ (x *ℕ d) a) -pr1 (pr1 (simplify-div-quotient-div-ℕ nz H) (u , p)) = u -pr2 (pr1 (simplify-div-quotient-div-ℕ {a} {d} {x} nz H) (u , p)) = - equational-reasoning - u *ℕ (x *ℕ d) - = (u *ℕ x) *ℕ d - by inv (associative-mul-ℕ u x d) - = (quotient-div-ℕ d a H) *ℕ d - by ap (_*ℕ d) p - = a - by eq-quotient-div-ℕ d a H -pr1 (pr2 (simplify-div-quotient-div-ℕ nz H) (u , p)) = u -pr2 (pr2 (simplify-div-quotient-div-ℕ {a} {d} {x} nz H) (u , p)) = - is-injective-right-mul-ℕ d nz - ( equational-reasoning - (u *ℕ x) *ℕ d - = u *ℕ (x *ℕ d) - by associative-mul-ℕ u x d - = a - by p - = (quotient-div-ℕ d a H) *ℕ d - by inv (eq-quotient-div-ℕ d a H)) -``` - -### Suppose `H : b | a` and `K : c | b`, where `c` is nonzero. If `d` divides `b/c` then `d` divides `a/c` - -```agda -div-quotient-div-div-quotient-div-ℕ : - {a b c d : ℕ} → is-nonzero-ℕ c → (H : div-ℕ b a) - (K : div-ℕ c b) (L : div-ℕ c a) → - div-ℕ d (quotient-div-ℕ c b K) → - div-ℕ d (quotient-div-ℕ c a L) -div-quotient-div-div-quotient-div-ℕ {a} {b} {c} {d} nz H K L M = - tr - ( div-ℕ d) - ( simplify-mul-quotient-div-ℕ nz H K L) - ( div-mul-ℕ - ( quotient-div-ℕ b a H) - ( d) - ( quotient-div-ℕ c b K) - ( M)) + {a d x : ℕ} (H : div-ℕ d a) → + div-ℕ x (quotient-div-ℕ d a H) ↔ div-ℕ (d *ℕ x) a +pr1 (simplify-div-quotient-div-ℕ {a} {d} {x} H) = + div-div-quotient-div-ℕ x a d H +pr2 (simplify-div-quotient-div-ℕ {a} {d} {x} H) = + div-quotient-div-div-ℕ x a d H + +simplify-div-quotient-div-ℕ' : + {a d x : ℕ} (H : div-ℕ d a) → + div-ℕ x (quotient-div-ℕ d a H) ↔ div-ℕ (x *ℕ d) a +pr1 (simplify-div-quotient-div-ℕ' {a} {d} {x} H) = + tr (λ u → div-ℕ u a) (commutative-mul-ℕ d x) ∘ div-div-quotient-div-ℕ x a d H +pr2 (simplify-div-quotient-div-ℕ' {a} {d} {x} H) = + div-quotient-div-div-ℕ x a d H ∘ tr (λ u → div-ℕ u a) (commutative-mul-ℕ x d) ``` ### If `x` is nonzero and `d | x`, then `x/d = 1` if and only if `d = x` @@ -714,7 +729,7 @@ is-one-divisor-ℕ d (succ-ℕ x) N (.(succ-ℕ x) , q) refl = ```agda le-quotient-div-ℕ : - (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → ¬ (is-one-ℕ d) → + (d x : ℕ) → is-nonzero-ℕ x → (H : div-ℕ d x) → ¬ is-one-ℕ d → le-ℕ (quotient-div-ℕ d x H) x le-quotient-div-ℕ d x f H g = map-left-unit-law-coproduct-is-empty @@ -727,29 +742,6 @@ le-quotient-div-ℕ d x f H g = ( upper-bound-quotient-div-ℕ d x H)) ``` -### If `m` is nonzero, then `a | b` if and only if `ma | mb` - -```agda -reflects-div-left-mul-ℕ : - (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (m *ℕ a) (m *ℕ b) → div-ℕ a b -pr1 (reflects-div-left-mul-ℕ m a b H K) = quotient-div-ℕ (m *ℕ a) (m *ℕ b) K -pr2 (reflects-div-left-mul-ℕ m a b H K) = - is-injective-left-mul-ℕ m H - ( ( left-swap-mul-ℕ m - ( quotient-div-ℕ (m *ℕ a) (m *ℕ b) K) - ( a)) ∙ - ( eq-quotient-div-ℕ (m *ℕ a) (m *ℕ b) K)) - -reflects-div-right-mul-ℕ : - (m a b : ℕ) → is-nonzero-ℕ m → div-ℕ (a *ℕ m) (b *ℕ m) → div-ℕ a b -reflects-div-right-mul-ℕ m a b H K = - reflects-div-left-mul-ℕ m a b H - ( concatenate-eq-div-eq-ℕ - ( commutative-mul-ℕ m a) - ( K) - ( commutative-mul-ℕ b m)) -``` - ## See also - [Bounded divisibility of natural numbers](elementary-number-theory.bounded-divisibility-natural-numbers.md) diff --git a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md index c8227757b4..075827b7ed 100644 --- a/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md +++ b/src/elementary-number-theory/euclidean-division-natural-numbers.lagda.md @@ -64,7 +64,7 @@ modulo `d` is the number `n` itself. The type of this data is therefore always ```agda euclidean-division-ℕ : - (k x : ℕ) → Σ ℕ (λ r → (cong-ℕ k x r) × (is-nonzero-ℕ k → le-ℕ r k)) + (k x : ℕ) → Σ ℕ (λ r → (x ≡ r mod-ℕ k) × (is-nonzero-ℕ k → le-ℕ r k)) pr1 (euclidean-division-ℕ zero-ℕ x) = x pr1 (pr2 (euclidean-division-ℕ zero-ℕ x)) = refl-cong-ℕ zero-ℕ x pr2 (pr2 (euclidean-division-ℕ zero-ℕ x)) f = ex-falso (f refl) @@ -83,7 +83,7 @@ remainder-euclidean-division-ℕ k x = pr1 (euclidean-division-ℕ k x) cong-euclidean-division-ℕ : - (k x : ℕ) → cong-ℕ k x (remainder-euclidean-division-ℕ k x) + (k x : ℕ) → x ≡ remainder-euclidean-division-ℕ k x mod-ℕ k cong-euclidean-division-ℕ k x = pr1 (pr2 (euclidean-division-ℕ k x)) @@ -98,17 +98,16 @@ quotient-euclidean-division-ℕ k x = eq-quotient-euclidean-division-ℕ : (k x : ℕ) → - ( (quotient-euclidean-division-ℕ k x) *ℕ k) = + ( quotient-euclidean-division-ℕ k x *ℕ k) = ( dist-ℕ x (remainder-euclidean-division-ℕ k x)) eq-quotient-euclidean-division-ℕ k x = pr2 (cong-euclidean-division-ℕ k x) eq-euclidean-division-ℕ : (k x : ℕ) → - ( add-ℕ - ( (quotient-euclidean-division-ℕ k x) *ℕ k) - ( remainder-euclidean-division-ℕ k x)) = - ( x) + quotient-euclidean-division-ℕ k x *ℕ k +ℕ + remainder-euclidean-division-ℕ k x = + x eq-euclidean-division-ℕ zero-ℕ x = ( inv ( ap @@ -117,7 +116,7 @@ eq-euclidean-division-ℕ zero-ℕ x = ( left-unit-law-add-ℕ x) eq-euclidean-division-ℕ (succ-ℕ k) x = ( ap - ( _+ℕ (remainder-euclidean-division-ℕ (succ-ℕ k) x)) + ( _+ℕ remainder-euclidean-division-ℕ (succ-ℕ k) x) ( ( pr2 (cong-euclidean-division-ℕ (succ-ℕ k) x)) ∙ ( symmetric-dist-ℕ x ( remainder-euclidean-division-ℕ (succ-ℕ k) x)))) ∙ diff --git a/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md b/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md index 0ae0e6fd99..846b4a93c9 100644 --- a/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md +++ b/src/elementary-number-theory/greatest-common-divisor-natural-numbers.lagda.md @@ -373,8 +373,8 @@ reflects-is-common-divisor-mul-ℕ : is-common-divisor-ℕ a b d reflects-is-common-divisor-mul-ℕ k a b d H = map-product - ( reflects-div-mul-ℕ k d a H) - ( reflects-div-mul-ℕ k d b H) + ( reflects-div-left-mul-ℕ k d a H) + ( reflects-div-left-mul-ℕ k d b H) ``` ### `gcd-ℕ 1 b = 1` @@ -423,54 +423,54 @@ is-id-is-gcd-zero-ℕ' {a} {x} H = is-id-is-gcd-zero-ℕ {a} {x} ```agda is-common-divisor-quotients-div-quotient-ℕ : - {a b d e n : ℕ} → is-nonzero-ℕ e → (H : is-common-divisor-ℕ a b d) + {a b d e n : ℕ} → (H : is-common-divisor-ℕ a b d) (K : div-ℕ e d) → div-ℕ n (quotient-div-ℕ e d K) → (M : is-common-divisor-ℕ a b e) → is-common-divisor-ℕ ( quotient-div-ℕ e a (pr1 M)) ( quotient-div-ℕ e b (pr2 M)) ( n) -pr1 (is-common-divisor-quotients-div-quotient-ℕ nz H K L M) = - div-quotient-div-div-quotient-div-ℕ nz (pr1 H) K (pr1 M) L -pr2 (is-common-divisor-quotients-div-quotient-ℕ nz H K L M) = - div-quotient-div-div-quotient-div-ℕ nz (pr2 H) K (pr2 M) L +pr1 (is-common-divisor-quotients-div-quotient-ℕ H K L M) = + div-quotient-div-div-quotient-div-ℕ (pr1 H) K (pr1 M) L +pr2 (is-common-divisor-quotients-div-quotient-ℕ H K L M) = + div-quotient-div-div-quotient-div-ℕ (pr2 H) K (pr2 M) L simplify-is-common-divisor-quotient-div-ℕ : - {a b d x : ℕ} → is-nonzero-ℕ d → (H : is-common-divisor-ℕ a b d) → + {a b d x : ℕ} (H : is-common-divisor-ℕ a b d) → is-common-divisor-ℕ ( quotient-div-ℕ d a (pr1 H)) ( quotient-div-ℕ d b (pr2 H)) ( x) ↔ is-common-divisor-ℕ a b (x *ℕ d) -pr1 (pr1 (simplify-is-common-divisor-quotient-div-ℕ nz H) K) = - forward-implication (simplify-div-quotient-div-ℕ nz (pr1 H)) (pr1 K) -pr2 (pr1 (simplify-is-common-divisor-quotient-div-ℕ nz H) K) = - forward-implication (simplify-div-quotient-div-ℕ nz (pr2 H)) (pr2 K) -pr1 (pr2 (simplify-is-common-divisor-quotient-div-ℕ nz H) K) = - backward-implication (simplify-div-quotient-div-ℕ nz (pr1 H)) (pr1 K) -pr2 (pr2 (simplify-is-common-divisor-quotient-div-ℕ nz H) K) = - backward-implication (simplify-div-quotient-div-ℕ nz (pr2 H)) (pr2 K) +pr1 (pr1 (simplify-is-common-divisor-quotient-div-ℕ H) K) = + forward-implication (simplify-div-quotient-div-ℕ' (pr1 H)) (pr1 K) +pr2 (pr1 (simplify-is-common-divisor-quotient-div-ℕ H) K) = + forward-implication (simplify-div-quotient-div-ℕ' (pr2 H)) (pr2 K) +pr1 (pr2 (simplify-is-common-divisor-quotient-div-ℕ H) K) = + backward-implication (simplify-div-quotient-div-ℕ' (pr1 H)) (pr1 K) +pr2 (pr2 (simplify-is-common-divisor-quotient-div-ℕ H) K) = + backward-implication (simplify-div-quotient-div-ℕ' (pr2 H)) (pr2 K) ``` ### The greatest common divisor of `a/d` and `b/d` is `gcd(a,b)/d` ```agda is-gcd-quotient-div-gcd-ℕ : - {a b d : ℕ} → is-nonzero-ℕ d → (H : is-common-divisor-ℕ a b d) → + {a b d : ℕ} (H : is-common-divisor-ℕ a b d) → is-gcd-ℕ ( quotient-div-ℕ d a (pr1 H)) ( quotient-div-ℕ d b (pr2 H)) ( quotient-div-ℕ d ( gcd-ℕ a b) ( div-gcd-is-common-divisor-ℕ a b d H)) -is-gcd-quotient-div-gcd-ℕ {a} {b} {d} nz H x = +is-gcd-quotient-div-gcd-ℕ {a} {b} {d} H x = logical-equivalence-reasoning is-common-divisor-ℕ ( quotient-div-ℕ d a (pr1 H)) ( quotient-div-ℕ d b (pr2 H)) ( x) ↔ is-common-divisor-ℕ a b (x *ℕ d) - by simplify-is-common-divisor-quotient-div-ℕ nz H + by simplify-is-common-divisor-quotient-div-ℕ H ↔ div-ℕ (x *ℕ d) (gcd-ℕ a b) by is-gcd-gcd-ℕ a b (x *ℕ d) ↔ div-ℕ x @@ -479,7 +479,7 @@ is-gcd-quotient-div-gcd-ℕ {a} {b} {d} nz H x = ( div-gcd-is-common-divisor-ℕ a b d H)) by inv-iff - ( simplify-div-quotient-div-ℕ nz + ( simplify-div-quotient-div-ℕ' ( div-gcd-is-common-divisor-ℕ a b d H)) ``` From c49bd091f2a6f5d6e49ec6975fe8f2105300b8f2 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 3 Feb 2025 14:43:02 -0500 Subject: [PATCH 59/72] typo --- references.bib | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/references.bib b/references.bib index 7898f0d4d7..6a3a187f68 100644 --- a/references.bib +++ b/references.bib @@ -1,4 +1,4 @@ -G@online{1000+theorems, +@online{1000+theorems, title = {1000+ theorems}, author = {Freek Wiedijk}, url = {https://1000-plus.github.io/} From c60e8b112e88efec45b9447dc07113ab06e1c620 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 3 Feb 2025 15:06:51 -0500 Subject: [PATCH 60/72] the largest power divisor is the largest power divisor --- ...st-power-divisors-natural-numbers.lagda.md | 42 +++++++++++++++---- 1 file changed, 33 insertions(+), 9 deletions(-) diff --git a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md index db024818bc..99262bcf48 100644 --- a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md @@ -18,6 +18,7 @@ open import elementary-number-theory.upper-bounds-natural-numbers open import foundation.cartesian-product-types open import foundation.dependent-pair-types +open import foundation.identity-types open import foundation.logical-equivalences open import foundation.universe-levels ``` @@ -41,6 +42,9 @@ of the form $m^k$ that ```agda is-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero is-power-divisor-ℕ m n x = is-exponential-ℕ m x × div-ℕ x n + +valuation-is-power-divisor-ℕ : (m n x : ℕ) → is-power-divisor-ℕ m n x → ℕ +valuation-is-power-divisor-ℕ m n x H = pr1 (pr1 H) ``` ### The predicate of being the largest power divisor of a natural number @@ -48,7 +52,11 @@ is-power-divisor-ℕ m n x = is-exponential-ℕ m x × div-ℕ x n ```agda is-largest-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero is-largest-power-divisor-ℕ m n x = - (y : ℕ) → is-power-divisor-ℕ m n y ↔ div-ℕ y x + Σ ( is-power-divisor-ℕ m n x) + ( λ H → + (y : ℕ) (K : is-power-divisor-ℕ m n y) → + valuation-is-power-divisor-ℕ m n y K ≤-ℕ + valuation-is-power-divisor-ℕ m n x H) ``` ### The largest power divisor of a natural number @@ -79,22 +87,36 @@ module _ ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) ( largest-power-divisor-ℕ) - upper-bound-exp-valuation-largest-power-divisor-ℕ : + nat-largest-power-divisor-ℕ : + ℕ + nat-largest-power-divisor-ℕ = + m ^ℕ valuation-largest-power-divisor-ℕ + + upper-bound-nat-largest-power-divisor-ℕ : m ^ℕ valuation-largest-power-divisor-ℕ ≤-ℕ n - upper-bound-exp-valuation-largest-power-divisor-ℕ = + upper-bound-nat-largest-power-divisor-ℕ = pr1 ( structure-maximal-element-ℕ ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) ( largest-power-divisor-ℕ)) - div-exp-valuation-largest-power-divisor-ℕ : + div-largest-power-divisor-ℕ : div-ℕ (m ^ℕ valuation-largest-power-divisor-ℕ) n - div-exp-valuation-largest-power-divisor-ℕ = + div-largest-power-divisor-ℕ = pr2 ( structure-maximal-element-ℕ ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) ( largest-power-divisor-ℕ)) + is-power-divisor-largest-power-divisor-ℕ : + is-power-divisor-ℕ m n nat-largest-power-divisor-ℕ + pr1 (pr1 is-power-divisor-largest-power-divisor-ℕ) = + valuation-largest-power-divisor-ℕ + pr2 (pr1 is-power-divisor-largest-power-divisor-ℕ) = + refl + pr2 is-power-divisor-largest-power-divisor-ℕ = + div-largest-power-divisor-ℕ + is-upper-bound-valuation-largest-power-divisor-ℕ : (k : ℕ) → div-ℕ (m ^ℕ k) n → k ≤-ℕ valuation-largest-power-divisor-ℕ is-upper-bound-valuation-largest-power-divisor-ℕ k L = @@ -104,10 +126,12 @@ module _ ( k) ( leq-div-ℕ (m ^ℕ k) n (is-nonzero-leq-one-ℕ n K) L , L) - is-largest-power-divisor-valuation-largest-power-divisor-ℕ : - is-largest-power-divisor-ℕ m n valuation-largest-power-divisor-ℕ - is-largest-power-divisor-valuation-largest-power-divisor-ℕ = - {!!} + is-largest-power-divisor-largest-power-divisor-ℕ : + is-largest-power-divisor-ℕ m n nat-largest-power-divisor-ℕ + pr1 is-largest-power-divisor-largest-power-divisor-ℕ = + is-power-divisor-largest-power-divisor-ℕ + pr2 is-largest-power-divisor-largest-power-divisor-ℕ y ((k , refl) , K) = + is-upper-bound-valuation-largest-power-divisor-ℕ k K ``` ## See also From 160188a543660cc9f7ce273816a56abea1d8987f Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Mon, 3 Feb 2025 16:04:32 -0500 Subject: [PATCH 61/72] parity results --- .../2-adic-decomposition.lagda.md | 42 ++++++++++--- .../parity-natural-numbers.lagda.md | 62 +++++++++++++++++++ 2 files changed, 95 insertions(+), 9 deletions(-) diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 752add93ec..0dc6a06ffe 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -83,6 +83,10 @@ module _ valuation-2-adic-decomposition-ℕ : ℕ valuation-2-adic-decomposition-ℕ = pr1 d + exp-valuation-2-adic-decomposition-ℕ : ℕ + exp-valuation-2-adic-decomposition-ℕ = + 2 ^ℕ valuation-2-adic-decomposition-ℕ + index-odd-factor-2-adic-decomposition-ℕ : ℕ index-odd-factor-2-adic-decomposition-ℕ = pr1 (pr2 d) @@ -96,6 +100,34 @@ module _ index-odd-factor-2-adic-decomposition-ℕ = n eq-2-adic-decomposition-ℕ = pr2 (pr2 d) + + div-exp-valuation-2-adic-decomposition-ℕ : + div-ℕ exp-valuation-2-adic-decomposition-ℕ n + pr1 div-exp-valuation-2-adic-decomposition-ℕ = + odd-factor-2-adic-decomposition-ℕ + pr2 div-exp-valuation-2-adic-decomposition-ℕ = + commutative-mul-ℕ + odd-factor-2-adic-decomposition-ℕ + exp-valuation-2-adic-decomposition-ℕ ∙ + eq-2-adic-decomposition-ℕ + + is-power-divisor-exp-valuation-2-adic-decomposition-ℕ : + is-power-divisor-ℕ 2 n + exp-valuation-2-adic-decomposition-ℕ + pr1 (pr1 (is-power-divisor-exp-valuation-2-adic-decomposition-ℕ)) = + valuation-2-adic-decomposition-ℕ + pr2 (pr1 (is-power-divisor-exp-valuation-2-adic-decomposition-ℕ)) = + refl + pr2 (is-power-divisor-exp-valuation-2-adic-decomposition-ℕ) = + div-exp-valuation-2-adic-decomposition-ℕ + + is-largest-power-divisor-2-adic-decomposition-ℕ : + is-largest-power-divisor-ℕ 2 n exp-valuation-2-adic-decomposition-ℕ + pr1 is-largest-power-divisor-2-adic-decomposition-ℕ = + is-power-divisor-exp-valuation-2-adic-decomposition-ℕ + pr2 is-largest-power-divisor-2-adic-decomposition-ℕ y ((k , refl) , K) = + {!!} + ``` ## Properties @@ -150,7 +182,7 @@ module _ div-exp-valuation-2-adic-decomposition-nat-ℕ : div-ℕ (2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) n div-exp-valuation-2-adic-decomposition-nat-ℕ = - div-exp-valuation-largest-power-divisor-ℕ 2 n star H + div-largest-power-divisor-ℕ 2 n star H is-upper-bound-valuation-2-adic-decomposition-nat-ℕ : (k : ℕ) → div-ℕ (2 ^ℕ k) n → k ≤-ℕ valuation-2-adic-decomposition-nat-ℕ @@ -222,14 +254,6 @@ module _ eq-2-adic-decomposition-nat-ℕ ``` -### If $2^km$ is a $2$-adic decomposition of a number $n$, then $2^k$ is the largest power divisor of $n$ base $2$ - -```agda -largest-power-divisor-2-adic-decomposition-ℕ : - (n : ℕ) → 2-adic-decomposition-ℕ n → {!!} -largest-power-divisor-2-adic-decomposition-ℕ = {!!} -``` - ### The type of 2-adic decompositions of any natural number is a proposition ```agda diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 96454b4e97..48cce71535 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -515,6 +515,68 @@ is-nonzero-is-odd-ℕ : is-nonzero-is-odd-ℕ .zero-ℕ H refl = H is-even-zero-ℕ ``` +### A product $mn$ is odd if and only if both factors are odd + +```agda +has-odd-expansion-mul-odd-number-ℕ : + (k l : ℕ) → has-odd-expansion-ℕ (odd-number-ℕ k *ℕ odd-number-ℕ l) +pr1 (has-odd-expansion-mul-odd-number-ℕ k l) = + 2 *ℕ (k *ℕ l) +ℕ k +ℕ l +pr2 (has-odd-expansion-mul-odd-number-ℕ k l) = + ( ap + ( succ-ℕ) + ( ( left-distributive-mul-add-ℕ 2 (2 *ℕ (k *ℕ l) +ℕ k) l) ∙ + ( ap-add-ℕ + ( ( left-distributive-mul-add-ℕ 2 (2 *ℕ (k *ℕ l)) k) ∙ + ( ap-add-ℕ + ( ap (2 *ℕ_) (left-swap-mul-ℕ 2 k l) ∙ + inv (associative-mul-ℕ 2 k (2 *ℕ l))) + ( inv (right-unit-law-mul-ℕ (2 *ℕ k))))) + ( inv (left-unit-law-mul-ℕ (2 *ℕ l)))))) ∙ + ( inv (double-distributive-mul-add-ℕ (2 *ℕ k) 1 (2 *ℕ l) 1)) + +has-odd-expansion-mul-has-odd-expansion-ℕ : + (m n : ℕ) → has-odd-expansion-ℕ m → has-odd-expansion-ℕ n → + has-odd-expansion-ℕ (m *ℕ n) +has-odd-expansion-mul-has-odd-expansion-ℕ m n (k , refl) (l , refl) = + has-odd-expansion-mul-odd-number-ℕ k l + +is-odd-mul-is-odd-ℕ : + (m n : ℕ) → is-odd-ℕ m → is-odd-ℕ n → is-odd-ℕ (m *ℕ n) +is-odd-mul-is-odd-ℕ m n H K = + is-odd-has-odd-expansion-ℕ + ( m *ℕ n) + ( has-odd-expansion-mul-has-odd-expansion-ℕ m n + ( has-odd-expansion-is-odd-ℕ m H) + ( has-odd-expansion-is-odd-ℕ n K)) + +is-odd-left-factor-is-odd-mul-ℕ : + (m n : ℕ) → is-odd-ℕ (m *ℕ n) → is-odd-ℕ m +is-odd-left-factor-is-odd-mul-ℕ m n H K = + H (is-even-div-is-even-ℕ m (m *ℕ n) K (n , commutative-mul-ℕ n m)) + +is-odd-right-factor-is-odd-mul-ℕ : + (m n : ℕ) → is-odd-ℕ (m *ℕ n) → is-odd-ℕ n +is-odd-right-factor-is-odd-mul-ℕ m n H K = + H (is-even-div-is-even-ℕ n (m *ℕ n) K (m , refl)) +``` + +### If a product $mn$ is even and one of the factors is odd, then the other is even + +```agda +is-even-left-factor-is-even-mul-ℕ : + (m n : ℕ) → is-even-ℕ (m *ℕ n) → is-odd-ℕ n → is-even-ℕ m +is-even-left-factor-is-even-mul-ℕ m n H K = + is-even-is-not-odd-ℕ m (λ L → is-odd-mul-is-odd-ℕ m n L K H) + +is-even-right-factor-is-even-mul-ℕ : + (m n : ℕ) → is-even-ℕ (m *ℕ n) → is-odd-ℕ m → is-even-ℕ n +is-even-right-factor-is-even-mul-ℕ m n H K = + is-even-is-not-odd-ℕ n (λ L → is-odd-mul-is-odd-ℕ m n K L H) +``` + +### If a power of $2$ divides $mn$ and one of the factors is odd, then the other is divisible by the power of $2$ + ## See also Further laws of parity are proven in other files, e.g.: From 00448cc95c194933c8fe9b5f30bf3f4add3ff2cd Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 4 Feb 2025 18:21:45 -0500 Subject: [PATCH 62/72] some divisibility infrastructure --- .../divisibility-natural-numbers.lagda.md | 60 +++++++++++++++++++ .../parity-natural-numbers.lagda.md | 34 +++++++++++ 2 files changed, 94 insertions(+) diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 2a3dc3dfff..19c7930e91 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -231,6 +231,33 @@ compute-quotient-div-ℕ q p H K = compute-quotient-bounded-div-ℕ q p ( bounded-div-div-ℕ _ _ H) ( bounded-div-div-ℕ _ _ K) + +compute-quotient-div-is-nonzero-dividend-ℕ : + {m m' n n' : ℕ} (q : m = m') (p : n = n') + (H : div-ℕ m n) (K : div-ℕ m' n') → + is-nonzero-ℕ n' → quotient-div-ℕ m n H = pr1 K +compute-quotient-div-is-nonzero-dividend-ℕ {m} {m'} {n} {zero-ℕ} q p H K f = + ex-falso (f refl) +compute-quotient-div-is-nonzero-dividend-ℕ {m} {m'} {n} {succ-ℕ n'} q p H K f = + compute-quotient-div-ℕ q p H K + +compute-quotient-div-is-nonzero-divisor-ℕ : + {m m' n n' : ℕ} (q : m = m') (p : n = n') + (H : div-ℕ m n) (K : div-ℕ m' n') → + is-nonzero-ℕ m' → quotient-div-ℕ m n H = pr1 K +compute-quotient-div-is-nonzero-divisor-ℕ {m} {zero-ℕ} {n} {n'} q p H K f = + ex-falso (f refl) +compute-quotient-div-is-nonzero-divisor-ℕ + {m} {succ-ℕ m'} {n} {zero-ℕ} q p H (zero-ℕ , K) f = + is-zero-leq-zero-ℕ + ( quotient-div-ℕ m n H) + ( concatenate-leq-eq-ℕ + ( quotient-div-ℕ m n H) + ( upper-bound-quotient-div-ℕ m n H) + ( p)) +compute-quotient-div-is-nonzero-divisor-ℕ + {m} {succ-ℕ m'} {n} {succ-ℕ n'} q p H K f = + compute-quotient-div-is-nonzero-dividend-ℕ q p H K (is-nonzero-succ-ℕ n') ``` ### The quotients of a natural number `n` by two natural numbers `c` and `d` are equal if `c` and `d` are equal @@ -472,6 +499,33 @@ div-mul-ℕ' : (k x y : ℕ) → div-ℕ x y → div-ℕ x (y *ℕ k) div-mul-ℕ' k x y H = tr (div-ℕ x) (commutative-mul-ℕ k y) (div-mul-ℕ k x y H) + +compute-quotient-div-mul-ℕ : + (k x y : ℕ) (H : div-ℕ x y) (K : div-ℕ x (k *ℕ y)) → + quotient-div-ℕ x (k *ℕ y) K = k *ℕ quotient-div-ℕ x y H +compute-quotient-div-mul-ℕ k zero-ℕ y H K = + is-zero-quotient-div-is-zero-divisor-ℕ 0 (k *ℕ y) K refl ∙ + inv (right-zero-law-mul-ℕ k) ∙ + ap (k *ℕ_) (inv (is-zero-quotient-div-is-zero-divisor-ℕ 0 y H refl)) +compute-quotient-div-mul-ℕ k (succ-ℕ x) y H K = + compute-quotient-div-is-nonzero-divisor-ℕ + ( refl) + ( refl) + ( K) + ( div-mul-ℕ k (succ-ℕ x) y H) + ( is-nonzero-succ-ℕ x) + +compute-quotient-div-mul-ℕ' : + (k x y : ℕ) (H : div-ℕ x y) (K : div-ℕ x (y *ℕ k)) → + quotient-div-ℕ x (y *ℕ k) K = quotient-div-ℕ x y H *ℕ k +compute-quotient-div-mul-ℕ' k x y H K = + ( compute-quotient-div-ℕ refl + ( commutative-mul-ℕ y k) + ( K) + ( concatenate-div-eq-ℕ K (commutative-mul-ℕ y k))) ∙ + ( compute-quotient-div-mul-ℕ k x y H + ( concatenate-div-eq-ℕ K (commutative-mul-ℕ y k))) ∙ + ( commutative-mul-ℕ k (quotient-div-ℕ x y H)) ``` ### Divisibility is decidable @@ -674,6 +728,12 @@ div-div-quotient-div-ℕ x y d H K = ( eq-quotient-div-ℕ' d y H) ( preserves-div-left-mul-ℕ d x (quotient-div-ℕ d y H) K) +div-div-quotient-div-ℕ' : + (x y d : ℕ) (H : div-ℕ d y) → + ((K : div-ℕ d y) → div-ℕ x (quotient-div-ℕ d y K)) → div-ℕ (d *ℕ x) y +div-div-quotient-div-ℕ' x y d H K = + div-div-quotient-div-ℕ x y d H (K H) + simplify-div-quotient-div-ℕ : {a d x : ℕ} (H : div-ℕ d a) → div-ℕ x (quotient-div-ℕ d a H) ↔ div-ℕ (d *ℕ x) a diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index 48cce71535..a1eadcd973 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -14,6 +14,7 @@ open import elementary-number-theory.distance-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.euclidean-division-natural-numbers +open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-natural-numbers @@ -577,6 +578,39 @@ is-even-right-factor-is-even-mul-ℕ m n H K = ### If a power of $2$ divides $mn$ and one of the factors is odd, then the other is divisible by the power of $2$ +The proof is by induction, and the base case is trivial. For the inductive step, assume that $2^{k+1} \mid mn$. Then the product $mn$ is even, and therefore $m$ is even. Let $m' := m/2$. Now it follows that $2^k \mid m'n$. By the inductive hypothesis, we conclude that $2^k \mid m'$, which is xequivalent to $2^{k+1} \mid m$. + +```agda +div-exp-2-left-factor-div-exp-2-mul-ℕ : + (k m n : ℕ) → div-ℕ (2 ^ℕ k) (m *ℕ n) → is-odd-ℕ n → div-ℕ (2 ^ℕ k) m +div-exp-2-left-factor-div-exp-2-mul-ℕ zero-ℕ m n H K = + div-one-ℕ m +div-exp-2-left-factor-div-exp-2-mul-ℕ (succ-ℕ k) m n H K = + concatenate-eq-div-ℕ + ( exp-succ-ℕ' 2 k) + ( div-div-quotient-div-ℕ' + ( 2 ^ℕ k) + ( m) + ( 2) + ( is-even-left-factor-is-even-mul-ℕ m n + ( is-even-div-is-even-ℕ + ( 2 ^ℕ succ-ℕ k) + ( m *ℕ n) + ( 2 ^ℕ k , inv (exp-succ-ℕ 2 k)) + ( H)) + ( K)) + ( λ L → + div-exp-2-left-factor-div-exp-2-mul-ℕ k + ( quotient-div-ℕ 2 m L) + ( n) + ( concatenate-div-eq-ℕ + {!!} + ( compute-quotient-div-mul-ℕ' n 2 m + {!!} + {!!})) + ( K))) +``` + ## See also Further laws of parity are proven in other files, e.g.: From 2cf37d651d237bba7c8bb273a94b6166293e1c47 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 5 Feb 2025 10:22:20 -0500 Subject: [PATCH 63/72] lemma about divisors of the form 2^k --- .../parity-natural-numbers.lagda.md | 28 ++++++++++++++++--- 1 file changed, 24 insertions(+), 4 deletions(-) diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index a1eadcd973..a2c35ea22c 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -580,6 +580,8 @@ is-even-right-factor-is-even-mul-ℕ m n H K = The proof is by induction, and the base case is trivial. For the inductive step, assume that $2^{k+1} \mid mn$. Then the product $mn$ is even, and therefore $m$ is even. Let $m' := m/2$. Now it follows that $2^k \mid m'n$. By the inductive hypothesis, we conclude that $2^k \mid m'$, which is xequivalent to $2^{k+1} \mid m$. +**Note.** This result should eventually be formalized as an instance of the following lemma: If $a \mid bc$ and $\gcd(a,c)=1$, then $a\mid b$. + ```agda div-exp-2-left-factor-div-exp-2-mul-ℕ : (k m n : ℕ) → div-ℕ (2 ^ℕ k) (m *ℕ n) → is-odd-ℕ n → div-ℕ (2 ^ℕ k) m @@ -604,11 +606,29 @@ div-exp-2-left-factor-div-exp-2-mul-ℕ (succ-ℕ k) m n H K = ( quotient-div-ℕ 2 m L) ( n) ( concatenate-div-eq-ℕ - {!!} - ( compute-quotient-div-mul-ℕ' n 2 m - {!!} - {!!})) + ( div-quotient-div-div-ℕ + ( 2 ^ℕ k) + ( m *ℕ n) + ( 2) + ( is-even-div-is-even-ℕ + ( 2 ^ℕ succ-ℕ k) + ( m *ℕ n) + ( div-exp-ℕ 2 (succ-ℕ k) (is-nonzero-succ-ℕ k)) + ( H)) + ( concatenate-eq-div-ℕ (inv (exp-succ-ℕ' 2 k)) H)) + ( compute-quotient-div-mul-ℕ' n 2 m L + ( is-even-div-is-even-ℕ + ( 2 ^ℕ succ-ℕ k) + ( m *ℕ n) + ( div-exp-ℕ 2 (succ-ℕ k) (is-nonzero-succ-ℕ k)) + ( H)))) ( K))) + +div-exp-2-right-factor-div-exp-2-mul-ℕ : + (k m n : ℕ) → div-ℕ (2 ^ℕ k) (m *ℕ n) → is-odd-ℕ m → div-ℕ (2 ^ℕ k) n +div-exp-2-right-factor-div-exp-2-mul-ℕ k m n H = + div-exp-2-left-factor-div-exp-2-mul-ℕ k n m + ( concatenate-div-eq-ℕ H (commutative-mul-ℕ m n)) ``` ## See also From 1b7b9b3f3e08e9ef203b147ccd3fc3fcdf387833 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Wed, 5 Feb 2025 14:53:58 -0500 Subject: [PATCH 64/72] Every nonzero natural number has a unique 2-adic decomposition --- .../2-adic-decomposition.lagda.md | 413 +++++++----------- ...st-power-divisors-natural-numbers.lagda.md | 95 +++- .../parity-natural-numbers.lagda.md | 24 +- .../type-arithmetic-natural-numbers.lagda.md | 53 +++ 4 files changed, 319 insertions(+), 266 deletions(-) diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 0dc6a06ffe..04d905144c 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -10,6 +10,7 @@ module elementary-number-theory.2-adic-decomposition where open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.based-strong-induction-natural-numbers open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.exponentiation-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.largest-power-divisors-natural-numbers @@ -19,12 +20,16 @@ open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.strong-induction-natural-numbers open import foundation.action-on-identifications-functions +open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types open import foundation.function-types open import foundation.functoriality-dependent-pair-types open import foundation.propositions +open import foundation.sets open import foundation.split-surjective-maps +open import foundation.subtypes open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels @@ -87,6 +92,11 @@ module _ exp-valuation-2-adic-decomposition-ℕ = 2 ^ℕ valuation-2-adic-decomposition-ℕ + is-nonzero-exp-valuation-2-adic-decomposition-ℕ : + is-nonzero-ℕ exp-valuation-2-adic-decomposition-ℕ + is-nonzero-exp-valuation-2-adic-decomposition-ℕ = + is-nonzero-exp-ℕ 2 valuation-2-adic-decomposition-ℕ (is-nonzero-succ-ℕ 1) + index-odd-factor-2-adic-decomposition-ℕ : ℕ index-odd-factor-2-adic-decomposition-ℕ = pr1 (pr2 d) @@ -94,6 +104,11 @@ module _ odd-factor-2-adic-decomposition-ℕ = odd-number-ℕ index-odd-factor-2-adic-decomposition-ℕ + is-odd-odd-factor-2-adic-decomposition-ℕ : + is-odd-ℕ odd-factor-2-adic-decomposition-ℕ + is-odd-odd-factor-2-adic-decomposition-ℕ = + is-odd-odd-number-ℕ index-odd-factor-2-adic-decomposition-ℕ + eq-2-adic-decomposition-ℕ : 2-adic-composition-ℕ valuation-2-adic-decomposition-ℕ @@ -121,17 +136,88 @@ module _ pr2 (is-power-divisor-exp-valuation-2-adic-decomposition-ℕ) = div-exp-valuation-2-adic-decomposition-ℕ - is-largest-power-divisor-2-adic-decomposition-ℕ : + is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ : is-largest-power-divisor-ℕ 2 n exp-valuation-2-adic-decomposition-ℕ - pr1 is-largest-power-divisor-2-adic-decomposition-ℕ = + pr1 is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ = is-power-divisor-exp-valuation-2-adic-decomposition-ℕ - pr2 is-largest-power-divisor-2-adic-decomposition-ℕ y ((k , refl) , K) = - {!!} - + pr2 is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ y + ((k , refl) , K) = + leq-exponent-div-exp-ℕ + ( 2) + ( valuation-2-adic-decomposition-ℕ) + ( k) + ( star) + ( div-exp-2-left-factor-div-exp-2-mul-ℕ k + ( exp-valuation-2-adic-decomposition-ℕ) + ( odd-factor-2-adic-decomposition-ℕ) + ( concatenate-div-eq-ℕ K (inv eq-2-adic-decomposition-ℕ)) + ( is-odd-odd-factor-2-adic-decomposition-ℕ)) ``` ## Properties +### The type of 2-adic decompositions of any natural number is a proposition + +```agda +eq-valuation-2-adic-decomposition-ℕ : + (n : ℕ) (H K : 2-adic-decomposition-ℕ n) → + valuation-2-adic-decomposition-ℕ n H = valuation-2-adic-decomposition-ℕ n K +eq-valuation-2-adic-decomposition-ℕ n H K = + eq-valuation-is-largest-power-divisor-ℕ 2 n + ( exp-valuation-2-adic-decomposition-ℕ n H) + ( exp-valuation-2-adic-decomposition-ℕ n K) + ( star) + ( is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ n H) + ( is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ n K) + +eq-exp-valuation-2-adic-decomposition-ℕ : + (n : ℕ) (H K : 2-adic-decomposition-ℕ n) → + exp-valuation-2-adic-decomposition-ℕ n H = + exp-valuation-2-adic-decomposition-ℕ n K +eq-exp-valuation-2-adic-decomposition-ℕ n H K = + ap (2 ^ℕ_) (eq-valuation-2-adic-decomposition-ℕ n H K) + +eq-odd-factor-2-adic-decomposition-ℕ : + (n : ℕ) (H K : 2-adic-decomposition-ℕ n) → + odd-factor-2-adic-decomposition-ℕ n H = odd-factor-2-adic-decomposition-ℕ n K +eq-odd-factor-2-adic-decomposition-ℕ n H K = + is-injective-left-mul-ℕ + ( exp-valuation-2-adic-decomposition-ℕ n H) + ( is-nonzero-exp-valuation-2-adic-decomposition-ℕ n H) + ( ( eq-2-adic-decomposition-ℕ n H) ∙ + ( inv (eq-2-adic-decomposition-ℕ n K)) ∙ + ( ap + ( _*ℕ odd-factor-2-adic-decomposition-ℕ n K) + ( inv (eq-exp-valuation-2-adic-decomposition-ℕ n H K)))) + +eq-index-odd-factor-2-adic-decomposition-ℕ : + (n : ℕ) (H K : 2-adic-decomposition-ℕ n) → + index-odd-factor-2-adic-decomposition-ℕ n H = + index-odd-factor-2-adic-decomposition-ℕ n K +eq-index-odd-factor-2-adic-decomposition-ℕ n H K = + is-injective-odd-number-ℕ (eq-odd-factor-2-adic-decomposition-ℕ n H K) + +all-elements-equal-2-adic-decomposition-ℕ : + (n : ℕ) → all-elements-equal (2-adic-decomposition-ℕ n) +all-elements-equal-2-adic-decomposition-ℕ n H@(k , m , p) K@(k' , m' , p') = + eq-pair-Σ + ( eq-valuation-2-adic-decomposition-ℕ n H K) + ( eq-type-subtype + ( λ x → Id-Prop ℕ-Set _ _) + ( eq-index-odd-factor-2-adic-decomposition-ℕ n + ( k' , + tr + ( λ x → Σ ℕ (λ y → 2-adic-composition-ℕ x y = n)) + ( eq-valuation-2-adic-decomposition-ℕ n H K) + ( m , p)) + ( K))) + +is-prop-2-adic-decomposition-ℕ : + (n : ℕ) → is-prop (2-adic-decomposition-ℕ n) +is-prop-2-adic-decomposition-ℕ n = + is-prop-all-elements-equal (all-elements-equal-2-adic-decomposition-ℕ n) +``` + ### The values of the $2$-adic composition function are nonzero ```agda @@ -171,286 +257,109 @@ pr2 (pr2 (2-adic-decomposition-is-odd-ℕ n H)) = ```agda module _ - (n : ℕ) (H : 1 ≤-ℕ n) + (n : ℕ) (H : is-nonzero-ℕ n) where - valuation-2-adic-decomposition-nat-ℕ : + valuation-2-adic-decomposition-is-nonzero-ℕ : ℕ - valuation-2-adic-decomposition-nat-ℕ = + valuation-2-adic-decomposition-is-nonzero-ℕ = valuation-largest-power-divisor-ℕ 2 n star H - div-exp-valuation-2-adic-decomposition-nat-ℕ : - div-ℕ (2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) n - div-exp-valuation-2-adic-decomposition-nat-ℕ = + exp-valuation-2-adic-decomposition-is-nonzero-ℕ : + ℕ + exp-valuation-2-adic-decomposition-is-nonzero-ℕ = + 2 ^ℕ valuation-2-adic-decomposition-is-nonzero-ℕ + + div-exp-valuation-2-adic-decomposition-is-nonzero-ℕ : + div-ℕ exp-valuation-2-adic-decomposition-is-nonzero-ℕ n + div-exp-valuation-2-adic-decomposition-is-nonzero-ℕ = div-largest-power-divisor-ℕ 2 n star H - is-upper-bound-valuation-2-adic-decomposition-nat-ℕ : - (k : ℕ) → div-ℕ (2 ^ℕ k) n → k ≤-ℕ valuation-2-adic-decomposition-nat-ℕ - is-upper-bound-valuation-2-adic-decomposition-nat-ℕ = + is-upper-bound-valuation-2-adic-decomposition-is-nonzero-ℕ : + (k : ℕ) → div-ℕ (2 ^ℕ k) n → + k ≤-ℕ valuation-2-adic-decomposition-is-nonzero-ℕ + is-upper-bound-valuation-2-adic-decomposition-is-nonzero-ℕ = is-upper-bound-valuation-largest-power-divisor-ℕ 2 n star H - odd-factor-2-adic-decomposition-nat-ℕ : + odd-factor-2-adic-decomposition-is-nonzero-ℕ : ℕ - odd-factor-2-adic-decomposition-nat-ℕ = + odd-factor-2-adic-decomposition-is-nonzero-ℕ = quotient-div-ℕ - ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) + ( exp-valuation-2-adic-decomposition-is-nonzero-ℕ) ( n) - ( div-exp-valuation-2-adic-decomposition-nat-ℕ) + ( div-exp-valuation-2-adic-decomposition-is-nonzero-ℕ) - is-odd-odd-factor-2-adic-decomposition-nat-ℕ : - is-odd-ℕ odd-factor-2-adic-decomposition-nat-ℕ - is-odd-odd-factor-2-adic-decomposition-nat-ℕ K = + is-odd-odd-factor-2-adic-decomposition-is-nonzero-ℕ : + is-odd-ℕ odd-factor-2-adic-decomposition-is-nonzero-ℕ + is-odd-odd-factor-2-adic-decomposition-is-nonzero-ℕ K = neg-succ-leq-ℕ - ( valuation-2-adic-decomposition-nat-ℕ) - ( is-upper-bound-valuation-2-adic-decomposition-nat-ℕ - ( succ-ℕ valuation-2-adic-decomposition-nat-ℕ) + ( valuation-2-adic-decomposition-is-nonzero-ℕ) + ( is-upper-bound-valuation-2-adic-decomposition-is-nonzero-ℕ + ( succ-ℕ valuation-2-adic-decomposition-is-nonzero-ℕ) ( tr ( is-divisor-ℕ n) - ( inv (exp-succ-ℕ 2 valuation-2-adic-decomposition-nat-ℕ)) + ( inv (exp-succ-ℕ 2 valuation-2-adic-decomposition-is-nonzero-ℕ)) ( div-div-quotient-div-ℕ 2 n - ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) - ( div-exp-valuation-2-adic-decomposition-nat-ℕ) + ( 2 ^ℕ valuation-2-adic-decomposition-is-nonzero-ℕ) + ( div-exp-valuation-2-adic-decomposition-is-nonzero-ℕ) ( K)))) - has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ : - has-odd-expansion-ℕ odd-factor-2-adic-decomposition-nat-ℕ - has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ = + has-odd-expansion-odd-factor-2-adic-decomposition-is-nonzero-ℕ : + has-odd-expansion-ℕ odd-factor-2-adic-decomposition-is-nonzero-ℕ + has-odd-expansion-odd-factor-2-adic-decomposition-is-nonzero-ℕ = has-odd-expansion-is-odd-ℕ - odd-factor-2-adic-decomposition-nat-ℕ - is-odd-odd-factor-2-adic-decomposition-nat-ℕ + odd-factor-2-adic-decomposition-is-nonzero-ℕ + is-odd-odd-factor-2-adic-decomposition-is-nonzero-ℕ - index-odd-factor-2-adic-decomposition-nat-ℕ : + index-odd-factor-2-adic-decomposition-is-nonzero-ℕ : ℕ - index-odd-factor-2-adic-decomposition-nat-ℕ = - pr1 has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ + index-odd-factor-2-adic-decomposition-is-nonzero-ℕ = + pr1 has-odd-expansion-odd-factor-2-adic-decomposition-is-nonzero-ℕ - eq-index-odd-factor-2-adic-decomposition-nat-ℕ : - odd-number-ℕ index-odd-factor-2-adic-decomposition-nat-ℕ = - odd-factor-2-adic-decomposition-nat-ℕ - eq-index-odd-factor-2-adic-decomposition-nat-ℕ = - pr2 has-odd-expansion-odd-factor-2-adic-decomposition-nat-ℕ + eq-index-odd-factor-2-adic-decomposition-is-nonzero-ℕ : + odd-number-ℕ index-odd-factor-2-adic-decomposition-is-nonzero-ℕ = + odd-factor-2-adic-decomposition-is-nonzero-ℕ + eq-index-odd-factor-2-adic-decomposition-is-nonzero-ℕ = + pr2 has-odd-expansion-odd-factor-2-adic-decomposition-is-nonzero-ℕ - eq-2-adic-decomposition-nat-ℕ : + eq-2-adic-decomposition-is-nonzero-ℕ : 2-adic-composition-ℕ - ( valuation-2-adic-decomposition-nat-ℕ) - ( index-odd-factor-2-adic-decomposition-nat-ℕ) = + ( valuation-2-adic-decomposition-is-nonzero-ℕ) + ( index-odd-factor-2-adic-decomposition-is-nonzero-ℕ) = n - eq-2-adic-decomposition-nat-ℕ = + eq-2-adic-decomposition-is-nonzero-ℕ = ( ap - ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ *ℕ_) - ( eq-index-odd-factor-2-adic-decomposition-nat-ℕ)) ∙ + ( 2 ^ℕ valuation-2-adic-decomposition-is-nonzero-ℕ *ℕ_) + ( eq-index-odd-factor-2-adic-decomposition-is-nonzero-ℕ)) ∙ ( eq-quotient-div-ℕ' - ( 2 ^ℕ valuation-2-adic-decomposition-nat-ℕ) + ( 2 ^ℕ valuation-2-adic-decomposition-is-nonzero-ℕ) ( n) - ( div-exp-valuation-2-adic-decomposition-nat-ℕ)) + ( div-exp-valuation-2-adic-decomposition-is-nonzero-ℕ)) - 2-adic-decomposition-nat-ℕ : + 2-adic-decomposition-is-nonzero-ℕ : 2-adic-decomposition-ℕ n - pr1 2-adic-decomposition-nat-ℕ = - valuation-2-adic-decomposition-nat-ℕ - pr1 (pr2 2-adic-decomposition-nat-ℕ) = - index-odd-factor-2-adic-decomposition-nat-ℕ - pr2 (pr2 2-adic-decomposition-nat-ℕ) = - eq-2-adic-decomposition-nat-ℕ + pr1 2-adic-decomposition-is-nonzero-ℕ = + valuation-2-adic-decomposition-is-nonzero-ℕ + pr1 (pr2 2-adic-decomposition-is-nonzero-ℕ) = + index-odd-factor-2-adic-decomposition-is-nonzero-ℕ + pr2 (pr2 2-adic-decomposition-is-nonzero-ℕ) = + eq-2-adic-decomposition-is-nonzero-ℕ + + is-largest-power-divisor-exp-valuation-2-adic-decomposition-is-nonzero-ℕ : + is-largest-power-divisor-ℕ 2 n + exp-valuation-2-adic-decomposition-is-nonzero-ℕ + is-largest-power-divisor-exp-valuation-2-adic-decomposition-is-nonzero-ℕ = + is-largest-power-divisor-exp-valuation-2-adic-decomposition-ℕ n + 2-adic-decomposition-is-nonzero-ℕ ``` -### The type of 2-adic decompositions of any natural number is a proposition +### The type of 2-adic decompositions of a nonzero natural number is contractible ```agda -all-elements-equal-2-adic-decomposition-ℕ : - (n : ℕ) → all-elements-equal (2-adic-decomposition-ℕ n) -all-elements-equal-2-adic-decomposition-ℕ n (k , m , p) (k' , m' , p') = {!!} -``` - -```text -pair-expansion : ℕ → UU lzero -pair-expansion n = - Σ (ℕ × ℕ) - ( λ p → - ( (exp-ℕ 2 (pr1 p)) *ℕ (succ-ℕ ((pr2 p) *ℕ 2))) = - succ-ℕ n) - -is-nonzero-pair-expansion : - (u v : ℕ) → - is-nonzero-ℕ ((exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) -is-nonzero-pair-expansion u v = - is-nonzero-mul-ℕ _ _ - ( is-nonzero-exp-ℕ 2 u is-nonzero-two-ℕ) - ( is-nonzero-succ-ℕ _) - -abstract - has-pair-expansion-is-even-or-odd : - (n : ℕ) → is-even-ℕ n + is-odd-ℕ n → pair-expansion n - has-pair-expansion-is-even-or-odd n = - strong-ind-ℕ - ( λ m → (is-even-ℕ m + is-odd-ℕ m) → (pair-expansion m)) - ( λ x → (0 , 0) , refl) - ( λ k f → - ( λ where - ( inl x) → - ( let s = has-odd-expansion-is-odd-ℕ k (is-odd-is-even-succ-ℕ k x) - in - ( ( 0 , succ-ℕ (pr1 s)) , - ( ap - ( succ-ℕ ∘ succ-ℕ) - ( left-unit-law-add-ℕ _ ∙ - ap succ-ℕ (commutative-mul-ℕ (pr1 s) 2) ∙ - pr2 s)))) - ( inr x) → - ( let - e : is-even-ℕ k - e = is-even-is-odd-succ-ℕ k x - - t : quotient-div-ℕ 2 k e ≤-ℕ k - t = upper-bound-quotient-div-ℕ 2 k e - - s : (pair-expansion (quotient-div-ℕ 2 k e)) - s = - f (quotient-div-ℕ 2 k e) - ( t) - ( is-decidable-is-even-ℕ (quotient-div-ℕ 2 k e)) - in - pair - ( succ-ℕ (pr1 (pr1 s)) , pr2 (pr1 s)) - ( ( ap - ( _*ℕ (succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) - ( commutative-mul-ℕ (exp-ℕ 2 (pr1 (pr1 s))) 2)) ∙ - ( ( associative-mul-ℕ - ( 2) - ( exp-ℕ 2 (pr1 (pr1 s))) - ( succ-ℕ ((pr2 (pr1 s)) *ℕ 2))) ∙ - ( ( ap (2 *ℕ_) (pr2 s)) ∙ - ( ( ap - ( succ-ℕ) - ( ( left-successor-law-add-ℕ - ( 0 +ℕ quotient-div-ℕ 2 k e) - ( quotient-div-ℕ 2 k e)))) ∙ - ( ( ap - ( succ-ℕ ∘ succ-ℕ) - ( ap - ( _+ℕ quotient-div-ℕ 2 k e) - ( left-unit-law-add-ℕ (quotient-div-ℕ 2 k e)))) ∙ - ( ( ap - ( succ-ℕ ∘ succ-ℕ) - ( inv - ( right-two-law-mul-ℕ - ( quotient-div-ℕ 2 k e)))) ∙ - ( ap - ( succ-ℕ ∘ succ-ℕ) - ( eq-quotient-div-ℕ 2 k e))))))))))) - ( n) - -has-pair-expansion : (n : ℕ) → pair-expansion n -has-pair-expansion n = - has-pair-expansion-is-even-or-odd n (is-decidable-is-even-ℕ n) -``` - -### If `(u , v)` and `(u' , v')` are the pairs corresponding the same number `x`, then `u = u'` and `v = v'` - -```text -is-pair-expansion-unique : - (u u' v v' : ℕ) → - ((exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = - ((exp-ℕ 2 u') *ℕ (succ-ℕ (v' *ℕ 2))) → - (u = u') × (v = v') -is-pair-expansion-unique zero-ℕ zero-ℕ v v' p = - ( pair refl - ( is-injective-right-mul-ℕ 2 is-nonzero-two-ℕ - ( is-injective-left-add-ℕ 0 (is-injective-succ-ℕ p)))) -is-pair-expansion-unique zero-ℕ (succ-ℕ u') v v' p = - ex-falso (s t) - where - s : is-odd-ℕ (succ-ℕ (0 +ℕ (v *ℕ 2))) - s = - is-odd-has-odd-expansion-ℕ _ - ( v , - ap - ( succ-ℕ) - ( commutative-mul-ℕ 2 v ∙ - inv (left-unit-law-add-ℕ _))) - t : is-even-ℕ (succ-ℕ (0 +ℕ (v *ℕ 2))) - t = tr is-even-ℕ (inv p) (div-mul-ℕ' _ 2 _ ((exp-ℕ 2 u') , refl)) -is-pair-expansion-unique (succ-ℕ u) zero-ℕ v v' p = - ex-falso (s t) - where - s : is-odd-ℕ (succ-ℕ (0 +ℕ (v' *ℕ 2))) - s = - is-odd-has-odd-expansion-ℕ _ - ( v' , - ap - ( succ-ℕ) - ( commutative-mul-ℕ 2 v' ∙ inv (left-unit-law-add-ℕ _))) - t : is-even-ℕ (succ-ℕ (0 +ℕ (v' *ℕ 2))) - t = tr is-even-ℕ p (div-mul-ℕ' _ 2 _ ((exp-ℕ 2 u) , refl)) -is-pair-expansion-unique (succ-ℕ u) (succ-ℕ u') v v' p = - pu , pv - where - q : - ((exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = - ((exp-ℕ 2 u') *ℕ (succ-ℕ (v' *ℕ 2))) - q = is-injective-left-mul-ℕ 2 is-nonzero-two-ℕ - ( inv (associative-mul-ℕ 2 (exp-ℕ 2 u) (succ-ℕ (v *ℕ 2))) ∙ - ( ( ap (_*ℕ (succ-ℕ (v *ℕ 2))) - ( commutative-mul-ℕ 2 (exp-ℕ 2 u))) ∙ - ( ( p) ∙ - ( ( ap (_*ℕ (succ-ℕ (v' *ℕ 2))) - ( commutative-mul-ℕ (exp-ℕ 2 u') 2)) ∙ - ( associative-mul-ℕ 2 (exp-ℕ 2 u') (succ-ℕ (v' *ℕ 2))))))) - pu : (succ-ℕ u) = (succ-ℕ u') - pu = ap succ-ℕ (pr1 (is-pair-expansion-unique u u' v v' q)) - pv : v = v' - pv = pr2 (is-pair-expansion-unique u u' v v' q) -``` - -A pairing function is a bijection between `ℕ × ℕ` and `ℕ`. - -## Definition - -```text -pairing-map : ℕ × ℕ → ℕ -pairing-map (u , v) = - pr1 (is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v)) -``` - -### Pairing function is split surjective - -```text -is-split-surjective-pairing-map : is-split-surjective pairing-map -is-split-surjective-pairing-map n = - (u , v) , is-injective-succ-ℕ (q ∙ s) - where - u = pr1 (pr1 (has-pair-expansion n)) - v = pr2 (pr1 (has-pair-expansion n)) - s = pr2 (has-pair-expansion n) - r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) - q : - ( succ-ℕ (pairing-map (u , v))) = - ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) - q = inv (pr2 r) -``` - -### Pairing function is injective - -```text -is-injective-pairing-map : is-injective pairing-map -is-injective-pairing-map {u , v} {u' , v'} p = - ( eq-pair' (is-pair-expansion-unique u u' v v' q)) - where - r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) - s = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u' v') - q : - ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = - ( (exp-ℕ 2 u') *ℕ (succ-ℕ (v' *ℕ 2))) - q = (pr2 r) ∙ (ap succ-ℕ p ∙ inv (pr2 s)) -``` - -### Pairing function is equivalence - -```text -is-equiv-pairing-map : is-equiv pairing-map -is-equiv-pairing-map = - is-equiv-is-split-surjective-is-injective - pairing-map - is-injective-pairing-map - is-split-surjective-pairing-map +is-contr-2-adic-decomposition-ℕ : + (n : ℕ) → is-nonzero-ℕ n → is-contr (2-adic-decomposition-ℕ n) +is-contr-2-adic-decomposition-ℕ n H = + is-proof-irrelevant-is-prop + ( is-prop-2-adic-decomposition-ℕ n) + ( 2-adic-decomposition-is-nonzero-ℕ n H) ``` diff --git a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md index 99262bcf48..52e6ce6bf8 100644 --- a/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/largest-power-divisors-natural-numbers.lagda.md @@ -16,6 +16,7 @@ open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import elementary-number-theory.upper-bounds-natural-numbers +open import foundation.action-on-identifications-functions open import foundation.cartesian-product-types open import foundation.dependent-pair-types open import foundation.identity-types @@ -43,27 +44,65 @@ of the form $m^k$ that is-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero is-power-divisor-ℕ m n x = is-exponential-ℕ m x × div-ℕ x n -valuation-is-power-divisor-ℕ : (m n x : ℕ) → is-power-divisor-ℕ m n x → ℕ -valuation-is-power-divisor-ℕ m n x H = pr1 (pr1 H) +valuation-is-power-divisor-ℕ : + (m n x : ℕ) → is-power-divisor-ℕ m n x → ℕ +valuation-is-power-divisor-ℕ m n x H = + pr1 (pr1 H) + +compute-exp-valuation-is-power-divisor-ℕ : + (m n x : ℕ) (H : is-power-divisor-ℕ m n x) → + m ^ℕ valuation-is-power-divisor-ℕ m n x H = x +compute-exp-valuation-is-power-divisor-ℕ m n x H = + pr2 (pr1 H) ``` ### The predicate of being the largest power divisor of a natural number ```agda -is-largest-power-divisor-ℕ : ℕ → ℕ → ℕ → UU lzero -is-largest-power-divisor-ℕ m n x = - Σ ( is-power-divisor-ℕ m n x) - ( λ H → - (y : ℕ) (K : is-power-divisor-ℕ m n y) → - valuation-is-power-divisor-ℕ m n y K ≤-ℕ - valuation-is-power-divisor-ℕ m n x H) +module _ + (m n x : ℕ) + where + + is-largest-power-divisor-ℕ : UU lzero + is-largest-power-divisor-ℕ = + Σ ( is-power-divisor-ℕ m n x) + ( λ H → + (y : ℕ) (K : is-power-divisor-ℕ m n y) → + valuation-is-power-divisor-ℕ m n y K ≤-ℕ + valuation-is-power-divisor-ℕ m n x H) + + is-power-divisor-is-largest-power-divisor-ℕ : + is-largest-power-divisor-ℕ → is-power-divisor-ℕ m n x + is-power-divisor-is-largest-power-divisor-ℕ = + pr1 + + valuation-is-largest-power-divisor-ℕ : + is-largest-power-divisor-ℕ → ℕ + valuation-is-largest-power-divisor-ℕ H = + valuation-is-power-divisor-ℕ m n x + ( is-power-divisor-is-largest-power-divisor-ℕ H) + + compute-exp-valuation-is-largest-power-divisor-ℕ : + (H : is-largest-power-divisor-ℕ) → + m ^ℕ valuation-is-largest-power-divisor-ℕ H = x + compute-exp-valuation-is-largest-power-divisor-ℕ H = + compute-exp-valuation-is-power-divisor-ℕ m n x + ( is-power-divisor-is-largest-power-divisor-ℕ H) + + is-upper-bound-valuation-is-largest-power-divisor-ℕ : + (H : is-largest-power-divisor-ℕ) → + (y : ℕ) (K : is-power-divisor-ℕ m n y) → + valuation-is-power-divisor-ℕ m n y K ≤-ℕ + valuation-is-largest-power-divisor-ℕ H + is-upper-bound-valuation-is-largest-power-divisor-ℕ = + pr2 ``` -### The largest power divisor of a natural number +### The largest power divisor of a nonzero natural number ```agda module _ - (m n : ℕ) (H : 1 <-ℕ m) (K : 1 ≤-ℕ n) + (m n : ℕ) (H : 1 <-ℕ m) (K : is-nonzero-ℕ n) where largest-power-divisor-ℕ : @@ -77,7 +116,7 @@ module _ ( is-finite-map-exp-ℕ m H) ( n) ( 0) - ( K) + ( leq-one-is-nonzero-ℕ n K) ( div-one-ℕ n) valuation-largest-power-divisor-ℕ : @@ -124,7 +163,7 @@ module _ ( is-structured-value-bound-input-ℕ (is-divisor-ℕ n) (m ^ℕ_) n) ( largest-power-divisor-ℕ) ( k) - ( leq-div-ℕ (m ^ℕ k) n (is-nonzero-leq-one-ℕ n K) L , L) + ( leq-div-ℕ (m ^ℕ k) n K L , L) is-largest-power-divisor-largest-power-divisor-ℕ : is-largest-power-divisor-ℕ m n nat-largest-power-divisor-ℕ @@ -134,6 +173,36 @@ module _ is-upper-bound-valuation-largest-power-divisor-ℕ k K ``` +### Any two largest power divisors are equal + +```agda +module _ + (m n x y : ℕ) (H : 1 <-ℕ m) + where + + eq-valuation-is-largest-power-divisor-ℕ : + (H : is-largest-power-divisor-ℕ m n x) + (K : is-largest-power-divisor-ℕ m n y) → + valuation-is-largest-power-divisor-ℕ m n x H = + valuation-is-largest-power-divisor-ℕ m n y K + eq-valuation-is-largest-power-divisor-ℕ H K = + antisymmetric-leq-ℕ + ( valuation-is-largest-power-divisor-ℕ m n x H) + ( valuation-is-largest-power-divisor-ℕ m n y K) + ( is-upper-bound-valuation-is-largest-power-divisor-ℕ m n y K x + ( is-power-divisor-is-largest-power-divisor-ℕ m n x H)) + ( is-upper-bound-valuation-is-largest-power-divisor-ℕ m n x H y + ( is-power-divisor-is-largest-power-divisor-ℕ m n y K)) + + eq-is-largest-power-divisor-ℕ : + is-largest-power-divisor-ℕ m n x → is-largest-power-divisor-ℕ m n y → + x = y + eq-is-largest-power-divisor-ℕ H K = + inv (compute-exp-valuation-is-largest-power-divisor-ℕ m n x H) ∙ + ap (m ^ℕ_) (eq-valuation-is-largest-power-divisor-ℕ H K) ∙ + compute-exp-valuation-is-largest-power-divisor-ℕ m n y K +``` + ## See also - [2-adic decomposition](elementary-number-theory.2-adic-decomposition.md) diff --git a/src/elementary-number-theory/parity-natural-numbers.lagda.md b/src/elementary-number-theory/parity-natural-numbers.lagda.md index a2c35ea22c..74e9e94987 100644 --- a/src/elementary-number-theory/parity-natural-numbers.lagda.md +++ b/src/elementary-number-theory/parity-natural-numbers.lagda.md @@ -32,6 +32,7 @@ open import foundation.empty-types open import foundation.fibers-of-maps open import foundation.function-types open import foundation.identity-types +open import foundation.injective-maps open import foundation.negation open import foundation.transport-along-identifications open import foundation.unit-type @@ -137,7 +138,7 @@ odd-number-succ-ℕ n = ap succ-ℕ (even-number-succ-ℕ n) ``` -### The `n`th even number is even +### The $n$th even number is even ```agda is-even-even-number-ℕ : @@ -146,6 +147,24 @@ pr1 (is-even-even-number-ℕ n) = n pr2 (is-even-even-number-ℕ n) = commutative-mul-ℕ n 2 ``` +### The even number function is injective + +```agda +is-injective-even-number-ℕ : + is-injective even-number-ℕ +is-injective-even-number-ℕ = + is-injective-left-mul-ℕ 2 (is-nonzero-succ-ℕ 1) +``` + +### The odd number function is injective + +```agda +is-injective-odd-number-ℕ : + is-injective odd-number-ℕ +is-injective-odd-number-ℕ = + is-injective-comp is-injective-even-number-ℕ is-injective-succ-ℕ +``` + ### Even natural numbers are closed under equality ```agda @@ -292,6 +311,9 @@ has-odd-expansion-is-odd-ℕ (succ-ℕ (succ-ℕ n)) p = where s : has-odd-expansion-ℕ n s = has-odd-expansion-is-odd-ℕ n (is-odd-is-odd-succ-succ-ℕ n p) + +is-odd-odd-number-ℕ : (n : ℕ) → is-odd-ℕ (odd-number-ℕ n) +is-odd-odd-number-ℕ n = is-odd-has-odd-expansion-ℕ (odd-number-ℕ n) (n , refl) ``` ### A number is even if and only if it is divisible by an even number diff --git a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md index 244a1fe92d..ab88982f69 100644 --- a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md +++ b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md @@ -167,6 +167,59 @@ equiv-iterated-coproduct-ℕ (succ-ℕ n) = ### The product `ℕ × ℕ` is equivalent to `ℕ` +A pairing function is a bijection between `ℕ × ℕ` and `ℕ`. + +## Definition + +```text +pairing-map : ℕ × ℕ → ℕ +pairing-map (u , v) = + pr1 (is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v)) +``` + +### Pairing function is split surjective + +```text +is-split-surjective-pairing-map : is-split-surjective pairing-map +is-split-surjective-pairing-map n = + (u , v) , is-injective-succ-ℕ (q ∙ s) + where + u = pr1 (pr1 (has-pair-expansion n)) + v = pr2 (pr1 (has-pair-expansion n)) + s = pr2 (has-pair-expansion n) + r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) + q : + ( succ-ℕ (pairing-map (u , v))) = + ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) + q = inv (pr2 r) +``` + +### Pairing function is injective + +```text +is-injective-pairing-map : is-injective pairing-map +is-injective-pairing-map {u , v} {u' , v'} p = + ( eq-pair' (is-pair-expansion-unique u u' v v' q)) + where + r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) + s = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u' v') + q : + ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = + ( (exp-ℕ 2 u') *ℕ (succ-ℕ (v' *ℕ 2))) + q = (pr2 r) ∙ (ap succ-ℕ p ∙ inv (pr2 s)) +``` + +### Pairing function is equivalence + +```text +is-equiv-pairing-map : is-equiv pairing-map +is-equiv-pairing-map = + is-equiv-is-split-surjective-is-injective + pairing-map + is-injective-pairing-map + is-split-surjective-pairing-map +``` + ```agda ℕ×ℕ≃ℕ : (ℕ × ℕ) ≃ ℕ ℕ×ℕ≃ℕ = pair pairing-map is-equiv-pairing-map From cec430bdcb0a523953bd806f6d5fa1e6f4638d6a Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 7 Feb 2025 14:20:42 -0500 Subject: [PATCH 65/72] NZM exercise 53 --- .../2-adic-decomposition.lagda.md | 27 +++ ...himedean-property-natural-numbers.lagda.md | 29 ++- .../bezouts-lemma-natural-numbers.lagda.md | 16 +- .../equality-natural-numbers.lagda.md | 10 + .../factorials.lagda.md | 79 ++++++- .../ideals-natural-numbers.lagda.md | 12 ++ .../inequality-integers.lagda.md | 199 ++++++++---------- .../natural-numbers.lagda.md | 5 + .../type-arithmetic-natural-numbers.lagda.md | 110 ++++------ src/foundation/embeddings.lagda.md | 46 ++++ src/foundation/subtype-duality.lagda.md | 14 ++ src/foundation/type-duality.lagda.md | 1 + src/set-theory/countable-sets.lagda.md | 4 +- 13 files changed, 347 insertions(+), 205 deletions(-) diff --git a/src/elementary-number-theory/2-adic-decomposition.lagda.md b/src/elementary-number-theory/2-adic-decomposition.lagda.md index 9532324119..7c2b8f7788 100644 --- a/src/elementary-number-theory/2-adic-decomposition.lagda.md +++ b/src/elementary-number-theory/2-adic-decomposition.lagda.md @@ -27,6 +27,7 @@ open import foundation.equality-cartesian-product-types open import foundation.equality-dependent-pair-types open import foundation.function-types open import foundation.functoriality-dependent-pair-types +open import foundation.logical-equivalences open import foundation.propositional-maps open import foundation.propositions open import foundation.sets @@ -41,6 +42,7 @@ open import foundation-core.cartesian-product-types open import foundation-core.coproduct-types open import foundation-core.empty-types open import foundation-core.equivalences +open import foundation-core.fibers-of-maps open import foundation-core.identity-types open import foundation-core.injective-maps ``` @@ -88,6 +90,11 @@ module _ (n : ℕ) (d : 2-adic-decomposition-ℕ n) where + fiber-2-adic-composition-2-adic-decomposition-ℕ : + fiber (λ x → 2-adic-composition-ℕ (pr1 x) (pr2 x)) n + fiber-2-adic-composition-2-adic-decomposition-ℕ = + ((pr1 d , pr1 (pr2 d)) , pr2 (pr2 d)) + valuation-2-adic-decomposition-ℕ : ℕ valuation-2-adic-decomposition-ℕ = pr1 d @@ -356,6 +363,21 @@ module _ 2-adic-decomposition-is-nonzero-ℕ ``` +### A logical equivalence between the type of 2-adic decompositions and the fibers of the successor function + +```agda +logical-equiv-fiber-2-adic-composition-fiber-succ-ℕ : + (n : ℕ) → + fiber (λ x → 2-adic-composition-ℕ (pr1 x) (pr2 x)) n ↔ fiber succ-ℕ n +pr1 (logical-equiv-fiber-2-adic-composition-fiber-succ-ℕ n) ((x , y) , p) = + fiber-succ-is-successor-ℕ + ( is-successor-is-nonzero-ℕ + ( is-nonzero-2-adic-decomposition-ℕ n (x , y , p))) +pr2 (logical-equiv-fiber-2-adic-composition-fiber-succ-ℕ n) (m , refl) = + fiber-2-adic-composition-2-adic-decomposition-ℕ n + ( 2-adic-decomposition-is-nonzero-ℕ (succ-ℕ m) (is-nonzero-succ-ℕ m)) +``` + ### The type of 2-adic decompositions of a nonzero natural number is contractible ```agda @@ -385,6 +407,11 @@ is-emb-2-adic-composition-ℕ : is-emb (λ x → 2-adic-composition-ℕ (pr1 x) (pr2 x)) is-emb-2-adic-composition-ℕ = is-emb-is-prop-map is-prop-map-2-adic-composition-ℕ + +2-adic-composition-emb-ℕ : + ℕ × ℕ ↪ ℕ +pr1 2-adic-composition-emb-ℕ (x , y) = 2-adic-composition-ℕ x y +pr2 2-adic-composition-emb-ℕ = is-emb-2-adic-composition-ℕ ``` ### The 2-adic composition function is injective diff --git a/src/elementary-number-theory/archimedean-property-natural-numbers.lagda.md b/src/elementary-number-theory/archimedean-property-natural-numbers.lagda.md index 0f84a8ced9..1251ff8154 100644 --- a/src/elementary-number-theory/archimedean-property-natural-numbers.lagda.md +++ b/src/elementary-number-theory/archimedean-property-natural-numbers.lagda.md @@ -7,6 +7,7 @@ module elementary-number-theory.archimedean-property-natural-numbers where
Imports ```agda +open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.euclidean-division-natural-numbers open import elementary-number-theory.multiplication-natural-numbers @@ -16,34 +17,44 @@ open import elementary-number-theory.strict-inequality-natural-numbers open import foundation.dependent-pair-types open import foundation.existential-quantification +open import foundation.identity-types open import foundation.transport-along-identifications ```
-## Definition +## Idea The {{#concept "Archimedean property" Disambiguation="natural numbers" Agda=archimedean-property-ℕ}} of the [natural numbers](elementary-number-theory.natural-numbers.md) is that -for any nonzero natural number `x` and any natural number `y`, there is an +for any [nonzero natural number](elementary-number-theory.nonzero-natural-numbers.md) `x` and any natural number `y`, there is an `n : ℕ` such that `y < n *ℕ x`. +## Definitions + +### The Archimedian property of the natural numbers + ```agda abstract archimedean-property-ℕ : - (x y : ℕ) → is-nonzero-ℕ x → exists ℕ (λ n → le-ℕ-Prop y (n *ℕ x)) + (x y : ℕ) → is-nonzero-ℕ x → exists-structure ℕ (λ n → y <-ℕ n *ℕ x) archimedean-property-ℕ x y nonzero-x = intro-exists - (succ-ℕ (quotient-euclidean-division-ℕ x y)) - ( tr - ( λ z → z <-ℕ succ-ℕ (quotient-euclidean-division-ℕ x y) *ℕ x) - ( eq-euclidean-division-ℕ x y) - ( preserves-le-left-add-ℕ + ( succ-ℕ (quotient-euclidean-division-ℕ x y)) + ( concatenate-eq-le-eq-ℕ + ( y) + ( quotient-euclidean-division-ℕ x y *ℕ x +ℕ + remainder-euclidean-division-ℕ x y) + ( quotient-euclidean-division-ℕ x y *ℕ x +ℕ x) + ( succ-ℕ (quotient-euclidean-division-ℕ x y) *ℕ x) + ( inv (eq-euclidean-division-ℕ x y)) + ( preserves-strict-order-right-add-ℕ ( quotient-euclidean-division-ℕ x y *ℕ x) ( remainder-euclidean-division-ℕ x y) ( x) - ( strict-upper-bound-remainder-euclidean-division-ℕ x y nonzero-x))) + ( strict-upper-bound-remainder-euclidean-division-ℕ x y nonzero-x)) + ( inv (left-successor-law-mul-ℕ (quotient-euclidean-division-ℕ x y) x))) ``` ## External links diff --git a/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md b/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md index 5dff84516e..102199e7f9 100755 --- a/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md +++ b/src/elementary-number-theory/bezouts-lemma-natural-numbers.lagda.md @@ -22,6 +22,7 @@ open import elementary-number-theory.inequality-integers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.integers open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic open import elementary-number-theory.multiplication-integers open import elementary-number-theory.multiplication-natural-numbers @@ -799,13 +800,14 @@ is-inhabited-pos-distance-between-multiples (succ-ℕ x) y = pair (succ-ℕ x) minimal-pos-distance-between-multiples : (x y : ℕ) → minimal-element-ℕ (pos-distance-between-multiples x y) -minimal-pos-distance-between-multiples x y = well-ordering-principle-ℕ - (pos-distance-between-multiples x y) - (λ z → is-decidable-function-type - (is-decidable-neg (is-decidable-is-zero-ℕ (x +ℕ y))) - (is-decidable-product (is-decidable-neg (is-decidable-is-zero-ℕ z)) - (is-decidable-is-distance-between-multiples-ℕ x y z))) - (is-inhabited-pos-distance-between-multiples x y) +minimal-pos-distance-between-multiples x y = + well-ordering-principle-ℕ + ( pos-distance-between-multiples x y) + ( λ z → is-decidable-function-type + ( is-decidable-neg (is-decidable-is-zero-ℕ (x +ℕ y))) + ( is-decidable-product (is-decidable-neg (is-decidable-is-zero-ℕ z)) + ( is-decidable-is-distance-between-multiples-ℕ x y z))) + ( pr2 (is-inhabited-pos-distance-between-multiples x y)) minimal-positive-distance : (x y : ℕ) → ℕ minimal-positive-distance x y = pr1 (minimal-pos-distance-between-multiples x y) diff --git a/src/elementary-number-theory/equality-natural-numbers.lagda.md b/src/elementary-number-theory/equality-natural-numbers.lagda.md index 06ac27ceeb..cafef33b84 100644 --- a/src/elementary-number-theory/equality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/equality-natural-numbers.lagda.md @@ -19,6 +19,7 @@ open import foundation.equivalences open import foundation.function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.identity-types +open import foundation.injective-maps open import foundation.propositions open import foundation.set-truncations open import foundation.sets @@ -27,6 +28,7 @@ open import foundation.universe-levels open import foundation-core.decidable-propositions open import foundation-core.discrete-types +open import foundation-core.embeddings open import foundation-core.torsorial-type-families ``` @@ -179,3 +181,11 @@ is-equiv-Eq-eq-ℕ {m} {n} = equiv-unit-trunc-ℕ-Set : ℕ ≃ type-trunc-Set ℕ equiv-unit-trunc-ℕ-Set = equiv-unit-trunc-Set ℕ-Set ``` + +### The successor embedding + +```agda +succ-emb-ℕ : ℕ ↪ ℕ +pr1 succ-emb-ℕ = succ-ℕ +pr2 succ-emb-ℕ = is-emb-is-injective is-set-ℕ is-injective-succ-ℕ +``` diff --git a/src/elementary-number-theory/factorials.lagda.md b/src/elementary-number-theory/factorials.lagda.md index a6b3dd39b7..19453e3b7c 100644 --- a/src/elementary-number-theory/factorials.lagda.md +++ b/src/elementary-number-theory/factorials.lagda.md @@ -7,8 +7,10 @@ module elementary-number-theory.factorials where
Imports ```agda +open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers +open import elementary-number-theory.greatest-common-divisor-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -83,21 +85,84 @@ leq-factorial-ℕ (succ-ℕ n) = ( is-nonzero-factorial-ℕ n) ``` -### Any two consecutive successor factorials are relatively prime +### For any nonzero natural number $n$, the successor factorial of $n$ and $n$ itself are relatively prime -This property is stated as exercise 53 in section 1.2 of {{#cite NZM}}. +Note that the assumption that $n$ is nonzero is necessary, because the numbers $0$ and $0! + 1$ have greatest common divisor $2$. ```agda +is-one-is-common-divisor-id-succ-factorial-ℕ : + (n d : ℕ) → is-nonzero-ℕ n → + is-common-divisor-ℕ n (succ-factorial-ℕ n) d → is-one-ℕ d +is-one-is-common-divisor-id-succ-factorial-ℕ n d H (K , L) = + is-one-div-one-ℕ d + ( div-right-summand-ℕ d + ( factorial-ℕ n) + ( 1) + ( transitive-div-ℕ d n + ( factorial-ℕ n) + ( div-factorial-ℕ n n (refl-leq-ℕ n) H) + ( K)) + ( L)) + +relatively-prime-id-succ-factorial-ℕ : + (n : ℕ) → is-nonzero-ℕ n → is-relatively-prime-ℕ n (succ-factorial-ℕ n) +relatively-prime-id-succ-factorial-ℕ n H = + is-one-is-common-divisor-id-succ-factorial-ℕ + ( n) + ( gcd-ℕ n (succ-factorial-ℕ n)) + ( H) + ( is-common-divisor-gcd-ℕ n (succ-factorial-ℕ n)) +``` + +### The successor factorials of $n$ and $n+1$ are relatively prime for any nonzero $n$ + +This property is stated as exercise 53 in section 1.2 of {{#cite NZM}}. This exercise was stated in a setting where the natural numbers start at $1$. The assumption that $n$ is nonzero is necessary, because $0! + 1 = 2$ and $1! + 1 = 2$. + +**Proof.** First, we observe that + +```text + ((n + 1)! + 1) + n = n!(n + 1) + (n + 1) + = (n! + 1)(n + 1). +``` + +This shows that if `d` is a common divisor of `n! + 1` and `(n + 1)! + 1`, then `d` is a divisor of `n`. However, `n` and `n! + 1` are relatively prime if `n` is nonzero, so it follows that `d = 1`. + +```agda +recursive-law-succ-factorial-ℕ : + (n : ℕ) → succ-factorial-ℕ (succ-ℕ n) +ℕ n = succ-factorial-ℕ n *ℕ succ-ℕ n +recursive-law-succ-factorial-ℕ n = + left-successor-law-add-ℕ (factorial-ℕ (succ-ℕ n)) n ∙ + inv (left-successor-law-mul-ℕ (factorial-ℕ n) (succ-ℕ n)) + is-one-is-common-divisor-consecutive-succ-factorial-ℕ : - (n d : ℕ) → - div-ℕ d (succ-factorial-ℕ n) → div-ℕ d (succ-factorial-ℕ (succ-ℕ n)) → + (n d : ℕ) → is-nonzero-ℕ n → + is-common-divisor-ℕ (succ-factorial-ℕ n) (succ-factorial-ℕ (succ-ℕ n)) d → is-one-ℕ d -is-one-is-common-divisor-consecutive-succ-factorial-ℕ n d H K = {!!} +is-one-is-common-divisor-consecutive-succ-factorial-ℕ n d H (K , L) = + is-one-is-common-divisor-id-succ-factorial-ℕ n d H + ( ( div-right-summand-ℕ + ( d) + ( succ-factorial-ℕ (succ-ℕ n)) + ( n) + ( L) + ( concatenate-div-eq-ℕ + ( div-mul-ℕ' (succ-ℕ n) d (succ-factorial-ℕ n) K) + ( inv (recursive-law-succ-factorial-ℕ n)))) , + ( K)) is-relatively-prime-consecutive-succ-factorial-ℕ : - (n : ℕ) → + (n : ℕ) → is-nonzero-ℕ n → is-relatively-prime-ℕ (succ-factorial-ℕ n) (succ-factorial-ℕ (succ-ℕ n)) -is-relatively-prime-consecutive-succ-factorial-ℕ n = {!!} +is-relatively-prime-consecutive-succ-factorial-ℕ n H = + is-one-is-common-divisor-consecutive-succ-factorial-ℕ + ( n) + ( gcd-ℕ + ( succ-factorial-ℕ n) + ( succ-factorial-ℕ (succ-ℕ n))) + ( H) + ( is-common-divisor-gcd-ℕ + ( succ-factorial-ℕ n) + ( succ-factorial-ℕ (succ-ℕ n))) ``` ## External links diff --git a/src/elementary-number-theory/ideals-natural-numbers.lagda.md b/src/elementary-number-theory/ideals-natural-numbers.lagda.md index 53272a3c39..9c0ab95264 100644 --- a/src/elementary-number-theory/ideals-natural-numbers.lagda.md +++ b/src/elementary-number-theory/ideals-natural-numbers.lagda.md @@ -3,3 +3,15 @@ ```agda module elementary-number-theory.ideals-natural-numbers where ``` + +
Imports + +```agda + +``` + +
+ +## Idea + +Since the [natural numbers](elementary-number-theory.natural-numbers.md) only form a [commutative semiring](elementary-number-theory.commutative-semiring-of-natural-numbers.md), their ideals aren't as easily defined as [ideals](ring-theory.ideals-rings.md) on a [ring](ring-theory.rings.md) diff --git a/src/elementary-number-theory/inequality-integers.lagda.md b/src/elementary-number-theory/inequality-integers.lagda.md index 015a013878..423e6669f5 100644 --- a/src/elementary-number-theory/inequality-integers.lagda.md +++ b/src/elementary-number-theory/inequality-integers.lagda.md @@ -22,6 +22,7 @@ open import elementary-number-theory.positive-integers open import foundation.action-on-identifications-functions open import foundation.coproduct-types open import foundation.decidable-propositions +open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.function-types open import foundation.functoriality-coproduct-types @@ -72,13 +73,13 @@ We will introduce both orderings and prove that they are equivalent. ```agda leq-ℤ-Prop : ℤ → ℤ → Prop lzero -leq-ℤ-Prop x y = subtype-nonnegative-ℤ (y -ℤ x) +leq-ℤ-Prop a b = subtype-nonnegative-ℤ (b -ℤ a) leq-ℤ : ℤ → ℤ → UU lzero -leq-ℤ x y = type-Prop (leq-ℤ-Prop x y) +leq-ℤ a b = type-Prop (leq-ℤ-Prop a b) -is-prop-leq-ℤ : (x y : ℤ) → is-prop (leq-ℤ x y) -is-prop-leq-ℤ x y = is-prop-type-Prop (leq-ℤ-Prop x y) +is-prop-leq-ℤ : (a b : ℤ) → is-prop (leq-ℤ a b) +is-prop-leq-ℤ a b = is-prop-type-Prop (leq-ℤ-Prop a b) infix 30 _≤-ℤ_ _≤-ℤ_ = leq-ℤ @@ -112,45 +113,44 @@ succ-inductive-leq-ℤ' a b .(succ-ℤ b) refl H = succ-inductive-leq-ℤ a b H ### Inequality on the integers is reflexive, antisymmetric and transitive ```agda -refl-leq-ℤ : (k : ℤ) → leq-ℤ k k -refl-leq-ℤ k = tr is-nonnegative-ℤ (inv (right-inverse-law-add-ℤ k)) star +refl-leq-ℤ : (a : ℤ) → a ≤-ℤ a +refl-leq-ℤ a = tr is-nonnegative-ℤ (inv (right-inverse-law-add-ℤ a)) star -antisymmetric-leq-ℤ : {x y : ℤ} → leq-ℤ x y → leq-ℤ y x → x = y -antisymmetric-leq-ℤ {x} {y} H K = +antisymmetric-leq-ℤ : {a b : ℤ} → a ≤-ℤ b → b ≤-ℤ a → a = b +antisymmetric-leq-ℤ {a} {b} H K = eq-diff-ℤ ( is-zero-is-nonnegative-neg-is-nonnegative-ℤ K - ( is-nonnegative-eq-ℤ (inv (distributive-neg-diff-ℤ x y)) H)) + ( is-nonnegative-eq-ℤ (inv (distributive-neg-diff-ℤ a b)) H)) -transitive-leq-ℤ : (k l m : ℤ) → leq-ℤ l m → leq-ℤ k l → leq-ℤ k m -transitive-leq-ℤ k l m H K = +transitive-leq-ℤ : (a b c : ℤ) → b ≤-ℤ c → a ≤-ℤ b → a ≤-ℤ c +transitive-leq-ℤ a b c H K = is-nonnegative-eq-ℤ - ( triangle-diff-ℤ m l k) + ( triangle-diff-ℤ c b a) ( is-nonnegative-add-ℤ H K) ``` ### Inequality on the integers is decidable ```agda -is-decidable-leq-ℤ : (x y : ℤ) → (leq-ℤ x y) + ¬ (leq-ℤ x y) -is-decidable-leq-ℤ x y = is-decidable-is-nonnegative-ℤ (y -ℤ x) - -leq-ℤ-Decidable-Prop : (x y : ℤ) → Decidable-Prop lzero -leq-ℤ-Decidable-Prop x y = - ( leq-ℤ x y , - is-prop-leq-ℤ x y , - is-decidable-leq-ℤ x y) +is-decidable-leq-ℤ : (a b : ℤ) → is-decidable (a ≤-ℤ b) +is-decidable-leq-ℤ a b = is-decidable-is-nonnegative-ℤ (b -ℤ a) + +leq-ℤ-Decidable-Prop : (a b : ℤ) → Decidable-Prop lzero +pr1 (leq-ℤ-Decidable-Prop a b) = a ≤-ℤ b +pr1 (pr2 (leq-ℤ-Decidable-Prop a b)) = is-prop-leq-ℤ a b +pr2 (pr2 (leq-ℤ-Decidable-Prop a b)) = is-decidable-leq-ℤ a b ``` ### Inequality on the integers is linear ```agda -linear-leq-ℤ : (x y : ℤ) → (leq-ℤ x y) + (leq-ℤ y x) -linear-leq-ℤ x y = +linear-leq-ℤ : (a b : ℤ) → (a ≤-ℤ b) + (b ≤-ℤ a) +linear-leq-ℤ a b = map-coproduct ( λ H → is-nonnegative-is-positive-ℤ ( is-positive-eq-ℤ - ( distributive-neg-diff-ℤ x y) + ( distributive-neg-diff-ℤ a b) ( is-positive-neg-is-negative-ℤ H))) ( id) ( decide-is-negative-is-nonnegative-ℤ) @@ -171,7 +171,7 @@ pr2 ℤ-Decidable-Preorder = is-decidable-leq-ℤ ℤ-Poset : Poset lzero lzero pr1 ℤ-Poset = ℤ-Preorder -pr2 ℤ-Poset x y = antisymmetric-leq-ℤ +pr2 ℤ-Poset a b = antisymmetric-leq-ℤ ℤ-Decidable-Poset : Decidable-Poset lzero lzero pr1 ℤ-Decidable-Poset = ℤ-Poset @@ -179,20 +179,20 @@ pr2 ℤ-Decidable-Poset = is-decidable-leq-ℤ ℤ-Total-Preorder : Total-Preorder lzero lzero pr1 ℤ-Total-Preorder = ℤ-Preorder -pr2 ℤ-Total-Preorder x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr2 ℤ-Total-Preorder a b = unit-trunc-Prop (linear-leq-ℤ a b) ℤ-Decidable-Total-Preorder : Decidable-Total-Preorder lzero lzero pr1 ℤ-Decidable-Total-Preorder = ℤ-Preorder -pr1 (pr2 ℤ-Decidable-Total-Preorder) x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr1 (pr2 ℤ-Decidable-Total-Preorder) a b = unit-trunc-Prop (linear-leq-ℤ a b) pr2 (pr2 ℤ-Decidable-Total-Preorder) = is-decidable-leq-ℤ ℤ-Total-Order : Total-Order lzero lzero pr1 ℤ-Total-Order = ℤ-Poset -pr2 ℤ-Total-Order x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr2 ℤ-Total-Order a b = unit-trunc-Prop (linear-leq-ℤ a b) ℤ-Decidable-Total-Order : Decidable-Total-Order lzero lzero pr1 ℤ-Decidable-Total-Order = ℤ-Poset -pr1 (pr2 ℤ-Decidable-Total-Order) x y = unit-trunc-Prop (linear-leq-ℤ x y) +pr1 (pr2 ℤ-Decidable-Total-Order) a b = unit-trunc-Prop (linear-leq-ℤ a b) pr2 (pr2 ℤ-Decidable-Total-Order) = is-decidable-leq-ℤ ``` @@ -200,56 +200,31 @@ pr2 (pr2 ℤ-Decidable-Total-Order) = is-decidable-leq-ℤ ```agda concatenate-eq-leq-eq-ℤ : - {x' x y y' : ℤ} → x' = x → leq-ℤ x y → y = y' → leq-ℤ x' y' + {a b c d : ℤ} → a = b → b ≤-ℤ c → c = d → a ≤-ℤ d concatenate-eq-leq-eq-ℤ refl H refl = H concatenate-leq-eq-ℤ : - (x : ℤ) {y y' : ℤ} → leq-ℤ x y → y = y' → leq-ℤ x y' -concatenate-leq-eq-ℤ x H refl = H + (a : ℤ) {b c : ℤ} → a ≤-ℤ b → b = c → a ≤-ℤ c +concatenate-leq-eq-ℤ a H refl = H concatenate-eq-leq-ℤ : - {x x' : ℤ} (y : ℤ) → x' = x → leq-ℤ x y → leq-ℤ x' y -concatenate-eq-leq-ℤ y refl H = H + {a b : ℤ} (c : ℤ) → a = b → b ≤-ℤ c → a ≤-ℤ c +concatenate-eq-leq-ℤ c refl H = H ``` ### An integer is lesser than its successor ```agda -succ-leq-ℤ : (k : ℤ) → leq-ℤ k (succ-ℤ k) -succ-leq-ℤ k = +succ-leq-ℤ : (a : ℤ) → a ≤-ℤ succ-ℤ a +succ-leq-ℤ a = is-nonnegative-eq-ℤ ( inv - ( ( left-successor-law-add-ℤ k (neg-ℤ k)) ∙ - ( ap succ-ℤ (right-inverse-law-add-ℤ k)))) + ( ( left-successor-law-add-ℤ a (neg-ℤ a)) ∙ + ( ap succ-ℤ (right-inverse-law-add-ℤ a)))) ( star) -leq-succ-leq-ℤ : (k l : ℤ) → leq-ℤ k l → leq-ℤ k (succ-ℤ l) -leq-succ-leq-ℤ k l = transitive-leq-ℤ k l (succ-ℤ l) (succ-leq-ℤ l) -``` - -### The inductively defined inequality on the integers is valued in propositions - -```agda -contradiction-inductive-leq-ℤ : - (x y : ℤ) (p : x = succ-ℤ y) → ¬ (inductive-leq-ℤ x y) -contradiction-inductive-leq-ℤ x .x p (refl-inductive-leq-ℤ .x) = - has-no-fixed-points-succ-ℤ x (inv p) -contradiction-inductive-leq-ℤ x .(succ-ℤ b) p (succ-inductive-leq-ℤ .x b H) = {!!} - -all-elements-equal-inductive-leq-ℤ : - (x y z : ℤ) (p : y = z) (H : inductive-leq-ℤ x y) (K : inductive-leq-ℤ x z) → - tr (inductive-leq-ℤ x) p H = K -all-elements-equal-inductive-leq-ℤ x .x .x p - ( refl-inductive-leq-ℤ .x) - ( refl-inductive-leq-ℤ .x) = - ap - ( λ t → tr (inductive-leq-ℤ x) t (refl-inductive-leq-ℤ x)) - ( eq-is-prop (is-set-ℤ x x) {p} {refl}) -all-elements-equal-inductive-leq-ℤ x .x ._ p - ( refl-inductive-leq-ℤ .x) - ( succ-inductive-leq-ℤ .x b K) = {!!} -all-elements-equal-inductive-leq-ℤ x .(succ-ℤ b) .x p (succ-inductive-leq-ℤ .x b H) (refl-inductive-leq-ℤ .x) = {!!} -all-elements-equal-inductive-leq-ℤ x .(succ-ℤ b) .(succ-ℤ b₁) p (succ-inductive-leq-ℤ .x b H) (succ-inductive-leq-ℤ .x b₁ K) = {!!} +leq-succ-leq-ℤ : (a b : ℤ) → a ≤-ℤ b → a ≤-ℤ succ-ℤ b +leq-succ-leq-ℤ a b = transitive-leq-ℤ a b (succ-ℤ b) (succ-leq-ℤ b) ``` ### Inequality on the integers is equivalent to the inductively defined inequality on the integers @@ -297,17 +272,17 @@ inductive-leq-leq-ℤ a b = ```agda preserves-order-left-add-ℤ : - (z x y : ℤ) → leq-ℤ x y → leq-ℤ (x +ℤ z) (y +ℤ z) -preserves-order-left-add-ℤ z x y = - is-nonnegative-eq-ℤ (inv (right-translation-diff-ℤ y x z)) + (c a b : ℤ) → a ≤-ℤ b → a +ℤ c ≤-ℤ b +ℤ c +preserves-order-left-add-ℤ c a b = + is-nonnegative-eq-ℤ (inv (right-translation-diff-ℤ b a c)) preserves-order-right-add-ℤ : - (z x y : ℤ) → leq-ℤ x y → leq-ℤ (z +ℤ x) (z +ℤ y) -preserves-order-right-add-ℤ z x y = - is-nonnegative-eq-ℤ (inv (left-translation-diff-ℤ y x z)) + (c a b : ℤ) → a ≤-ℤ b → c +ℤ a ≤-ℤ c +ℤ b +preserves-order-right-add-ℤ c a b = + is-nonnegative-eq-ℤ (inv (left-translation-diff-ℤ b a c)) preserves-order-add-ℤ : - {a b c d : ℤ} → leq-ℤ a b → leq-ℤ c d → leq-ℤ (a +ℤ c) (b +ℤ d) + {a b c d : ℤ} → a ≤-ℤ b → c ≤-ℤ d → a +ℤ c ≤-ℤ b +ℤ d preserves-order-add-ℤ {a} {b} {c} {d} H K = transitive-leq-ℤ ( a +ℤ c) @@ -321,30 +296,30 @@ preserves-order-add-ℤ {a} {b} {c} {d} H K = ```agda reflects-order-left-add-ℤ : - (z x y : ℤ) → leq-ℤ (x +ℤ z) (y +ℤ z) → leq-ℤ x y -reflects-order-left-add-ℤ z x y = - is-nonnegative-eq-ℤ (right-translation-diff-ℤ y x z) + (c a b : ℤ) → a +ℤ c ≤-ℤ b +ℤ c → a ≤-ℤ b +reflects-order-left-add-ℤ c a b = + is-nonnegative-eq-ℤ (right-translation-diff-ℤ b a c) reflects-order-right-add-ℤ : - (z x y : ℤ) → leq-ℤ (z +ℤ x) (z +ℤ y) → leq-ℤ x y -reflects-order-right-add-ℤ z x y = - is-nonnegative-eq-ℤ (left-translation-diff-ℤ y x z) + (c a b : ℤ) → c +ℤ a ≤-ℤ c +ℤ b → a ≤-ℤ b +reflects-order-right-add-ℤ c a b = + is-nonnegative-eq-ℤ (left-translation-diff-ℤ b a c) ``` ### The inclusion of ℕ into ℤ preserves inequality ```agda -leq-int-ℕ : (x y : ℕ) → leq-ℕ x y → leq-ℤ (int-ℕ x) (int-ℕ y) -leq-int-ℕ zero-ℕ y H = +leq-int-ℕ : (m n : ℕ) → m ≤-ℕ n → int-ℕ m ≤-ℤ int-ℕ n +leq-int-ℕ zero-ℕ n H = tr ( is-nonnegative-ℤ) - ( inv (right-unit-law-add-ℤ (int-ℕ y))) - ( is-nonnegative-int-ℕ y) -leq-int-ℕ (succ-ℕ x) (succ-ℕ y) H = tr (is-nonnegative-ℤ) - ( inv (diff-succ-ℤ (int-ℕ y) (int-ℕ x)) ∙ - ( ap (_-ℤ (succ-ℤ (int-ℕ x))) (succ-int-ℕ y) ∙ - ap ((int-ℕ (succ-ℕ y)) -ℤ_) (succ-int-ℕ x))) - ( leq-int-ℕ x y H) + ( inv (right-unit-law-add-ℤ (int-ℕ n))) + ( is-nonnegative-int-ℕ n) +leq-int-ℕ (succ-ℕ m) (succ-ℕ n) H = tr (is-nonnegative-ℤ) + ( inv (diff-succ-ℤ (int-ℕ n) (int-ℕ m)) ∙ + ( ap (_-ℤ (succ-ℤ (int-ℕ m))) (succ-int-ℕ n) ∙ + ap ((int-ℕ (succ-ℕ n)) -ℤ_) (succ-int-ℕ m))) + ( leq-int-ℕ m n H) ``` ### Transposing summands in inequalities @@ -456,57 +431,57 @@ transpose-left-neg-leq-ℤ a b H = ( concatenate-leq-eq-ℤ (neg-ℤ b) H (inv (neg-neg-ℤ a))) ``` -### An integer `x` is nonnegative if and only if `0 ≤ x` +### An integer `a` is nonnegative if and only if `0 ≤ a` ```agda module _ - (x : ℤ) + (a : ℤ) where abstract - leq-zero-is-nonnegative-ℤ : is-nonnegative-ℤ x → leq-ℤ zero-ℤ x + leq-zero-is-nonnegative-ℤ : is-nonnegative-ℤ a → zero-ℤ ≤-ℤ a leq-zero-is-nonnegative-ℤ = - is-nonnegative-eq-ℤ (inv (right-zero-law-diff-ℤ x)) + is-nonnegative-eq-ℤ (inv (right-zero-law-diff-ℤ a)) - is-nonnegative-leq-zero-ℤ : leq-ℤ zero-ℤ x → is-nonnegative-ℤ x + is-nonnegative-leq-zero-ℤ : zero-ℤ ≤-ℤ a → is-nonnegative-ℤ a is-nonnegative-leq-zero-ℤ = - is-nonnegative-eq-ℤ (right-zero-law-diff-ℤ x) + is-nonnegative-eq-ℤ (right-zero-law-diff-ℤ a) ``` ### An integer greater than or equal to a nonnegative integer is nonnegative ```agda module _ - (x y : ℤ) (I : leq-ℤ x y) + (a b : ℤ) (I : a ≤-ℤ b) where abstract - is-nonnegative-leq-nonnegative-ℤ : is-nonnegative-ℤ x → is-nonnegative-ℤ y + is-nonnegative-leq-nonnegative-ℤ : is-nonnegative-ℤ a → is-nonnegative-ℤ b is-nonnegative-leq-nonnegative-ℤ H = - is-nonnegative-leq-zero-ℤ y + is-nonnegative-leq-zero-ℤ b ( transitive-leq-ℤ ( zero-ℤ) - ( x) - ( y) + ( a) + ( b) ( I) - ( leq-zero-is-nonnegative-ℤ x H)) + ( leq-zero-is-nonnegative-ℤ a H)) ``` -### An integer `x` is nonpositive if and only if `x ≤ 0` +### An integer `a` is nonpositive if and only if `a ≤ 0` ```agda module _ - (x : ℤ) + (a : ℤ) where abstract - leq-zero-is-nonpositive-ℤ : is-nonpositive-ℤ x → leq-ℤ x zero-ℤ + leq-zero-is-nonpositive-ℤ : is-nonpositive-ℤ a → a ≤-ℤ zero-ℤ leq-zero-is-nonpositive-ℤ = is-nonnegative-neg-is-nonpositive-ℤ - is-nonpositive-leq-zero-ℤ : leq-ℤ x zero-ℤ → is-nonpositive-ℤ x + is-nonpositive-leq-zero-ℤ : a ≤-ℤ zero-ℤ → is-nonpositive-ℤ a is-nonpositive-leq-zero-ℤ H = is-nonpositive-eq-ℤ - ( neg-neg-ℤ x) + ( neg-neg-ℤ a) ( is-nonpositive-neg-is-nonnegative-ℤ H) ``` @@ -514,18 +489,18 @@ module _ ```agda module _ - (x y : ℤ) (I : leq-ℤ x y) + (a b : ℤ) (I : a ≤-ℤ b) where abstract - is-nonpositive-leq-nonpositive-ℤ : is-nonpositive-ℤ y → is-nonpositive-ℤ x + is-nonpositive-leq-nonpositive-ℤ : is-nonpositive-ℤ b → is-nonpositive-ℤ a is-nonpositive-leq-nonpositive-ℤ H = - is-nonpositive-leq-zero-ℤ x + is-nonpositive-leq-zero-ℤ a ( transitive-leq-ℤ - ( x) - ( y) + ( a) + ( b) ( zero-ℤ) - ( leq-zero-is-nonpositive-ℤ y H) + ( leq-zero-is-nonpositive-ℤ b H) ( I)) ``` @@ -533,9 +508,9 @@ module _ ```agda leq-zero-is-positive-ℤ : - (x : ℤ) → is-positive-ℤ x → zero-ℤ ≤-ℤ x -leq-zero-is-positive-ℤ x H = - leq-zero-is-nonnegative-ℤ x (is-nonnegative-is-positive-ℤ H) + (a : ℤ) → is-positive-ℤ a → zero-ℤ ≤-ℤ a +leq-zero-is-positive-ℤ a H = + leq-zero-is-nonnegative-ℤ a (is-nonnegative-is-positive-ℤ H) ``` ## See also diff --git a/src/elementary-number-theory/natural-numbers.lagda.md b/src/elementary-number-theory/natural-numbers.lagda.md index af0b132b43..e6b5dfe66c 100644 --- a/src/elementary-number-theory/natural-numbers.lagda.md +++ b/src/elementary-number-theory/natural-numbers.lagda.md @@ -11,6 +11,7 @@ open import foundation.dependent-pair-types open import foundation.universe-levels open import foundation-core.empty-types +open import foundation-core.fibers-of-maps open import foundation-core.function-types open import foundation-core.identity-types open import foundation-core.injective-maps @@ -74,6 +75,10 @@ is-zero-ℕ' n = (zero-ℕ = n) is-successor-ℕ : ℕ → UU lzero is-successor-ℕ n = Σ ℕ (λ y → n = succ-ℕ y) +fiber-succ-is-successor-ℕ : + {n : ℕ} → is-successor-ℕ n → fiber succ-ℕ n +fiber-succ-is-successor-ℕ (x , p) = (x , inv p) + pred-is-successor-ℕ : (n : ℕ) → is-successor-ℕ n → ℕ pred-is-successor-ℕ n H = pr1 H diff --git a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md index 784d97a48d..7f367cf3d4 100644 --- a/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md +++ b/src/elementary-number-theory/type-arithmetic-natural-numbers.lagda.md @@ -9,6 +9,7 @@ module elementary-number-theory.type-arithmetic-natural-numbers where ```agda open import elementary-number-theory.addition-natural-numbers open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.integers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers @@ -16,8 +17,10 @@ open import elementary-number-theory.parity-natural-numbers open import elementary-number-theory.2-adic-decomposition open import foundation.action-on-identifications-functions +open import foundation.commuting-triangles-of-maps open import foundation.contractible-maps open import foundation.dependent-pair-types +open import foundation.embeddings open import foundation.function-types open import foundation.functoriality-cartesian-product-types open import foundation.functoriality-coproduct-types @@ -168,7 +171,10 @@ equiv-iterated-coproduct-ℕ (succ-ℕ n) = ### The product `ℕ × ℕ` is equivalent to `ℕ` -The [2-adic composition map](elementary-number-theory.2-adic-decomposition.md) $\mathbb{N}\times\mathbb{N} \to \mathbb{N}$ is an [embedding](foundation-core.embeddings.md) which fits in a [commuting triangle](foundation-core.commuting-triangles-of-maps.md) +The [2-adic composition map](elementary-number-theory.2-adic-decomposition.md) +$\mathbb{N}\times\mathbb{N} \to \mathbb{N}$ is an +[embedding](foundation-core.embeddings.md) which fits in a +[commuting triangle](foundation-core.commuting-triangles-of-maps.md) ```text ℕ × ℕ -----> ℕ @@ -179,88 +185,56 @@ The [2-adic composition map](elementary-number-theory.2-adic-decomposition.md) $ ℕ. ``` -Since the [image](foundation.images.md) of the 2-adic composition function consists precisely of the [nonzero natural numbers](elementary-number-theory.nonzero-natural-numbers.md), the top map in this triangle is an equivalence. +Since the [image](foundation.images.md) of the 2-adic composition function +consists precisely of the +[nonzero natural numbers](elementary-number-theory.nonzero-natural-numbers.md), +the top map in this triangle is an equivalence. ```agda -pairing-map-ℕ : ℕ × ℕ → ℕ -pairing-map-ℕ (u , v) = - pr1 (is-successor-is-nonzero-ℕ (is-nonzero-2-adic-composition-ℕ u v)) - -compute-succ-pairing-map-ℕ : - (x : ℕ × ℕ) → succ-ℕ (pairing-map-ℕ x) = 2-adic-composition-ℕ (pr1 x) (pr2 x) -compute-succ-pairing-map-ℕ (u , v) = - inv (pr2 (is-successor-is-nonzero-ℕ (is-nonzero-2-adic-composition-ℕ u v))) -``` - -### Pairing function is split surjective - -```text -is-split-surjective-pairing-map : is-split-surjective pairing-map -is-split-surjective-pairing-map n = - (u , v) , is-injective-succ-ℕ (q ∙ s) - where - u = pr1 (pr1 (has-pair-expansion n)) - v = pr2 (pr1 (has-pair-expansion n)) - s = pr2 (has-pair-expansion n) - r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) - q : - ( succ-ℕ (pairing-map (u , v))) = - ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) - q = inv (pr2 r) -``` - -### Pairing function is injective - -```text -is-injective-pairing-map : is-injective pairing-map -is-injective-pairing-map {u , v} {u' , v'} p = - ( eq-pair' (is-pair-expansion-unique u u' v v' q)) - where - r = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u v) - s = is-successor-is-nonzero-ℕ (is-nonzero-pair-expansion u' v') - q : - ( (exp-ℕ 2 u) *ℕ (succ-ℕ (v *ℕ 2))) = - ( (exp-ℕ 2 u') *ℕ (succ-ℕ (v' *ℕ 2))) - q = (pr2 r) ∙ (ap succ-ℕ p ∙ inv (pr2 s)) -``` - -### Pairing function is equivalence - -```text -is-equiv-pairing-map : is-equiv pairing-map -is-equiv-pairing-map = - is-equiv-is-split-surjective-is-injective - pairing-map - is-injective-pairing-map - is-split-surjective-pairing-map -``` +pairing-equiv-ℕ : ℕ × ℕ ≃ ℕ +pairing-equiv-ℕ = + equiv-domain-logical-equiv-fiber-emb + ( 2-adic-composition-emb-ℕ) + ( succ-emb-ℕ) + ( logical-equiv-fiber-2-adic-composition-fiber-succ-ℕ) -```text -ℕ×ℕ≃ℕ : (ℕ × ℕ) ≃ ℕ -ℕ×ℕ≃ℕ = pair pairing-map is-equiv-pairing-map - -map-ℕ-to-ℕ×ℕ : ℕ → ℕ × ℕ -map-ℕ-to-ℕ×ℕ = map-inv-is-equiv (pr2 ℕ×ℕ≃ℕ) - -is-equiv-map-ℕ-to-ℕ×ℕ : is-equiv map-ℕ-to-ℕ×ℕ -is-equiv-map-ℕ-to-ℕ×ℕ = is-equiv-map-inv-is-equiv (pr2 ℕ×ℕ≃ℕ) +pairing-map-ℕ : ℕ × ℕ → ℕ +pairing-map-ℕ = map-equiv pairing-equiv-ℕ + +map-inv-pairing-equiv-ℕ : ℕ → ℕ × ℕ +map-inv-pairing-equiv-ℕ = map-inv-equiv pairing-equiv-ℕ + +is-equiv-map-inv-pairing-equiv-ℕ : is-equiv map-inv-pairing-equiv-ℕ +is-equiv-map-inv-pairing-equiv-ℕ = + is-equiv-map-inv-equiv pairing-equiv-ℕ + +coherence-triangle-pairing-map-ℕ : + coherence-triangle-maps + ( λ x → 2-adic-composition-ℕ (pr1 x) (pr2 x)) + ( succ-ℕ) + ( pairing-map-ℕ) +coherence-triangle-pairing-map-ℕ = + coherence-triangle-equiv-domain-logical-equiv-fiber-emb + ( 2-adic-composition-emb-ℕ) + ( succ-emb-ℕ) + ( logical-equiv-fiber-2-adic-composition-fiber-succ-ℕ) ``` ### The iterated coproduct `ℕ × ℕ × ... × ℕ` is equivalent to `ℕ` for any n -```text +```agda equiv-iterated-product-ℕ : - (n : ℕ) → (iterate n (_×_ ℕ) ℕ) ≃ ℕ + (n : ℕ) → iterate n (ℕ ×_) ℕ ≃ ℕ equiv-iterated-product-ℕ zero-ℕ = id-equiv equiv-iterated-product-ℕ (succ-ℕ n) = - ( ℕ×ℕ≃ℕ) ∘e + ( pairing-equiv-ℕ) ∘e ( equiv-product id-equiv (equiv-iterated-product-ℕ n)) ``` ### The coproduct `(Fin n) + ℕ` is equivalent to `N` for any standard finite `Fin n` ```agda -equiv-coproduct-Fin-ℕ : (n : ℕ) → ((Fin n) + ℕ) ≃ ℕ +equiv-coproduct-Fin-ℕ : (n : ℕ) → Fin n + ℕ ≃ ℕ equiv-coproduct-Fin-ℕ zero-ℕ = left-unit-law-coproduct ℕ equiv-coproduct-Fin-ℕ (succ-ℕ n) = ( equiv-coproduct-Fin-ℕ n) ∘e @@ -271,7 +245,7 @@ equiv-coproduct-Fin-ℕ (succ-ℕ n) = ### The product `(Fin n) × ℕ` is equivalent to `N` for any standard finite `Fin n` where n is nonzero ```agda -equiv-product-Fin-ℕ : (n : ℕ) → ((Fin (succ-ℕ n)) × ℕ) ≃ ℕ +equiv-product-Fin-ℕ : (n : ℕ) → Fin (succ-ℕ n) × ℕ ≃ ℕ equiv-product-Fin-ℕ zero-ℕ = ( left-unit-law-coproduct ℕ) ∘e ( ( equiv-coproduct (left-absorption-product ℕ) left-unit-law-product) ∘e diff --git a/src/foundation/embeddings.lagda.md b/src/foundation/embeddings.lagda.md index 6631d6bead..1c1f5587fc 100644 --- a/src/foundation/embeddings.lagda.md +++ b/src/foundation/embeddings.lagda.md @@ -17,6 +17,7 @@ open import foundation.functoriality-cartesian-product-types open import foundation.functoriality-dependent-pair-types open import foundation.fundamental-theorem-of-identity-types open import foundation.identity-types +open import foundation.logical-equivalences open import foundation.transport-along-identifications open import foundation.truncated-maps open import foundation.universe-levels @@ -454,3 +455,48 @@ module _ ( equiv-fiber f (f a)) ( is-contr-fibers-values-is-emb' e a) ``` + +### A family of logical equivalences between the fibers of two embeddings into a type induces an equivalence between their domain types + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} + (f : A ↪ X) (g : B ↪ X) + (H : (x : X) → fiber (map-emb f) x ↔ fiber (map-emb g) x) + where + + fam-equiv-logical-equiv-fiber-emb : + (x : X) → fiber (map-emb f) x ≃ fiber (map-emb g) x + fam-equiv-logical-equiv-fiber-emb x = + equiv-iff + ( fiber-emb-Prop f x) + ( fiber-emb-Prop g x) + ( forward-implication (H x)) + ( backward-implication (H x)) + + equiv-domain-logical-equiv-fiber-emb : + A ≃ B + equiv-domain-logical-equiv-fiber-emb = + equiv-total-fiber (map-emb g) ∘e + equiv-tot fam-equiv-logical-equiv-fiber-emb ∘e + inv-equiv-total-fiber (map-emb f) + + map-equiv-domain-logical-equiv-fiber-emb : + A → B + map-equiv-domain-logical-equiv-fiber-emb = + map-equiv equiv-domain-logical-equiv-fiber-emb + + coherence-triangle-equiv-domain-logical-equiv-fiber-emb : + coherence-triangle-maps + ( map-emb f) + ( map-emb g) + ( map-equiv-domain-logical-equiv-fiber-emb) + coherence-triangle-equiv-domain-logical-equiv-fiber-emb a = + inv (pr2 (forward-implication (H (map-emb f a)) (a , refl))) +``` + +## See also + +- [Propositional maps (core)](foundation-core.propositional-maps.md) +- [Propositional-maps](foundation.propositional-maps.md) +- [Subtype duality](foundation.subtype-duality.md) diff --git a/src/foundation/subtype-duality.lagda.md b/src/foundation/subtype-duality.lagda.md index abd64501e7..597d3853c9 100644 --- a/src/foundation/subtype-duality.lagda.md +++ b/src/foundation/subtype-duality.lagda.md @@ -22,6 +22,16 @@ open import foundation-core.propositions
+## Idea + +The {{#concept "subtype duality" Agda=equiv-Fiber-Prop}} principle establishes an equivalence + +$$ + \left(\sum_{A:\mathcal{U}} X\hookrightarrow A\right) \simeq (A \to \mathrm{Prop}), +$$ + +In other words, it establishes a unique correspondence between [embeddings](foundation-core.embeddings.md) into a type $A$, and [subtypes](foundation-core.subtypes.md) of $A$. + ## Theorem ### Subtype duality @@ -36,7 +46,11 @@ equiv-Fiber-Prop : equiv-Fiber-Prop l A = ( equiv-Fiber-structure l is-prop A) ∘e ( equiv-tot (λ X → equiv-tot equiv-is-prop-map-is-emb)) +``` + +### Inhabited type duality +```agda Slice-surjection : (l : Level) {l1 : Level} (A : UU l1) → UU (lsuc l ⊔ l1) Slice-surjection l A = Σ (UU l) (λ X → X ↠ A) diff --git a/src/foundation/type-duality.lagda.md b/src/foundation/type-duality.lagda.md index c2beea330e..7582ccd9e2 100644 --- a/src/foundation/type-duality.lagda.md +++ b/src/foundation/type-duality.lagda.md @@ -314,3 +314,4 @@ fiber-Σ {l1} {l2} X A = - In [`foundation.binary-type-duality`](foundation.binary-type-duality.md) we show that [binary relations](foundation.binary-relations.md) are equivalently described as [spans of types](foundation.spans.md). +- [Subtype duality](foundation.subtype-duality.md) diff --git a/src/set-theory/countable-sets.lagda.md b/src/set-theory/countable-sets.lagda.md index 023911e269..68bf2d9ffd 100644 --- a/src/set-theory/countable-sets.lagda.md +++ b/src/set-theory/countable-sets.lagda.md @@ -459,12 +459,12 @@ module _ ( unit-trunc-Prop ( pair ( map-maybe-product ∘ - ( map-product (pr1 h) (pr1 h') ∘ map-ℕ-to-ℕ×ℕ)) + ( map-product (pr1 h) (pr1 h') ∘ map-inv-pairing-equiv-ℕ)) ( is-surjective-comp ( is-surjective-map-maybe-product) ( is-surjective-comp ( is-surjective-map-product (pr2 h) (pr2 h')) - ( is-surjective-is-equiv (is-equiv-map-ℕ-to-ℕ×ℕ))))))) + ( is-surjective-is-equiv (is-equiv-map-inv-pairing-equiv-ℕ))))))) ``` In particular, the sets ℕ + ℕ, ℕ × ℕ, and ℤ are countable. From b4b86ac4de0628beb83fb3894d721c8e4f444336 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Fri, 7 Feb 2025 14:54:41 -0500 Subject: [PATCH 66/72] refactoring up to the fundamental theorem --- .../cofibonacci.lagda.md | 3 +- .../decidable-total-order-integers.lagda.md | 42 ++++++++++++++-- .../fermat-numbers.lagda.md | 2 +- .../finitary-natural-numbers.lagda.md | 9 +++- ...fundamental-theorem-of-arithmetic.lagda.md | 7 ++- .../inequality-integers.lagda.md | 49 ------------------- ...ntrivial-divisors-natural-numbers.lagda.md | 7 ++- .../pisano-periods.lagda.md | 17 +++++-- ...-ordered-pairs-of-natural-numbers.lagda.md | 4 +- src/lists/sort-by-insertion-lists.lagda.md | 3 ++ src/lists/sorted-lists.lagda.md | 1 + .../pigeonhole-principle.lagda.md | 4 ++ .../sequences-finite-types.lagda.md | 3 +- 13 files changed, 83 insertions(+), 68 deletions(-) diff --git a/src/elementary-number-theory/cofibonacci.lagda.md b/src/elementary-number-theory/cofibonacci.lagda.md index 9d2112effe..d88d3387be 100644 --- a/src/elementary-number-theory/cofibonacci.lagda.md +++ b/src/elementary-number-theory/cofibonacci.lagda.md @@ -11,6 +11,7 @@ open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.fibonacci-sequence open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.natural-numbers open import elementary-number-theory.pisano-periods @@ -79,7 +80,7 @@ least-multiple-of-cofibonacci m = well-ordering-principle-ℕ ( is-multiple-of-cofibonacci m) ( is-decidable-is-multiple-of-cofibonacci m) - ( cofibonacci-multiple m) + ( pr2 (cofibonacci-multiple m)) cofibonacci : ℕ → ℕ cofibonacci m = pr1 (least-multiple-of-cofibonacci m) diff --git a/src/elementary-number-theory/decidable-total-order-integers.lagda.md b/src/elementary-number-theory/decidable-total-order-integers.lagda.md index f4e2772ec7..49dab8a4a9 100644 --- a/src/elementary-number-theory/decidable-total-order-integers.lagda.md +++ b/src/elementary-number-theory/decidable-total-order-integers.lagda.md @@ -8,13 +8,20 @@ module elementary-number-theory.decidable-total-order-integers where ```agda open import elementary-number-theory.inequality-integers +open import elementary-number-theory.integers open import foundation.dependent-pair-types open import foundation.propositional-truncations open import foundation.universe-levels +open import order-theory.decidable-posets +open import order-theory.decidable-preorders open import order-theory.decidable-total-orders +open import order-theory.decidable-total-preorders +open import order-theory.posets +open import order-theory.preorders open import order-theory.total-orders +open import order-theory.total-preorders ```
@@ -26,18 +33,47 @@ The type of [integers](elementary-number-theory.integers.md) [standard ordering relation](elementary-number-theory.inequality-integers.md) forms a [decidable total order](order-theory.decidable-total-orders.md). -## Definition +## Definitions + +### The totally ordered set of integers ordered by inequality ```agda +ℤ-Preorder : Preorder lzero lzero +pr1 ℤ-Preorder = ℤ +pr1 (pr2 ℤ-Preorder) = leq-ℤ-Prop +pr1 (pr2 (pr2 ℤ-Preorder)) = refl-leq-ℤ +pr2 (pr2 (pr2 ℤ-Preorder)) = transitive-leq-ℤ + +ℤ-Decidable-Preorder : Decidable-Preorder lzero lzero +pr1 ℤ-Decidable-Preorder = ℤ-Preorder +pr2 ℤ-Decidable-Preorder = is-decidable-leq-ℤ + +ℤ-Poset : Poset lzero lzero +pr1 ℤ-Poset = ℤ-Preorder +pr2 ℤ-Poset a b = antisymmetric-leq-ℤ + +ℤ-Decidable-Poset : Decidable-Poset lzero lzero +pr1 ℤ-Decidable-Poset = ℤ-Poset +pr2 ℤ-Decidable-Poset = is-decidable-leq-ℤ + is-total-leq-ℤ : is-total-Poset ℤ-Poset is-total-leq-ℤ x y = unit-trunc-Prop (linear-leq-ℤ x y) +ℤ-Total-Preorder : Total-Preorder lzero lzero +pr1 ℤ-Total-Preorder = ℤ-Preorder +pr2 ℤ-Total-Preorder a b = unit-trunc-Prop (linear-leq-ℤ a b) + +ℤ-Decidable-Total-Preorder : Decidable-Total-Preorder lzero lzero +pr1 ℤ-Decidable-Total-Preorder = ℤ-Preorder +pr1 (pr2 ℤ-Decidable-Total-Preorder) a b = unit-trunc-Prop (linear-leq-ℤ a b) +pr2 (pr2 ℤ-Decidable-Total-Preorder) = is-decidable-leq-ℤ + ℤ-Total-Order : Total-Order lzero lzero pr1 ℤ-Total-Order = ℤ-Poset -pr2 ℤ-Total-Order = is-total-leq-ℤ +pr2 ℤ-Total-Order a b = unit-trunc-Prop (linear-leq-ℤ a b) ℤ-Decidable-Total-Order : Decidable-Total-Order lzero lzero pr1 ℤ-Decidable-Total-Order = ℤ-Poset -pr1 (pr2 ℤ-Decidable-Total-Order) = is-total-leq-ℤ +pr1 (pr2 ℤ-Decidable-Total-Order) a b = unit-trunc-Prop (linear-leq-ℤ a b) pr2 (pr2 ℤ-Decidable-Total-Order) = is-decidable-leq-ℤ ``` diff --git a/src/elementary-number-theory/fermat-numbers.lagda.md b/src/elementary-number-theory/fermat-numbers.lagda.md index 45eabce82a..9e777d0092 100644 --- a/src/elementary-number-theory/fermat-numbers.lagda.md +++ b/src/elementary-number-theory/fermat-numbers.lagda.md @@ -147,7 +147,7 @@ dist-fermat-number-two-ℕ : Π-ℕ n (λ k → fermat-number-ℕ (nat-Fin n k)) dist-fermat-number-two-ℕ zero-ℕ = refl dist-fermat-number-two-ℕ (succ-ℕ n) = - ap (λ x → dist-ℕ x 1) (ap (2 ^ℕ_) (exp-succ-ℕ n 2) ∙ exp-mul-ℕ (2 ^ℕ n) 2) ∙ + ap (λ x → dist-ℕ x 1) (ap (2 ^ℕ_) (exp-succ-ℕ 2 n) ∙ exp-mul-ℕ (2 ^ℕ n) 2) ∙ distance-of-squares-ℕ (2 ^ℕ 2 ^ℕ n) 1 ∙ ap (_*ℕ fermat-number-ℕ n) (dist-fermat-number-two-ℕ n) ``` diff --git a/src/elementary-number-theory/finitary-natural-numbers.lagda.md b/src/elementary-number-theory/finitary-natural-numbers.lagda.md index 6e6f189d42..8a90e3a2f7 100644 --- a/src/elementary-number-theory/finitary-natural-numbers.lagda.md +++ b/src/elementary-number-theory/finitary-natural-numbers.lagda.md @@ -88,7 +88,10 @@ cong-unary-op-ℕ (succ-ℕ k) x n = le-constant-unary-op-ℕ : (k : ℕ) (x y : Fin k) (m : ℕ) → le-ℕ (constant-ℕ k x) (unary-op-ℕ k y m) le-constant-unary-op-ℕ k x y m = - concatenate-le-leq-ℕ {nat-Fin k x} {k} {unary-op-ℕ k y m} + concatenate-le-leq-ℕ + ( nat-Fin k x) + ( k) + ( unary-op-ℕ k y m) ( strict-upper-bound-nat-Fin k x) ( transitive-leq-ℕ ( k) @@ -110,6 +113,8 @@ is-injective-convert-based-ℕ { unary-op-based-ℕ .(succ-ℕ k) y m} p = ex-falso ( neq-le-ℕ + ( convert-based-ℕ (succ-ℕ k) (constant-based-ℕ (succ-ℕ k) x)) + ( convert-based-ℕ (succ-ℕ k) (unary-op-based-ℕ (succ-ℕ k) y m)) ( le-constant-unary-op-ℕ (succ-ℕ k) x y (convert-based-ℕ (succ-ℕ k) m)) ( p)) is-injective-convert-based-ℕ @@ -118,6 +123,8 @@ is-injective-convert-based-ℕ { constant-based-ℕ .(succ-ℕ k) y} p = ex-falso ( neq-le-ℕ + ( convert-based-ℕ (succ-ℕ k) (constant-based-ℕ (succ-ℕ k) y)) + ( convert-based-ℕ (succ-ℕ k) (unary-op-based-ℕ (succ-ℕ k) x n)) ( le-constant-unary-op-ℕ (succ-ℕ k) y x (convert-based-ℕ (succ-ℕ k) n)) ( inv p)) is-injective-convert-based-ℕ diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index 40da861feb..b2b7578234 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -15,6 +15,7 @@ open import elementary-number-theory.equality-natural-numbers open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lists-of-prime-numbers open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.multiplication-lists-of-natural-numbers open import elementary-number-theory.multiplication-natural-numbers @@ -44,6 +45,7 @@ open import foundation.unit-type open import foundation.universe-levels open import lists.concatenation-lists +open import lists.elementhood-relation-lists open import lists.functoriality-lists open import lists.lists open import lists.permutation-lists @@ -243,10 +245,7 @@ module _ le-one-is-prime-decomposition-cons-list-ℕ : le-ℕ 1 x le-one-is-prime-decomposition-cons-list-ℕ = - concatenate-le-leq-ℕ - {x = 1} - {y = y} - {z = x} + concatenate-le-leq-ℕ 1 y x ( le-one-head-is-prime-decomposition-cons-list-ℕ) ( leq-div-ℕ ( y) diff --git a/src/elementary-number-theory/inequality-integers.lagda.md b/src/elementary-number-theory/inequality-integers.lagda.md index 423e6669f5..14159cf347 100644 --- a/src/elementary-number-theory/inequality-integers.lagda.md +++ b/src/elementary-number-theory/inequality-integers.lagda.md @@ -34,15 +34,6 @@ open import foundation.propositions open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels - -open import order-theory.decidable-posets -open import order-theory.decidable-preorders -open import order-theory.decidable-total-preorders -open import order-theory.decidable-total-orders -open import order-theory.posets -open import order-theory.preorders -open import order-theory.total-preorders -open import order-theory.total-orders ```
@@ -156,46 +147,6 @@ linear-leq-ℤ a b = ( decide-is-negative-is-nonnegative-ℤ) ``` -### The partially ordered set of integers ordered by inequality - -```agda -ℤ-Preorder : Preorder lzero lzero -pr1 ℤ-Preorder = ℤ -pr1 (pr2 ℤ-Preorder) = leq-ℤ-Prop -pr1 (pr2 (pr2 ℤ-Preorder)) = refl-leq-ℤ -pr2 (pr2 (pr2 ℤ-Preorder)) = transitive-leq-ℤ - -ℤ-Decidable-Preorder : Decidable-Preorder lzero lzero -pr1 ℤ-Decidable-Preorder = ℤ-Preorder -pr2 ℤ-Decidable-Preorder = is-decidable-leq-ℤ - -ℤ-Poset : Poset lzero lzero -pr1 ℤ-Poset = ℤ-Preorder -pr2 ℤ-Poset a b = antisymmetric-leq-ℤ - -ℤ-Decidable-Poset : Decidable-Poset lzero lzero -pr1 ℤ-Decidable-Poset = ℤ-Poset -pr2 ℤ-Decidable-Poset = is-decidable-leq-ℤ - -ℤ-Total-Preorder : Total-Preorder lzero lzero -pr1 ℤ-Total-Preorder = ℤ-Preorder -pr2 ℤ-Total-Preorder a b = unit-trunc-Prop (linear-leq-ℤ a b) - -ℤ-Decidable-Total-Preorder : Decidable-Total-Preorder lzero lzero -pr1 ℤ-Decidable-Total-Preorder = ℤ-Preorder -pr1 (pr2 ℤ-Decidable-Total-Preorder) a b = unit-trunc-Prop (linear-leq-ℤ a b) -pr2 (pr2 ℤ-Decidable-Total-Preorder) = is-decidable-leq-ℤ - -ℤ-Total-Order : Total-Order lzero lzero -pr1 ℤ-Total-Order = ℤ-Poset -pr2 ℤ-Total-Order a b = unit-trunc-Prop (linear-leq-ℤ a b) - -ℤ-Decidable-Total-Order : Decidable-Total-Order lzero lzero -pr1 ℤ-Decidable-Total-Order = ℤ-Poset -pr1 (pr2 ℤ-Decidable-Total-Order) a b = unit-trunc-Prop (linear-leq-ℤ a b) -pr2 (pr2 ℤ-Decidable-Total-Order) = is-decidable-leq-ℤ -``` - ### Chaining rules for equality and inequality ```agda diff --git a/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md index ba53142781..9852a68d93 100644 --- a/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md +++ b/src/elementary-number-theory/nontrivial-divisors-natural-numbers.lagda.md @@ -9,6 +9,7 @@ module elementary-number-theory.nontrivial-divisors-natural-numbers where ```agda open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.inequality-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.prime-numbers open import elementary-number-theory.strict-inequality-natural-numbers @@ -139,7 +140,7 @@ least-nontrivial-divisor-ℕ n H = well-ordering-principle-ℕ ( is-nontrivial-divisor-ℕ n) ( is-decidable-is-nontrivial-divisor-ℕ n) - ( n , is-nontrivial-divisor-diagonal-ℕ n H) + ( is-nontrivial-divisor-diagonal-ℕ n H) nat-least-nontrivial-divisor-ℕ : (n : ℕ) → le-ℕ 1 n → ℕ nat-least-nontrivial-divisor-ℕ n H = pr1 (least-nontrivial-divisor-ℕ n H) @@ -213,7 +214,9 @@ pr1 (is-prime-least-nontrivial-divisor-ℕ n H x) (K , L) = ( L)) ( leq-le-ℕ 1 n H))) pr1 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = - neq-le-ℕ (le-one-least-nontrivial-divisor-ℕ n H) + neq-le-ℕ 1 + ( nat-least-nontrivial-divisor-ℕ n H) + ( le-one-least-nontrivial-divisor-ℕ n H) pr2 (pr2 (is-prime-least-nontrivial-divisor-ℕ n H .1) refl) = div-one-ℕ _ ``` diff --git a/src/elementary-number-theory/pisano-periods.lagda.md b/src/elementary-number-theory/pisano-periods.lagda.md index 690e753fa1..e6720052f1 100644 --- a/src/elementary-number-theory/pisano-periods.lagda.md +++ b/src/elementary-number-theory/pisano-periods.lagda.md @@ -11,6 +11,7 @@ open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.fibonacci-sequence open import elementary-number-theory.inequality-natural-numbers open import elementary-number-theory.lower-bounds-natural-numbers +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers @@ -173,6 +174,9 @@ cases-is-repetition-of-zero-pisano-period k zero-ℕ y p q = refl cases-is-repetition-of-zero-pisano-period k (succ-ℕ x) zero-ℕ p q = ex-falso ( concatenate-eq-le-eq-ℕ + ( succ-ℕ x) + ( pr1 (is-ordered-repetition-pisano-period k)) + ( pisano-period k) zero-ℕ ( inv p) ( pr1 (pr2 (is-ordered-repetition-pisano-period k))) ( q)) @@ -184,6 +188,10 @@ cases-is-repetition-of-zero-pisano-period k (succ-ℕ x) (succ-ℕ y) p q = R : is-ordered-repetition-fibonacci-pair-Fin k y R = triple x ( concatenate-eq-le-eq-ℕ + ( succ-ℕ x) + ( pr1 (is-ordered-repetition-pisano-period k)) + ( pisano-period k) + ( succ-ℕ y) ( inv p) ( pr1 (pr2 (is-ordered-repetition-pisano-period k))) ( q)) @@ -219,16 +227,17 @@ le-zero-pisano-period : (k : ℕ) → le-ℕ zero-ℕ (pisano-period k) le-zero-pisano-period k = concatenate-eq-le-ℕ - { x = zero-ℕ} - { y = pr1 (is-ordered-repetition-pisano-period k)} - { z = pisano-period k} + ( zero-ℕ) + ( pr1 (is-ordered-repetition-pisano-period k)) + ( pisano-period k) ( inv (is-repetition-of-zero-pisano-period k)) ( pr1 (pr2 (is-ordered-repetition-pisano-period k))) is-nonzero-pisano-period : (k : ℕ) → is-nonzero-ℕ (pisano-period k) is-nonzero-pisano-period k p = - ex-falso (concatenate-le-eq-ℕ (le-zero-pisano-period k) p) + ex-falso + ( concatenate-le-eq-ℕ 0 (pisano-period k) 0 (le-zero-pisano-period k) p) ``` ### `k + 1` divides the Fibonacci number at `pisano-period k` diff --git a/src/elementary-number-theory/strictly-ordered-pairs-of-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-ordered-pairs-of-natural-numbers.lagda.md index faab496e28..fa7d1c929b 100644 --- a/src/elementary-number-theory/strictly-ordered-pairs-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-ordered-pairs-of-natural-numbers.lagda.md @@ -26,7 +26,7 @@ open import foundation.universe-levels ## Idea -A strictly ordered pair of natural numbers consists of `x y : ℕ` such that +A {{#concept "strictly ordered pair of natural numbers" Agda=strictly-ordered-pair-ℕ}} consists of `x y : ℕ` satisfying the [strict inequality](elementary-number-theory.strict-inequality-natural-numbers.md) `x < y`. ## Definition @@ -58,7 +58,7 @@ module _ pair-of-distinct-elements-strictly-ordered-pair-ℕ : strictly-ordered-pair-ℕ → pair-of-distinct-elements ℕ pair-of-distinct-elements-strictly-ordered-pair-ℕ (a , b , H) = - (a , b , neq-le-ℕ H) + (a , b , neq-le-ℕ a b H) ``` ### Any pair of distinct elements of natural numbers can be strictly ordered diff --git a/src/lists/sort-by-insertion-lists.lagda.md b/src/lists/sort-by-insertion-lists.lagda.md index c14433df58..deca8a7272 100644 --- a/src/lists/sort-by-insertion-lists.lagda.md +++ b/src/lists/sort-by-insertion-lists.lagda.md @@ -13,6 +13,9 @@ open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.universe-levels +open import linear-algebra.vectors + + open import lists.arrays open import lists.lists open import lists.permutation-lists diff --git a/src/lists/sorted-lists.lagda.md b/src/lists/sorted-lists.lagda.md index 1ece7eb3c0..ccdb8a6727 100644 --- a/src/lists/sorted-lists.lagda.md +++ b/src/lists/sorted-lists.lagda.md @@ -17,6 +17,7 @@ open import foundation.universe-levels open import linear-algebra.vectors open import lists.arrays +open import lists.elementhood-relation-lists open import lists.lists open import lists.sorted-vectors diff --git a/src/univalent-combinatorics/pigeonhole-principle.lagda.md b/src/univalent-combinatorics/pigeonhole-principle.lagda.md index 8c52386070..e845f9d55a 100644 --- a/src/univalent-combinatorics/pigeonhole-principle.lagda.md +++ b/src/univalent-combinatorics/pigeonhole-principle.lagda.md @@ -347,6 +347,10 @@ module _ apply-universal-property-trunc-Prop K empty-Prop ( λ d → is-not-emb-le-count e d f ( concatenate-eq-le-eq-ℕ + ( number-of-elements-count d) + ( number-of-elements-is-finite K) + ( number-of-elements-is-finite H) + ( number-of-elements-count e) ( compute-number-of-elements-is-finite d K) ( p) ( inv (compute-number-of-elements-is-finite e H))) diff --git a/src/univalent-combinatorics/sequences-finite-types.lagda.md b/src/univalent-combinatorics/sequences-finite-types.lagda.md index e9cd8106af..d908fcbaea 100644 --- a/src/univalent-combinatorics/sequences-finite-types.lagda.md +++ b/src/univalent-combinatorics/sequences-finite-types.lagda.md @@ -8,6 +8,7 @@ module univalent-combinatorics.sequences-finite-types where ```agda open import elementary-number-theory.decidable-types +open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.strict-inequality-natural-numbers open import elementary-number-theory.well-ordering-principle-natural-numbers @@ -135,7 +136,7 @@ minimal-element-repetition-of-values-sequence-Fin k f = is-decidable-strictly-bounded-Σ-ℕ' x ( λ y → f y = f x) ( λ y → has-decidable-equality-Fin k (f y) (f x))) - ( v , u , H , p) + ( u , H , p) where r = ordered-repetition-of-values-sequence-Fin k f u = pr1 (pr1 r) From 0661b4b56ef8927f72dca77234d2ba078af49830 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 9 Feb 2025 13:41:41 -0500 Subject: [PATCH 67/72] some work towards Bezout's identity in the natural numbers --- .../divisibility-natural-numbers.lagda.md | 1 + .../prime-divisors-natural-numbers.lagda.md | 77 +++ src/group-theory/commutative-monoids.lagda.md | 6 + src/group-theory/monoids.lagda.md | 8 + src/group-theory/semigroups.lagda.md | 6 + src/ring-theory/ideals-semirings.lagda.md | 271 +++------ ...ideals-generated-by-subsets-rings.lagda.md | 6 +- ...ls-generated-by-subsets-semirings.lagda.md | 527 ++++++++++++++++++ src/ring-theory/left-ideals-rings.lagda.md | 101 ++-- .../left-ideals-semirings.lagda.md | 194 +++++++ ...ar-combinations-of-elements-rings.lagda.md | 164 ++++++ ...ombinations-of-elements-semirings.lagda.md | 306 ++++++++++ ...ar-combinations-of-elements-rings.lagda.md | 200 +++++++ ...ombinations-of-elements-semirings.lagda.md | 200 +++++++ .../monoids-with-semiring-actions.lagda.md | 207 +++++++ .../poset-of-left-ideals-semirings.lagda.md | 178 ++++++ .../right-ideals-semirings.lagda.md | 195 +++++++ ...ar-combinations-of-elements-rings.lagda.md | 163 ++++++ ...ombinations-of-elements-semirings.lagda.md | 163 ++++++ src/ring-theory/rings.lagda.md | 5 + src/ring-theory/semirings.lagda.md | 6 + src/ring-theory/subsets-rings.lagda.md | 165 ++++-- src/ring-theory/subsets-semirings.lagda.md | 78 ++- src/structured-types/magmas.lagda.md | 42 +- 24 files changed, 2971 insertions(+), 298 deletions(-) create mode 100644 src/elementary-number-theory/prime-divisors-natural-numbers.lagda.md create mode 100644 src/ring-theory/left-ideals-generated-by-subsets-semirings.lagda.md create mode 100644 src/ring-theory/left-ideals-semirings.lagda.md create mode 100644 src/ring-theory/left-linear-combinations-of-elements-rings.lagda.md create mode 100644 src/ring-theory/left-linear-combinations-of-elements-semirings.lagda.md create mode 100644 src/ring-theory/linear-combinations-of-elements-rings.lagda.md create mode 100644 src/ring-theory/linear-combinations-of-elements-semirings.lagda.md create mode 100644 src/ring-theory/monoids-with-semiring-actions.lagda.md create mode 100644 src/ring-theory/poset-of-left-ideals-semirings.lagda.md create mode 100644 src/ring-theory/right-ideals-semirings.lagda.md create mode 100644 src/ring-theory/right-linear-combinations-of-elements-rings.lagda.md create mode 100644 src/ring-theory/right-linear-combinations-of-elements-semirings.lagda.md diff --git a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md index 19c7930e91..83d2779a26 100644 --- a/src/elementary-number-theory/divisibility-natural-numbers.lagda.md +++ b/src/elementary-number-theory/divisibility-natural-numbers.lagda.md @@ -807,5 +807,6 @@ le-quotient-div-ℕ d x f H g = - [Bounded divisibility of natural numbers](elementary-number-theory.bounded-divisibility-natural-numbers.md) - [Greatest common divisors of natural numbers](elementary-number-theory.greatest-common-divisors-natural-numbers.md) - [Nontrivial divisors of natural numbers](elementary-number-theory.nontrivial-divisors-natural-numbers.md) +- [Prime divisors](elementary-number-theory.prime-divisors-natural-numbers.md) - [Proper divisors of natural numbers](elementary-number-theory.proper-divisors-natural-numbers.md) - [The poset of natural numbers ordered by divisibility](elementary-number-theory.poset-of-natural-numbers-ordered-by-divisibility.md) diff --git a/src/elementary-number-theory/prime-divisors-natural-numbers.lagda.md b/src/elementary-number-theory/prime-divisors-natural-numbers.lagda.md new file mode 100644 index 0000000000..25a07823ea --- /dev/null +++ b/src/elementary-number-theory/prime-divisors-natural-numbers.lagda.md @@ -0,0 +1,77 @@ +# Prime divisors of natural numbers + +```agda +module elementary-number-theory.prime-divisors-natural-numbers where +``` + +
Imports + +```agda +open import elementary-number-theory.divisibility-natural-numbers +open import elementary-number-theory.natural-numbers +open import elementary-number-theory.prime-numbers + +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-propositions +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.universe-levels +``` + +
+ +## Idea + +A {{#concept "prime divisor" Disambiguation="natural numbers" Agda=is-prime-divisor-ℕ}} of a [natural number](elementary-number-theory.natural-numbers.md) $n$ is a [prime number](elementary-number-theory.prime-numbers.md) $p$ which is a [divisor](elementary-number-theory.divisibility-natural-numbers.md) of $n$. + +## Definitions + +### The predicate of being a prime divisor + +```agda +module _ + (n d : ℕ) + where + + is-prime-divisor-ℕ : + UU lzero + is-prime-divisor-ℕ = + is-prime-ℕ d × is-divisor-ℕ n d + + is-prop-is-prime-divisor-ℕ : + is-prop is-prime-divisor-ℕ + is-prop-is-prime-divisor-ℕ = + is-prop-Σ + ( is-prop-is-prime-ℕ d) + ( λ H → is-prop-div-ℕ d n (inl (is-nonzero-is-prime-ℕ d H))) + + is-prime-divisor-prop-ℕ : + Prop lzero + pr1 is-prime-divisor-prop-ℕ = + is-prime-divisor-ℕ + pr2 is-prime-divisor-prop-ℕ = + is-prop-is-prime-divisor-ℕ + + is-decidable-is-prime-divisor-ℕ : + is-decidable is-prime-divisor-ℕ + is-decidable-is-prime-divisor-ℕ = + is-decidable-product (is-decidable-is-prime-ℕ d) (is-decidable-div-ℕ d n) + + is-decidable-prop-is-prime-divisor-ℕ : + is-decidable-prop is-prime-divisor-ℕ + pr1 is-decidable-prop-is-prime-divisor-ℕ = + is-prop-is-prime-divisor-ℕ + pr2 is-decidable-prop-is-prime-divisor-ℕ = + is-decidable-is-prime-divisor-ℕ + + is-prime-divisor-decidable-prop-ℕ : + Decidable-Prop lzero + pr1 is-prime-divisor-decidable-prop-ℕ = + is-prime-divisor-ℕ + pr2 is-prime-divisor-decidable-prop-ℕ = + is-decidable-prop-is-prime-divisor-ℕ +``` + + diff --git a/src/group-theory/commutative-monoids.lagda.md b/src/group-theory/commutative-monoids.lagda.md index 89b8b3f6aa..a3cde60757 100644 --- a/src/group-theory/commutative-monoids.lagda.md +++ b/src/group-theory/commutative-monoids.lagda.md @@ -19,6 +19,8 @@ open import foundation.universe-levels open import group-theory.monoids open import group-theory.semigroups + +open import structured-types.magmas ``` @@ -66,6 +68,10 @@ module _ monoid-Commutative-Monoid : Monoid l monoid-Commutative-Monoid = pr1 M + unital-magma-Commutative-Monoid : Unital-Magma l + unital-magma-Commutative-Monoid = + unital-magma-Monoid monoid-Commutative-Monoid + semigroup-Commutative-Monoid : Semigroup l semigroup-Commutative-Monoid = semigroup-Monoid monoid-Commutative-Monoid diff --git a/src/group-theory/monoids.lagda.md b/src/group-theory/monoids.lagda.md index 1e76dbc225..0afefc88de 100644 --- a/src/group-theory/monoids.lagda.md +++ b/src/group-theory/monoids.lagda.md @@ -18,6 +18,7 @@ open import foundation.universe-levels open import group-theory.semigroups +open import structured-types.magmas open import structured-types.h-spaces open import structured-types.wild-monoids ``` @@ -48,6 +49,9 @@ module _ semigroup-Monoid : Semigroup l semigroup-Monoid = pr1 M + magma-Monoid : Magma l + magma-Monoid = magma-Semigroup semigroup-Monoid + is-unital-Monoid : is-unital-Semigroup semigroup-Monoid is-unital-Monoid = pr2 M @@ -88,6 +92,10 @@ module _ right-unit-law-mul-Monoid : (x : type-Monoid) → mul-Monoid x unit-Monoid = x right-unit-law-mul-Monoid = pr2 (pr2 has-unit-Monoid) + unital-magma-Monoid : Unital-Magma l + pr1 unital-magma-Monoid = magma-Monoid + pr2 unital-magma-Monoid = has-unit-Monoid + left-swap-mul-Monoid : {x y z : type-Monoid} → mul-Monoid x y = mul-Monoid y x → mul-Monoid x (mul-Monoid y z) = diff --git a/src/group-theory/semigroups.lagda.md b/src/group-theory/semigroups.lagda.md index 4709573fcf..fda03b20cd 100644 --- a/src/group-theory/semigroups.lagda.md +++ b/src/group-theory/semigroups.lagda.md @@ -13,6 +13,8 @@ open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.sets open import foundation.universe-levels + +open import structured-types.magmas ``` @@ -99,6 +101,10 @@ module _ ( associative-mul-Semigroup _ _ _) ∙ ( ap (mul-Semigroup _) (left-swap-mul-Semigroup H)) ∙ ( inv (associative-mul-Semigroup _ _ _)) + + magma-Semigroup : Magma l + pr1 magma-Semigroup = type-Semigroup + pr2 magma-Semigroup = mul-Semigroup ``` ### The structure of a semigroup diff --git a/src/ring-theory/ideals-semirings.lagda.md b/src/ring-theory/ideals-semirings.lagda.md index 801c43f0b9..48c987fdd1 100644 --- a/src/ring-theory/ideals-semirings.lagda.md +++ b/src/ring-theory/ideals-semirings.lagda.md @@ -9,8 +9,13 @@ module ring-theory.ideals-semirings where ```agda open import foundation.cartesian-product-types open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fundamental-theorem-of-identity-types open import foundation.identity-types open import foundation.propositions +open import foundation.subtype-identity-principle +open import foundation.subtypes +open import foundation.torsorial-type-families open import foundation.universe-levels open import group-theory.submonoids @@ -36,200 +41,26 @@ correspondence to congruences. We will call such ideals **normal**. ## Definitions -### Additive submonoids - -```agda -module _ - {l1 : Level} (R : Semiring l1) - where - - is-additive-submonoid-Semiring : - {l2 : Level} → subset-Semiring l2 R → UU (l1 ⊔ l2) - is-additive-submonoid-Semiring = - is-submonoid-subset-Monoid (additive-monoid-Semiring R) -``` - -### Left ideals - -```agda -module _ - {l1 : Level} (R : Semiring l1) - where - - is-left-ideal-subset-Semiring : - {l2 : Level} → subset-Semiring l2 R → UU (l1 ⊔ l2) - is-left-ideal-subset-Semiring P = - is-additive-submonoid-Semiring R P × - is-closed-under-left-multiplication-subset-Semiring R P - -left-ideal-Semiring : - (l : Level) {l1 : Level} (R : Semiring l1) → UU (lsuc l ⊔ l1) -left-ideal-Semiring l R = - Σ (subset-Semiring l R) (is-left-ideal-subset-Semiring R) - -module _ - {l1 l2 : Level} (R : Semiring l1) (I : left-ideal-Semiring l2 R) - where - - subset-left-ideal-Semiring : subset-Semiring l2 R - subset-left-ideal-Semiring = pr1 I - - is-in-left-ideal-Semiring : type-Semiring R → UU l2 - is-in-left-ideal-Semiring x = type-Prop (subset-left-ideal-Semiring x) - - type-left-ideal-Semiring : UU (l1 ⊔ l2) - type-left-ideal-Semiring = type-subset-Semiring R subset-left-ideal-Semiring - - inclusion-left-ideal-Semiring : type-left-ideal-Semiring → type-Semiring R - inclusion-left-ideal-Semiring = - inclusion-subset-Semiring R subset-left-ideal-Semiring - - ap-inclusion-left-ideal-Semiring : - (x y : type-left-ideal-Semiring) → x = y → - inclusion-left-ideal-Semiring x = inclusion-left-ideal-Semiring y - ap-inclusion-left-ideal-Semiring = - ap-inclusion-subset-Semiring R subset-left-ideal-Semiring - - is-in-subset-inclusion-left-ideal-Semiring : - (x : type-left-ideal-Semiring) → - is-in-left-ideal-Semiring (inclusion-left-ideal-Semiring x) - is-in-subset-inclusion-left-ideal-Semiring = - is-in-subset-inclusion-subset-Semiring R subset-left-ideal-Semiring - - is-closed-under-eq-left-ideal-Semiring : - {x y : type-Semiring R} → is-in-left-ideal-Semiring x → - (x = y) → is-in-left-ideal-Semiring y - is-closed-under-eq-left-ideal-Semiring = - is-closed-under-eq-subset-Semiring R subset-left-ideal-Semiring - - is-closed-under-eq-left-ideal-Semiring' : - {x y : type-Semiring R} → is-in-left-ideal-Semiring y → - (x = y) → is-in-left-ideal-Semiring x - is-closed-under-eq-left-ideal-Semiring' = - is-closed-under-eq-subset-Semiring' R subset-left-ideal-Semiring - - is-left-ideal-subset-left-ideal-Semiring : - is-left-ideal-subset-Semiring R subset-left-ideal-Semiring - is-left-ideal-subset-left-ideal-Semiring = pr2 I - - is-additive-submonoid-left-ideal-Semiring : - is-additive-submonoid-Semiring R subset-left-ideal-Semiring - is-additive-submonoid-left-ideal-Semiring = - pr1 is-left-ideal-subset-left-ideal-Semiring - - contains-zero-left-ideal-Semiring : - is-in-left-ideal-Semiring (zero-Semiring R) - contains-zero-left-ideal-Semiring = - pr1 is-additive-submonoid-left-ideal-Semiring - - is-closed-under-addition-left-ideal-Semiring : - is-closed-under-addition-subset-Semiring R subset-left-ideal-Semiring - is-closed-under-addition-left-ideal-Semiring = - pr2 is-additive-submonoid-left-ideal-Semiring - - is-closed-under-left-multiplication-left-ideal-Semiring : - is-closed-under-left-multiplication-subset-Semiring R - subset-left-ideal-Semiring - is-closed-under-left-multiplication-left-ideal-Semiring = - pr2 is-left-ideal-subset-left-ideal-Semiring -``` - -### Right ideals - -```agda -module _ - {l1 : Level} (R : Semiring l1) - where - - is-right-ideal-subset-Semiring : - {l2 : Level} → subset-Semiring l2 R → UU (l1 ⊔ l2) - is-right-ideal-subset-Semiring P = - is-additive-submonoid-Semiring R P × - is-closed-under-right-multiplication-subset-Semiring R P - -right-ideal-Semiring : - (l : Level) {l1 : Level} (R : Semiring l1) → UU (lsuc l ⊔ l1) -right-ideal-Semiring l R = - Σ (subset-Semiring l R) (is-right-ideal-subset-Semiring R) - -module _ - {l1 l2 : Level} (R : Semiring l1) (I : right-ideal-Semiring l2 R) - where - - subset-right-ideal-Semiring : subset-Semiring l2 R - subset-right-ideal-Semiring = pr1 I - - is-in-right-ideal-Semiring : type-Semiring R → UU l2 - is-in-right-ideal-Semiring x = type-Prop (subset-right-ideal-Semiring x) - - type-right-ideal-Semiring : UU (l1 ⊔ l2) - type-right-ideal-Semiring = - type-subset-Semiring R subset-right-ideal-Semiring - - inclusion-right-ideal-Semiring : type-right-ideal-Semiring → type-Semiring R - inclusion-right-ideal-Semiring = - inclusion-subset-Semiring R subset-right-ideal-Semiring - - ap-inclusion-right-ideal-Semiring : - (x y : type-right-ideal-Semiring) → x = y → - inclusion-right-ideal-Semiring x = inclusion-right-ideal-Semiring y - ap-inclusion-right-ideal-Semiring = - ap-inclusion-subset-Semiring R subset-right-ideal-Semiring - - is-in-subset-inclusion-right-ideal-Semiring : - (x : type-right-ideal-Semiring) → - is-in-right-ideal-Semiring (inclusion-right-ideal-Semiring x) - is-in-subset-inclusion-right-ideal-Semiring = - is-in-subset-inclusion-subset-Semiring R subset-right-ideal-Semiring - - is-closed-under-eq-right-ideal-Semiring : - {x y : type-Semiring R} → is-in-right-ideal-Semiring x → - (x = y) → is-in-right-ideal-Semiring y - is-closed-under-eq-right-ideal-Semiring = - is-closed-under-eq-subset-Semiring R subset-right-ideal-Semiring - - is-closed-under-eq-right-ideal-Semiring' : - {x y : type-Semiring R} → is-in-right-ideal-Semiring y → - (x = y) → is-in-right-ideal-Semiring x - is-closed-under-eq-right-ideal-Semiring' = - is-closed-under-eq-subset-Semiring' R subset-right-ideal-Semiring - - is-right-ideal-subset-right-ideal-Semiring : - is-right-ideal-subset-Semiring R subset-right-ideal-Semiring - is-right-ideal-subset-right-ideal-Semiring = pr2 I - - is-additive-submonoid-right-ideal-Semiring : - is-additive-submonoid-Semiring R subset-right-ideal-Semiring - is-additive-submonoid-right-ideal-Semiring = - pr1 is-right-ideal-subset-right-ideal-Semiring - - contains-zero-right-ideal-Semiring : - is-in-right-ideal-Semiring (zero-Semiring R) - contains-zero-right-ideal-Semiring = - pr1 is-additive-submonoid-right-ideal-Semiring - - is-closed-under-addition-right-ideal-Semiring : - is-closed-under-addition-subset-Semiring R subset-right-ideal-Semiring - is-closed-under-addition-right-ideal-Semiring = - pr2 is-additive-submonoid-right-ideal-Semiring - - is-closed-under-right-multiplication-right-ideal-Semiring : - is-closed-under-right-multiplication-subset-Semiring R - subset-right-ideal-Semiring - is-closed-under-right-multiplication-right-ideal-Semiring = - pr2 is-right-ideal-subset-right-ideal-Semiring -``` - -### Two-sided ideals +### Ideals of semirings ```agda is-ideal-subset-Semiring : {l1 l2 : Level} (R : Semiring l1) (P : subset-Semiring l2 R) → UU (l1 ⊔ l2) is-ideal-subset-Semiring R P = - is-additive-submonoid-Semiring R P × + is-additive-submonoid-subset-Semiring R P × ( is-closed-under-left-multiplication-subset-Semiring R P × is-closed-under-right-multiplication-subset-Semiring R P) +is-prop-is-ideal-subset-Semiring : + {l1 l2 : Level} (R : Semiring l1) (P : subset-Semiring l2 R) → + is-prop (is-ideal-subset-Semiring R P) +is-prop-is-ideal-subset-Semiring R P = + is-prop-product + ( is-prop-is-additive-submonoid-subset-Semiring R P) + ( is-prop-product + ( is-prop-is-closed-under-left-multiplication-subset-Semiring R P) + ( is-prop-is-closed-under-right-multiplication-subset-Semiring R P)) + ideal-Semiring : (l : Level) {l1 : Level} (R : Semiring l1) → UU (lsuc l ⊔ l1) ideal-Semiring l R = @@ -242,6 +73,10 @@ module _ subset-ideal-Semiring : subset-Semiring l2 R subset-ideal-Semiring = pr1 I + is-ideal-ideal-Semiring : + is-ideal-subset-Semiring R subset-ideal-Semiring + is-ideal-ideal-Semiring = pr2 I + is-in-ideal-Semiring : type-Semiring R → UU l2 is-in-ideal-Semiring = is-in-subset-Semiring R subset-ideal-Semiring @@ -289,7 +124,7 @@ module _ is-ideal-subset-ideal-Semiring = pr2 I is-additive-submonoid-ideal-Semiring : - is-additive-submonoid-Semiring R subset-ideal-Semiring + is-additive-submonoid-subset-Semiring R subset-ideal-Semiring is-additive-submonoid-ideal-Semiring = pr1 is-ideal-subset-ideal-Semiring @@ -318,3 +153,65 @@ module _ pr1 (pr2 submonoid-ideal-Semiring) = contains-zero-ideal-Semiring pr2 (pr2 submonoid-ideal-Semiring) = is-closed-under-addition-ideal-Semiring ``` + +## Properties + +### Characterizing equality of ideals in semirings + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) (I : ideal-Semiring l2 R) + where + + has-same-elements-ideal-Semiring : + (J : ideal-Semiring l3 R) → UU (l1 ⊔ l2 ⊔ l3) + has-same-elements-ideal-Semiring J = + has-same-elements-subtype + ( subset-ideal-Semiring R I) + ( subset-ideal-Semiring R J) + +module _ + {l1 l2 : Level} (R : Semiring l1) (I : ideal-Semiring l2 R) + where + + refl-has-same-elements-ideal-Semiring : + has-same-elements-ideal-Semiring R I I + refl-has-same-elements-ideal-Semiring = + refl-has-same-elements-subtype (subset-ideal-Semiring R I) + + is-torsorial-has-same-elements-ideal-Semiring : + is-torsorial (has-same-elements-ideal-Semiring R I) + is-torsorial-has-same-elements-ideal-Semiring = + is-torsorial-Eq-subtype + ( is-torsorial-has-same-elements-subtype (subset-ideal-Semiring R I)) + ( is-prop-is-ideal-subset-Semiring R) + ( subset-ideal-Semiring R I) + ( refl-has-same-elements-ideal-Semiring) + ( is-ideal-ideal-Semiring R I) + + has-same-elements-eq-ideal-Semiring : + (J : ideal-Semiring l2 R) → + (I = J) → has-same-elements-ideal-Semiring R I J + has-same-elements-eq-ideal-Semiring .I refl = + refl-has-same-elements-ideal-Semiring + + is-equiv-has-same-elements-eq-ideal-Semiring : + (J : ideal-Semiring l2 R) → is-equiv (has-same-elements-eq-ideal-Semiring J) + is-equiv-has-same-elements-eq-ideal-Semiring = + fundamental-theorem-id + is-torsorial-has-same-elements-ideal-Semiring + has-same-elements-eq-ideal-Semiring + + extensionality-ideal-Semiring : + (J : ideal-Semiring l2 R) → + (I = J) ≃ has-same-elements-ideal-Semiring R I J + pr1 (extensionality-ideal-Semiring J) = + has-same-elements-eq-ideal-Semiring J + pr2 (extensionality-ideal-Semiring J) = + is-equiv-has-same-elements-eq-ideal-Semiring J + + eq-has-same-elements-ideal-Semiring : + (J : ideal-Semiring l2 R) → has-same-elements-ideal-Semiring R I J → I = J + eq-has-same-elements-ideal-Semiring J = + map-inv-equiv (extensionality-ideal-Semiring J) +``` diff --git a/src/ring-theory/left-ideals-generated-by-subsets-rings.lagda.md b/src/ring-theory/left-ideals-generated-by-subsets-rings.lagda.md index 45001b74cb..2ee0d560d7 100644 --- a/src/ring-theory/left-ideals-generated-by-subsets-rings.lagda.md +++ b/src/ring-theory/left-ideals-generated-by-subsets-rings.lagda.md @@ -42,9 +42,9 @@ open import ring-theory.subsets-rings ## Idea -The **left ideal generated by a [subset](ring-theory.subsets-rings.md)** `S` of -a [ring](ring-theory.rings.md) `R` is the least left ideal in `R` containing -`S`. +The {{#concept "left ideal generated by a subset" Disambiguation="rings" Agda=left-ideal-subset-Ring}} `S` of +a [ring](ring-theory.rings.md) `R` is the least [left ideal](ring-theory.left-ideals-rings.md) in `R` containing +the [subset](ring-theory.subsets-rings.md) `S`. ## Definitions diff --git a/src/ring-theory/left-ideals-generated-by-subsets-semirings.lagda.md b/src/ring-theory/left-ideals-generated-by-subsets-semirings.lagda.md new file mode 100644 index 0000000000..0f69449590 --- /dev/null +++ b/src/ring-theory/left-ideals-generated-by-subsets-semirings.lagda.md @@ -0,0 +1,527 @@ +# Left ideals generated by subsets of semirings + +```agda +module ring-theory.left-ideals-generated-by-subsets-semirings where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.identity-types +open import foundation.logical-equivalences +open import foundation.powersets +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import lists.concatenation-lists +open import lists.functoriality-lists +open import lists.lists + +open import order-theory.galois-connections-large-posets +open import order-theory.least-upper-bounds-large-posets +open import order-theory.order-preserving-maps-large-posets +open import order-theory.order-preserving-maps-large-preorders +open import order-theory.reflective-galois-connections-large-posets + +open import ring-theory.left-ideals-semirings +open import ring-theory.left-linear-combinations-of-elements-semirings +open import ring-theory.monoids-with-semiring-actions +open import ring-theory.poset-of-left-ideals-semirings +open import ring-theory.semirings +open import ring-theory.subsets-semirings +``` + +
+ +## Idea + +The {{#concept "left ideal generated by a subset" Disambiguation="semirings" Agda=left-ideal-subset-Semiring}} `S` of +a [semiring](ring-theory.semirings.md) `R` is the least [left ideal](ring-theory.left-ideals-semirings.md) in `R` containing +the [subset](ring-theory.subsets-semirings.md) `S`. + +## Definitions + +### The universal property of left ideals generated by a subset of a semiring + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) (S : subset-Semiring l2 R) + (I : left-ideal-Semiring l3 R) (H : S ⊆ subset-left-ideal-Semiring R I) + where + + is-left-ideal-generated-by-subset-Semiring : UUω + is-left-ideal-generated-by-subset-Semiring = + {l : Level} (J : left-ideal-Semiring l R) → + S ⊆ subset-left-ideal-Semiring R J → + subset-left-ideal-Semiring R I ⊆ subset-left-ideal-Semiring R J +``` + +### The universal property of left ideals generated by a family of subsets of a semiring + +```agda +module _ + {l1 l2 l3 l4 : Level} (R : Semiring l1) {U : UU l2} + (S : U → subset-Semiring l3 R) + (I : left-ideal-Semiring l4 R) + (H : (α : U) → S α ⊆ subset-left-ideal-Semiring R I) + where + + is-left-ideal-generated-by-family-of-subsets-Semiring : UUω + is-left-ideal-generated-by-family-of-subsets-Semiring = + {l : Level} (J : left-ideal-Semiring l R) → + ((α : U) → S α ⊆ subset-left-ideal-Semiring R J) → + leq-left-ideal-Semiring R I J +``` + +### The universal property of left ideals generated by a family of elements in a semiring + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) {U : UU l2} (a : U → type-Semiring R) + (I : left-ideal-Semiring l3 R) (H : (α : U) → is-in-left-ideal-Semiring R I (a α)) + where + + is-left-ideal-generated-by-family-of-elements-Semiring : UUω + is-left-ideal-generated-by-family-of-elements-Semiring = + {l : Level} (J : left-ideal-Semiring l R) → + ((α : U) → is-in-left-ideal-Semiring R J (a α)) → + leq-left-ideal-Semiring R I J +``` + +### Construction of the Galois connection of left ideals generated by subsets + +#### Left ideals generated by subsets + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (S : subset-Semiring l2 R) + where + + subset-left-ideal-subset-Semiring' : type-Semiring R → UU (l1 ⊔ l2) + subset-left-ideal-subset-Semiring' = + is-left-linear-combination-subset-Semiring R S + + subset-left-ideal-subset-Semiring : subset-Semiring (l1 ⊔ l2) R + subset-left-ideal-subset-Semiring = + is-mere-left-linear-combination-prop-subset-Semiring R S + + is-in-left-ideal-subset-Semiring : type-Semiring R → UU (l1 ⊔ l2) + is-in-left-ideal-subset-Semiring = + is-in-subtype subset-left-ideal-subset-Semiring + + contains-zero-left-ideal-subset-Semiring : + contains-zero-subset-Semiring R subset-left-ideal-subset-Semiring + contains-zero-left-ideal-subset-Semiring = + unit-trunc-Prop (nil , refl) + + is-closed-under-addition-left-ideal-subset-Semiring : + is-closed-under-addition-subset-Semiring R subset-left-ideal-subset-Semiring + is-closed-under-addition-left-ideal-subset-Semiring x y H K = + apply-twice-universal-property-trunc-Prop H K + ( subset-left-ideal-subset-Semiring _) + ( λ H' K' → + unit-trunc-Prop + ( is-left-linear-combination-mul-monoid-Semiring R + ( additive-monoid-Semiring R) + ( λ r (s , p) → mul-Semiring R r s) + ( x) + ( y) + ( H') + ( K'))) + + is-closed-under-left-multiplication-left-ideal-subset-Semiring : + is-closed-under-left-multiplication-subset-Semiring R + ( subset-left-ideal-subset-Semiring) + is-closed-under-left-multiplication-left-ideal-subset-Semiring r x H = + apply-universal-property-trunc-Prop H + ( subset-left-ideal-subset-Semiring _) + ( λ H' → + unit-trunc-Prop + ( is-left-linear-combination-action-Semiring + ( R) + ( additive-monoid-with-left-semiring-action-Semiring R) + ( λ t (s , H) → mul-Semiring R t s) + ( λ s t (a , H) → associative-mul-Semiring R s t a) + ( r) + ( x) + ( H'))) + + left-ideal-subset-Semiring : + left-ideal-Semiring (l1 ⊔ l2) R + pr1 left-ideal-subset-Semiring = + subset-left-ideal-subset-Semiring + pr1 (pr1 (pr2 left-ideal-subset-Semiring)) = + contains-zero-left-ideal-subset-Semiring + pr2 (pr1 (pr2 left-ideal-subset-Semiring)) = + is-closed-under-addition-left-ideal-subset-Semiring + pr2 (pr2 left-ideal-subset-Semiring) = + is-closed-under-left-multiplication-left-ideal-subset-Semiring + + contains-subset-left-ideal-subset-Semiring : + S ⊆ subset-left-ideal-subset-Semiring + contains-subset-left-ideal-subset-Semiring s H = + unit-trunc-Prop + ( ( unit-list (one-Semiring R , (s , H))) , + ( right-unit-law-add-Semiring R (mul-Semiring R (one-Semiring R) s)) ∙ + ( left-unit-law-mul-Semiring R s)) + + contains-left-linear-combinations-left-ideal-subset-Semiring : + {l3 : Level} (I : left-ideal-Semiring l3 R) → + S ⊆ subset-left-ideal-Semiring R I → + (x : left-linear-combination-subset-Semiring R S) → + is-in-left-ideal-Semiring R I + ( ev-left-linear-combination-subset-Semiring R S x) + contains-left-linear-combinations-left-ideal-subset-Semiring I H nil = + contains-zero-left-ideal-Semiring R I + contains-left-linear-combinations-left-ideal-subset-Semiring I H + ( cons (r , (s , K)) l) = + is-closed-under-addition-left-ideal-Semiring R I + ( mul-Semiring R r s) + ( ev-left-linear-combination-subset-Semiring R S l) + ( is-closed-under-left-multiplication-left-ideal-Semiring R I r s (H s K)) + ( contains-left-linear-combinations-left-ideal-subset-Semiring I H l) + + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring : + is-left-ideal-generated-by-subset-Semiring R S + ( left-ideal-subset-Semiring) + ( contains-subset-left-ideal-subset-Semiring) + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring {l} J K x H = + apply-universal-property-trunc-Prop H (subset-left-ideal-Semiring R J x) P + where + P : subset-left-ideal-subset-Semiring' x → is-in-left-ideal-Semiring R J x + P (pair c refl) = + contains-left-linear-combinations-left-ideal-subset-Semiring J K c + + is-closed-under-eq-left-ideal-subset-Semiring : + {x y : type-Semiring R} → is-in-left-ideal-subset-Semiring x → + (x = y) → is-in-left-ideal-subset-Semiring y + is-closed-under-eq-left-ideal-subset-Semiring = + is-closed-under-eq-left-ideal-Semiring R left-ideal-subset-Semiring + + is-closed-under-eq-left-ideal-subset-Semiring' : + {x y : type-Semiring R} → is-in-left-ideal-subset-Semiring y → + (x = y) → is-in-left-ideal-subset-Semiring x + is-closed-under-eq-left-ideal-subset-Semiring' = + is-closed-under-eq-left-ideal-Semiring' R left-ideal-subset-Semiring +``` + +#### The subset relation is preserved by generating left ideals + +```agda +module _ + {l1 : Level} (A : Semiring l1) + where + + preserves-order-left-ideal-subset-Semiring : + {l2 l3 : Level} (S : subset-Semiring l2 A) (T : subset-Semiring l3 A) → + S ⊆ T → + leq-left-ideal-Semiring A + ( left-ideal-subset-Semiring A S) + ( left-ideal-subset-Semiring A T) + preserves-order-left-ideal-subset-Semiring S T H = + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring A S + ( left-ideal-subset-Semiring A T) + ( transitive-leq-subtype S T + ( subset-left-ideal-subset-Semiring A T) + ( contains-subset-left-ideal-subset-Semiring A T) + ( H)) + + left-ideal-subset-hom-large-poset-Semiring : + hom-Large-Poset + ( λ l2 → l1 ⊔ l2) + ( powerset-Large-Poset (type-Semiring A)) + ( left-ideal-Semiring-Large-Poset A) + map-hom-Large-Preorder left-ideal-subset-hom-large-poset-Semiring = + left-ideal-subset-Semiring A + preserves-order-hom-Large-Preorder + left-ideal-subset-hom-large-poset-Semiring = + preserves-order-left-ideal-subset-Semiring +``` + +#### The Galois connection `S ↦ (S)` + +```agda +module _ + {l1 : Level} (A : Semiring l1) + where + + adjoint-relation-left-ideal-subset-Semiring : + {l2 l3 : Level} (S : subset-Semiring l2 A) (I : left-ideal-Semiring l3 A) → + leq-left-ideal-Semiring A (left-ideal-subset-Semiring A S) I ↔ + (S ⊆ subset-left-ideal-Semiring A I) + pr1 (adjoint-relation-left-ideal-subset-Semiring S I) H = + transitive-leq-subtype S + ( subset-left-ideal-subset-Semiring A S) + ( subset-left-ideal-Semiring A I) + ( H) + ( contains-subset-left-ideal-subset-Semiring A S) + pr2 (adjoint-relation-left-ideal-subset-Semiring S I) = + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring A S I + + left-ideal-subset-galois-connection-Semiring : + galois-connection-Large-Poset + ( l1 ⊔_) (λ l → l) + ( powerset-Large-Poset (type-Semiring A)) + ( left-ideal-Semiring-Large-Poset A) + lower-adjoint-galois-connection-Large-Poset + left-ideal-subset-galois-connection-Semiring = + left-ideal-subset-hom-large-poset-Semiring A + upper-adjoint-galois-connection-Large-Poset + left-ideal-subset-galois-connection-Semiring = + subset-left-ideal-hom-large-poset-Semiring A + adjoint-relation-galois-connection-Large-Poset + left-ideal-subset-galois-connection-Semiring = + adjoint-relation-left-ideal-subset-Semiring +``` + +#### The reflective Galois connection `S ↦ (S)` + +```agda +module _ + {l1 : Level} (A : Semiring l1) + where + + forward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring : + {l2 : Level} (I : left-ideal-Semiring l2 A) → + leq-left-ideal-Semiring A + ( left-ideal-subset-Semiring A (subset-left-ideal-Semiring A I)) + ( I) + forward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring + I = + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring A + ( subset-left-ideal-Semiring A I) + ( I) + ( refl-leq-left-ideal-Semiring A I) + + backward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring : + {l2 : Level} (I : left-ideal-Semiring l2 A) → + leq-left-ideal-Semiring A I + ( left-ideal-subset-Semiring A (subset-left-ideal-Semiring A I)) + backward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring + I = + contains-subset-left-ideal-subset-Semiring A + ( subset-left-ideal-Semiring A I) + + is-reflective-left-ideal-subset-galois-connection-Semiring : + is-reflective-galois-connection-Large-Poset + ( powerset-Large-Poset (type-Semiring A)) + ( left-ideal-Semiring-Large-Poset A) + ( left-ideal-subset-galois-connection-Semiring A) + pr1 (is-reflective-left-ideal-subset-galois-connection-Semiring I) = + forward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring + I + pr2 (is-reflective-left-ideal-subset-galois-connection-Semiring I) = + backward-inclusion-is-reflective-left-ideal-subset-galois-connection-Semiring + I + + left-ideal-subset-reflective-galois-connection-Semiring : + reflective-galois-connection-Large-Poset + ( powerset-Large-Poset (type-Semiring A)) + ( left-ideal-Semiring-Large-Poset A) + galois-connection-reflective-galois-connection-Large-Poset + left-ideal-subset-reflective-galois-connection-Semiring = + left-ideal-subset-galois-connection-Semiring A + is-reflective-reflective-galois-connection-Large-Poset + left-ideal-subset-reflective-galois-connection-Semiring = + is-reflective-left-ideal-subset-galois-connection-Semiring +``` + +### The left ideal generated by a family of subsets of a semiring + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) {U : UU l2} (S : U → subset-Semiring l3 R) + where + + generating-subset-left-ideal-family-of-subsets-Semiring : + subset-Semiring (l2 ⊔ l3) R + generating-subset-left-ideal-family-of-subsets-Semiring x = + union-family-of-subtypes S x + + left-ideal-family-of-subsets-Semiring : left-ideal-Semiring (l1 ⊔ l2 ⊔ l3) R + left-ideal-family-of-subsets-Semiring = + left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-subsets-Semiring + + subset-left-ideal-family-of-subsets-Semiring : + subset-Semiring (l1 ⊔ l2 ⊔ l3) R + subset-left-ideal-family-of-subsets-Semiring = + subset-left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-subsets-Semiring + + is-in-left-ideal-family-of-subsets-Semiring : + type-Semiring R → UU (l1 ⊔ l2 ⊔ l3) + is-in-left-ideal-family-of-subsets-Semiring = + is-in-left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-subsets-Semiring + + contains-subset-left-ideal-family-of-subsets-Semiring : + {α : U} → (S α) ⊆ subset-left-ideal-family-of-subsets-Semiring + contains-subset-left-ideal-family-of-subsets-Semiring {α} x H = + contains-subset-left-ideal-subset-Semiring R + ( generating-subset-left-ideal-family-of-subsets-Semiring) + ( x) + ( unit-trunc-Prop (α , H)) + + is-left-ideal-generated-by-family-of-subsets-left-ideal-family-of-subsets-Semiring : + is-left-ideal-generated-by-family-of-subsets-Semiring R S + ( left-ideal-family-of-subsets-Semiring) + ( λ α → contains-subset-left-ideal-family-of-subsets-Semiring) + is-left-ideal-generated-by-family-of-subsets-left-ideal-family-of-subsets-Semiring + J H = + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring R + ( generating-subset-left-ideal-family-of-subsets-Semiring) + ( J) + ( λ y q → + apply-universal-property-trunc-Prop q + ( subset-left-ideal-Semiring R J y) + ( λ (α , K) → H α y K)) +``` + +### The left ideal generated by a family of elements in a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) {I : UU l1} (a : I → type-Semiring R) + where + + generating-subset-left-ideal-family-of-elements-Semiring : + subset-Semiring l1 R + generating-subset-left-ideal-family-of-elements-Semiring x = + trunc-Prop (fiber a x) + + left-ideal-family-of-elements-Semiring : + left-ideal-Semiring l1 R + left-ideal-family-of-elements-Semiring = + left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-elements-Semiring + + subset-left-ideal-family-of-elements-Semiring : + subset-Semiring l1 R + subset-left-ideal-family-of-elements-Semiring = + subset-left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-elements-Semiring + + is-in-left-ideal-family-of-elements-Semiring : + type-Semiring R → UU l1 + is-in-left-ideal-family-of-elements-Semiring = + is-in-left-ideal-subset-Semiring R + generating-subset-left-ideal-family-of-elements-Semiring + + contains-element-left-ideal-family-of-elements-Semiring : + (i : I) → is-in-left-ideal-family-of-elements-Semiring (a i) + contains-element-left-ideal-family-of-elements-Semiring i = + contains-subset-left-ideal-subset-Semiring R + ( generating-subset-left-ideal-family-of-elements-Semiring) + ( a i) + ( unit-trunc-Prop (i , refl)) + + abstract + is-left-ideal-generated-by-family-of-elements-left-ideal-family-of-elements-Semiring : + is-left-ideal-generated-by-family-of-elements-Semiring R a + left-ideal-family-of-elements-Semiring + contains-element-left-ideal-family-of-elements-Semiring + is-left-ideal-generated-by-family-of-elements-left-ideal-family-of-elements-Semiring + J H = + is-left-ideal-generated-by-subset-left-ideal-subset-Semiring R + ( generating-subset-left-ideal-family-of-elements-Semiring) + ( J) + ( λ x p → + apply-universal-property-trunc-Prop p + ( subset-left-ideal-Semiring R J x) + ( λ where (i , refl) → H i)) +``` + +## Properties + +### The left ideal generated by the underlying subset of an left ideal `I` is `I` itself + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (I : left-ideal-Semiring l2 R) + where + + cases-forward-inclusion-idempotent-left-ideal-subset-Semiring : + (l : left-linear-combination-subset-Semiring R (subset-left-ideal-Semiring R I)) → + is-in-left-ideal-Semiring R I + ( ev-left-linear-combination-subset-Semiring R (subset-left-ideal-Semiring R I) l) + cases-forward-inclusion-idempotent-left-ideal-subset-Semiring nil = + contains-zero-left-ideal-Semiring R I + cases-forward-inclusion-idempotent-left-ideal-subset-Semiring + ( cons (x , y , u) l) = + is-closed-under-addition-left-ideal-Semiring R I + ( mul-Semiring R x y) + ( ev-left-linear-combination-subset-Semiring R + ( subset-left-ideal-Semiring R I) + ( l)) + ( is-closed-under-left-multiplication-left-ideal-Semiring R I x y u) + ( cases-forward-inclusion-idempotent-left-ideal-subset-Semiring l) + + abstract + forward-inclusion-idempotent-left-ideal-subset-Semiring : + leq-left-ideal-Semiring R + ( left-ideal-subset-Semiring R (subset-left-ideal-Semiring R I)) + ( I) + forward-inclusion-idempotent-left-ideal-subset-Semiring x H = + apply-universal-property-trunc-Prop H + ( subset-left-ideal-Semiring R I x) + ( λ where + ( l , refl) → + cases-forward-inclusion-idempotent-left-ideal-subset-Semiring l) + + backward-inclusion-idempotent-left-ideal-subset-Semiring : + leq-left-ideal-Semiring R + ( I) + ( left-ideal-subset-Semiring R (subset-left-ideal-Semiring R I)) + backward-inclusion-idempotent-left-ideal-subset-Semiring = + contains-subset-left-ideal-subset-Semiring R + ( subset-left-ideal-Semiring R I) + + idempotent-left-ideal-subset-Semiring : + has-same-elements-left-ideal-Semiring R + ( left-ideal-subset-Semiring R (subset-left-ideal-Semiring R I)) + ( I) + pr1 (idempotent-left-ideal-subset-Semiring x) = + forward-inclusion-idempotent-left-ideal-subset-Semiring x + pr2 (idempotent-left-ideal-subset-Semiring x) = + backward-inclusion-idempotent-left-ideal-subset-Semiring x +``` + +### The operation `S ↦ (S)` preserves least upper bounds + +In +[`ring-theory.joins-left-ideals-semirings`](ring-theory.joins-left-ideals-semirings.md) +we will convert this fact to the fact that `S ↦ (S)` preserves joins. + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {U : UU l2} (S : U → subset-Semiring l3 R) + where + + preserves-least-upper-bounds-left-ideal-subset-Semiring : + {l4 : Level} (T : subset-Semiring l4 R) → + is-least-upper-bound-family-of-elements-Large-Poset + ( powerset-Large-Poset (type-Semiring R)) + ( S) + ( T) → + is-least-upper-bound-family-of-elements-Large-Poset + ( left-ideal-Semiring-Large-Poset R) + ( λ α → left-ideal-subset-Semiring R (S α)) + ( left-ideal-subset-Semiring R T) + preserves-least-upper-bounds-left-ideal-subset-Semiring = + preserves-join-lower-adjoint-galois-connection-Large-Poset + ( powerset-Large-Poset (type-Semiring R)) + ( left-ideal-Semiring-Large-Poset R) + ( left-ideal-subset-galois-connection-Semiring R) + ( S) +``` diff --git a/src/ring-theory/left-ideals-rings.lagda.md b/src/ring-theory/left-ideals-rings.lagda.md index c0d2710a89..b7668f4c21 100644 --- a/src/ring-theory/left-ideals-rings.lagda.md +++ b/src/ring-theory/left-ideals-rings.lagda.md @@ -18,6 +18,7 @@ open import foundation.subtypes open import foundation.torsorial-type-families open import foundation.universe-levels +open import ring-theory.left-ideals-semirings open import ring-theory.rings open import ring-theory.subsets-rings ``` @@ -26,7 +27,7 @@ open import ring-theory.subsets-rings ## Idea -A **left ideal** in a [ring](ring-theory.rings.md) `R` is a left submodule of +A {{#concept "left ideal" Disambiguation="rings" Agda=left-ideal-Ring}} in a [ring](ring-theory.rings.md) `R` is a left submodule of `R`. ## Definitions @@ -40,92 +41,99 @@ module _ is-left-ideal-subset-Ring : {l2 : Level} → subset-Ring l2 R → UU (l1 ⊔ l2) - is-left-ideal-subset-Ring S = - is-additive-subgroup-subset-Ring R S × - is-closed-under-left-multiplication-subset-Ring R S + is-left-ideal-subset-Ring = + is-left-ideal-subset-Semiring (semiring-Ring R) is-prop-is-left-ideal-subset-Ring : {l2 : Level} (S : subset-Ring l2 R) → is-prop (is-left-ideal-subset-Ring S) - is-prop-is-left-ideal-subset-Ring S = - is-prop-product - ( is-prop-is-additive-subgroup-subset-Ring R S) - ( is-prop-is-closed-under-left-multiplication-subset-Ring R S) + is-prop-is-left-ideal-subset-Ring = + is-prop-is-left-ideal-subset-Semiring (semiring-Ring R) left-ideal-Ring : (l : Level) {l1 : Level} (R : Ring l1) → UU (lsuc l ⊔ l1) -left-ideal-Ring l R = Σ (subset-Ring l R) (is-left-ideal-subset-Ring R) +left-ideal-Ring l R = + left-ideal-Semiring l (semiring-Ring R) module _ {l1 l2 : Level} (R : Ring l1) (I : left-ideal-Ring l2 R) where subset-left-ideal-Ring : subset-Ring l2 R - subset-left-ideal-Ring = pr1 I + subset-left-ideal-Ring = + subset-left-ideal-Semiring (semiring-Ring R) I is-in-left-ideal-Ring : type-Ring R → UU l2 - is-in-left-ideal-Ring x = type-Prop (subset-left-ideal-Ring x) + is-in-left-ideal-Ring = + is-in-left-ideal-Semiring (semiring-Ring R) I type-left-ideal-Ring : UU (l1 ⊔ l2) type-left-ideal-Ring = type-subset-Ring R subset-left-ideal-Ring inclusion-left-ideal-Ring : type-left-ideal-Ring → type-Ring R - inclusion-left-ideal-Ring = inclusion-subset-Ring R subset-left-ideal-Ring + inclusion-left-ideal-Ring = + inclusion-left-ideal-Semiring (semiring-Ring R) I ap-inclusion-left-ideal-Ring : (x y : type-left-ideal-Ring) → x = y → inclusion-left-ideal-Ring x = inclusion-left-ideal-Ring y ap-inclusion-left-ideal-Ring = - ap-inclusion-subset-Ring R subset-left-ideal-Ring + ap-inclusion-left-ideal-Semiring (semiring-Ring R) I is-in-subset-inclusion-left-ideal-Ring : (x : type-left-ideal-Ring) → is-in-left-ideal-Ring (inclusion-left-ideal-Ring x) is-in-subset-inclusion-left-ideal-Ring = - is-in-subset-inclusion-subset-Ring R subset-left-ideal-Ring + is-in-subset-inclusion-left-ideal-Semiring (semiring-Ring R) I is-closed-under-eq-left-ideal-Ring : {x y : type-Ring R} → is-in-left-ideal-Ring x → (x = y) → is-in-left-ideal-Ring y is-closed-under-eq-left-ideal-Ring = - is-closed-under-eq-subset-Ring R subset-left-ideal-Ring + is-closed-under-eq-left-ideal-Semiring (semiring-Ring R) I is-closed-under-eq-left-ideal-Ring' : {x y : type-Ring R} → is-in-left-ideal-Ring y → (x = y) → is-in-left-ideal-Ring x is-closed-under-eq-left-ideal-Ring' = - is-closed-under-eq-subset-Ring' R subset-left-ideal-Ring + is-closed-under-eq-left-ideal-Semiring' (semiring-Ring R) I is-left-ideal-subset-left-ideal-Ring : is-left-ideal-subset-Ring R subset-left-ideal-Ring - is-left-ideal-subset-left-ideal-Ring = pr2 I + is-left-ideal-subset-left-ideal-Ring = + is-left-ideal-subset-left-ideal-Semiring (semiring-Ring R) I - is-additive-subgroup-subset-left-ideal-Ring : - is-additive-subgroup-subset-Ring R subset-left-ideal-Ring - is-additive-subgroup-subset-left-ideal-Ring = - pr1 is-left-ideal-subset-left-ideal-Ring + is-additive-submonoid-subset-left-ideal-Ring : + is-additive-submonoid-subset-Ring R subset-left-ideal-Ring + is-additive-submonoid-subset-left-ideal-Ring = + is-additive-submonoid-subset-left-ideal-Semiring (semiring-Ring R) I contains-zero-left-ideal-Ring : is-in-left-ideal-Ring (zero-Ring R) contains-zero-left-ideal-Ring = - pr1 is-additive-subgroup-subset-left-ideal-Ring + contains-zero-left-ideal-Semiring (semiring-Ring R) I is-closed-under-addition-left-ideal-Ring : is-closed-under-addition-subset-Ring R subset-left-ideal-Ring is-closed-under-addition-left-ideal-Ring = - pr1 (pr2 is-additive-subgroup-subset-left-ideal-Ring) - - is-closed-under-negatives-left-ideal-Ring : - is-closed-under-negatives-subset-Ring R subset-left-ideal-Ring - is-closed-under-negatives-left-ideal-Ring = - pr2 (pr2 is-additive-subgroup-subset-left-ideal-Ring) + is-closed-under-addition-left-ideal-Semiring (semiring-Ring R) I is-closed-under-left-multiplication-left-ideal-Ring : is-closed-under-left-multiplication-subset-Ring R subset-left-ideal-Ring is-closed-under-left-multiplication-left-ideal-Ring = - pr2 is-left-ideal-subset-left-ideal-Ring + is-closed-under-left-multiplication-left-ideal-Semiring (semiring-Ring R) I + + is-additive-subgroup-left-ideal-Ring : + is-additive-subgroup-subset-Ring R subset-left-ideal-Ring + is-additive-subgroup-left-ideal-Ring = + is-additive-subgroup-is-closed-under-left-multiplication-subset-Ring + ( R) + ( subset-left-ideal-Ring) + ( is-additive-submonoid-subset-left-ideal-Ring) + ( is-closed-under-left-multiplication-left-ideal-Ring) is-left-ideal-left-ideal-Ring : is-left-ideal-subset-Ring R subset-left-ideal-Ring - is-left-ideal-left-ideal-Ring = pr2 I + is-left-ideal-left-ideal-Ring = + is-left-ideal-left-ideal-Semiring (semiring-Ring R) I ``` ## Properties @@ -139,10 +147,8 @@ module _ has-same-elements-left-ideal-Ring : (J : left-ideal-Ring l3 R) → UU (l1 ⊔ l2 ⊔ l3) - has-same-elements-left-ideal-Ring J = - has-same-elements-subtype - ( subset-left-ideal-Ring R I) - ( subset-left-ideal-Ring R J) + has-same-elements-left-ideal-Ring = + has-same-elements-left-ideal-Semiring (semiring-Ring R) I module _ {l1 l2 : Level} (R : Ring l1) (I : left-ideal-Ring l2 R) @@ -151,43 +157,34 @@ module _ refl-has-same-elements-left-ideal-Ring : has-same-elements-left-ideal-Ring R I I refl-has-same-elements-left-ideal-Ring = - refl-has-same-elements-subtype (subset-left-ideal-Ring R I) + refl-has-same-elements-left-ideal-Semiring (semiring-Ring R) I is-torsorial-has-same-elements-left-ideal-Ring : is-torsorial (has-same-elements-left-ideal-Ring R I) is-torsorial-has-same-elements-left-ideal-Ring = - is-torsorial-Eq-subtype - ( is-torsorial-has-same-elements-subtype (subset-left-ideal-Ring R I)) - ( is-prop-is-left-ideal-subset-Ring R) - ( subset-left-ideal-Ring R I) - ( refl-has-same-elements-left-ideal-Ring) - ( is-left-ideal-left-ideal-Ring R I) + is-torsorial-has-same-elements-left-ideal-Semiring (semiring-Ring R) I has-same-elements-eq-left-ideal-Ring : (J : left-ideal-Ring l2 R) → (I = J) → has-same-elements-left-ideal-Ring R I J - has-same-elements-eq-left-ideal-Ring .I refl = - refl-has-same-elements-left-ideal-Ring + has-same-elements-eq-left-ideal-Ring = + has-same-elements-eq-left-ideal-Semiring (semiring-Ring R) I is-equiv-has-same-elements-eq-left-ideal-Ring : (J : left-ideal-Ring l2 R) → is-equiv (has-same-elements-eq-left-ideal-Ring J) is-equiv-has-same-elements-eq-left-ideal-Ring = - fundamental-theorem-id - is-torsorial-has-same-elements-left-ideal-Ring - has-same-elements-eq-left-ideal-Ring + is-equiv-has-same-elements-eq-left-ideal-Semiring (semiring-Ring R) I extensionality-left-ideal-Ring : (J : left-ideal-Ring l2 R) → (I = J) ≃ has-same-elements-left-ideal-Ring R I J - pr1 (extensionality-left-ideal-Ring J) = - has-same-elements-eq-left-ideal-Ring J - pr2 (extensionality-left-ideal-Ring J) = - is-equiv-has-same-elements-eq-left-ideal-Ring J + extensionality-left-ideal-Ring = + extensionality-left-ideal-Semiring (semiring-Ring R) I eq-has-same-elements-left-ideal-Ring : (J : left-ideal-Ring l2 R) → has-same-elements-left-ideal-Ring R I J → I = J - eq-has-same-elements-left-ideal-Ring J = - map-inv-equiv (extensionality-left-ideal-Ring J) + eq-has-same-elements-left-ideal-Ring = + eq-has-same-elements-left-ideal-Semiring (semiring-Ring R) I ``` diff --git a/src/ring-theory/left-ideals-semirings.lagda.md b/src/ring-theory/left-ideals-semirings.lagda.md new file mode 100644 index 0000000000..fd5167ec4d --- /dev/null +++ b/src/ring-theory/left-ideals-semirings.lagda.md @@ -0,0 +1,194 @@ +# Left ideals of semirings + +```agda +module ring-theory.left-ideals-semirings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fundamental-theorem-of-identity-types +open import foundation.identity-types +open import foundation.propositions +open import foundation.subtype-identity-principle +open import foundation.subtypes +open import foundation.torsorial-type-families +open import foundation.universe-levels + +open import ring-theory.semirings +open import ring-theory.subsets-semirings +``` + +
+ +## Idea + +A {{#concept "left ideal" Disambiguation="semiring" Agda=left-ideal-Semiring}} in a [semiring](ring-theory.semirings.md) `R` is a left submodule of +`R`. + +## Definitions + +### Left ideals + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + is-left-ideal-subset-Semiring : + {l2 : Level} → subset-Semiring l2 R → UU (l1 ⊔ l2) + is-left-ideal-subset-Semiring S = + is-additive-submonoid-subset-Semiring R S × + is-closed-under-left-multiplication-subset-Semiring R S + + is-prop-is-left-ideal-subset-Semiring : + {l2 : Level} (S : subset-Semiring l2 R) → + is-prop (is-left-ideal-subset-Semiring S) + is-prop-is-left-ideal-subset-Semiring S = + is-prop-product + ( is-prop-is-additive-submonoid-subset-Semiring R S) + ( is-prop-is-closed-under-left-multiplication-subset-Semiring R S) + +left-ideal-Semiring : + (l : Level) {l1 : Level} (R : Semiring l1) → UU (lsuc l ⊔ l1) +left-ideal-Semiring l R = + Σ (subset-Semiring l R) (is-left-ideal-subset-Semiring R) + +module _ + {l1 l2 : Level} (R : Semiring l1) (I : left-ideal-Semiring l2 R) + where + + subset-left-ideal-Semiring : subset-Semiring l2 R + subset-left-ideal-Semiring = pr1 I + + is-in-left-ideal-Semiring : type-Semiring R → UU l2 + is-in-left-ideal-Semiring x = type-Prop (subset-left-ideal-Semiring x) + + type-left-ideal-Semiring : UU (l1 ⊔ l2) + type-left-ideal-Semiring = type-subset-Semiring R subset-left-ideal-Semiring + + inclusion-left-ideal-Semiring : + type-left-ideal-Semiring → type-Semiring R + inclusion-left-ideal-Semiring = + inclusion-subset-Semiring R subset-left-ideal-Semiring + + ap-inclusion-left-ideal-Semiring : + (x y : type-left-ideal-Semiring) → x = y → + inclusion-left-ideal-Semiring x = inclusion-left-ideal-Semiring y + ap-inclusion-left-ideal-Semiring = + ap-inclusion-subset-Semiring R subset-left-ideal-Semiring + + is-in-subset-inclusion-left-ideal-Semiring : + (x : type-left-ideal-Semiring) → + is-in-left-ideal-Semiring (inclusion-left-ideal-Semiring x) + is-in-subset-inclusion-left-ideal-Semiring = + is-in-subset-inclusion-subset-Semiring R subset-left-ideal-Semiring + + is-closed-under-eq-left-ideal-Semiring : + {x y : type-Semiring R} → is-in-left-ideal-Semiring x → + (x = y) → is-in-left-ideal-Semiring y + is-closed-under-eq-left-ideal-Semiring = + is-closed-under-eq-subset-Semiring R subset-left-ideal-Semiring + + is-closed-under-eq-left-ideal-Semiring' : + {x y : type-Semiring R} → is-in-left-ideal-Semiring y → + (x = y) → is-in-left-ideal-Semiring x + is-closed-under-eq-left-ideal-Semiring' = + is-closed-under-eq-subset-Semiring' R subset-left-ideal-Semiring + + is-left-ideal-subset-left-ideal-Semiring : + is-left-ideal-subset-Semiring R subset-left-ideal-Semiring + is-left-ideal-subset-left-ideal-Semiring = pr2 I + + is-additive-submonoid-subset-left-ideal-Semiring : + is-additive-submonoid-subset-Semiring R subset-left-ideal-Semiring + is-additive-submonoid-subset-left-ideal-Semiring = + pr1 is-left-ideal-subset-left-ideal-Semiring + + contains-zero-left-ideal-Semiring : + is-in-left-ideal-Semiring (zero-Semiring R) + contains-zero-left-ideal-Semiring = + pr1 is-additive-submonoid-subset-left-ideal-Semiring + + is-closed-under-addition-left-ideal-Semiring : + is-closed-under-addition-subset-Semiring R subset-left-ideal-Semiring + is-closed-under-addition-left-ideal-Semiring = + pr2 is-additive-submonoid-subset-left-ideal-Semiring + + is-closed-under-left-multiplication-left-ideal-Semiring : + is-closed-under-left-multiplication-subset-Semiring R + ( subset-left-ideal-Semiring) + is-closed-under-left-multiplication-left-ideal-Semiring = + pr2 is-left-ideal-subset-left-ideal-Semiring + + is-left-ideal-left-ideal-Semiring : + is-left-ideal-subset-Semiring R subset-left-ideal-Semiring + is-left-ideal-left-ideal-Semiring = pr2 I +``` + +## Properties + +### Characterizing equality of left ideals in semirings + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) (I : left-ideal-Semiring l2 R) + where + + has-same-elements-left-ideal-Semiring : + (J : left-ideal-Semiring l3 R) → UU (l1 ⊔ l2 ⊔ l3) + has-same-elements-left-ideal-Semiring J = + has-same-elements-subtype + ( subset-left-ideal-Semiring R I) + ( subset-left-ideal-Semiring R J) + +module _ + {l1 l2 : Level} (R : Semiring l1) (I : left-ideal-Semiring l2 R) + where + + refl-has-same-elements-left-ideal-Semiring : + has-same-elements-left-ideal-Semiring R I I + refl-has-same-elements-left-ideal-Semiring = + refl-has-same-elements-subtype (subset-left-ideal-Semiring R I) + + is-torsorial-has-same-elements-left-ideal-Semiring : + is-torsorial (has-same-elements-left-ideal-Semiring R I) + is-torsorial-has-same-elements-left-ideal-Semiring = + is-torsorial-Eq-subtype + ( is-torsorial-has-same-elements-subtype (subset-left-ideal-Semiring R I)) + ( is-prop-is-left-ideal-subset-Semiring R) + ( subset-left-ideal-Semiring R I) + ( refl-has-same-elements-left-ideal-Semiring) + ( is-left-ideal-left-ideal-Semiring R I) + + has-same-elements-eq-left-ideal-Semiring : + (J : left-ideal-Semiring l2 R) → + (I = J) → has-same-elements-left-ideal-Semiring R I J + has-same-elements-eq-left-ideal-Semiring .I refl = + refl-has-same-elements-left-ideal-Semiring + + is-equiv-has-same-elements-eq-left-ideal-Semiring : + (J : left-ideal-Semiring l2 R) → + is-equiv (has-same-elements-eq-left-ideal-Semiring J) + is-equiv-has-same-elements-eq-left-ideal-Semiring = + fundamental-theorem-id + is-torsorial-has-same-elements-left-ideal-Semiring + has-same-elements-eq-left-ideal-Semiring + + extensionality-left-ideal-Semiring : + (J : left-ideal-Semiring l2 R) → + (I = J) ≃ has-same-elements-left-ideal-Semiring R I J + pr1 (extensionality-left-ideal-Semiring J) = + has-same-elements-eq-left-ideal-Semiring J + pr2 (extensionality-left-ideal-Semiring J) = + is-equiv-has-same-elements-eq-left-ideal-Semiring J + + eq-has-same-elements-left-ideal-Semiring : + (J : left-ideal-Semiring l2 R) → + has-same-elements-left-ideal-Semiring R I J → I = J + eq-has-same-elements-left-ideal-Semiring J = + map-inv-equiv (extensionality-left-ideal-Semiring J) +``` diff --git a/src/ring-theory/left-linear-combinations-of-elements-rings.lagda.md b/src/ring-theory/left-linear-combinations-of-elements-rings.lagda.md new file mode 100644 index 0000000000..7c27f978b9 --- /dev/null +++ b/src/ring-theory/left-linear-combinations-of-elements-rings.lagda.md @@ -0,0 +1,164 @@ +# Left linear combinations with respect to rings + +```agda +module ring-theory.left-linear-combinations-of-elements-rings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.universe-levels + +open import ring-theory.rings +open import ring-theory.subsets-rings + +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [ring](ring-theory.rings.md) $R$ and a type $A$. A {#concept "left linear combination"}} of elements of $R$ is a [list](lists.lists.md) of pairs $(r,a)$ consisting of an element $r:R$ and an element $a:A$. + +Furthermore, if we are given an action $\mu : R \to A \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a left linear combination $((r_0,a_0),\ldots,(r_{n-1},a_{n-1}))$ by defining + +$$ + ev((r_0,a_0),\ldots,(r_{n-1},a_{n-1})) := \sum_{i=0}^{n-1} \mu(r_i,a_i). +$$ + +To be explicit, left linear combinations of elements of a type $A$ have the ring coefficients on the left. + + +## Definitions + +### The type of left linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (A : UU l2) + where + + left-linear-combination-Ring : + UU (l1 ⊔ l2) + left-linear-combination-Ring = + list (type-Ring R × A) +``` + +### Evaluating left linear combinations of elements in unital magmas + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Unital-Magma M) + where + + ev-left-linear-combination-Ring : + left-linear-combination-Ring R A → type-Unital-Magma M + ev-left-linear-combination-Ring nil = + unit-Unital-Magma M + ev-left-linear-combination-Ring (cons (r , a) l) = + mul-Unital-Magma M (μ r a) (ev-left-linear-combination-Ring l) +``` + +### The predicate of being a left linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Unital-Magma M) + where + + is-left-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-left-linear-combination-Ring = + fiber (ev-left-linear-combination-Ring R M μ) +``` + +### The predicate of being a mere left linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Unital-Magma M) + where + + is-mere-left-linear-combination-prop-Ring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-left-linear-combination-prop-Ring x = + trunc-Prop (is-left-linear-combination-Ring R M μ x) + + is-mere-left-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-left-linear-combination-Ring x = + type-trunc-Prop (is-left-linear-combination-Ring R M μ x) + + is-prop-is-mere-left-linear-combination-Ring : + (x : type-Unital-Magma M) → + is-prop (is-mere-left-linear-combination-Ring x) + is-prop-is-mere-left-linear-combination-Ring x = + is-prop-type-trunc-Prop +``` + +### Left linear combinations of subsets of a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) + where + + left-linear-combination-subset-Ring : + UU (l1 ⊔ l2) + left-linear-combination-subset-Ring = + left-linear-combination-Ring R (type-subset-Ring R S) + + ev-left-linear-combination-subset-Ring : + left-linear-combination-subset-Ring → type-Ring R + ev-left-linear-combination-subset-Ring = + ev-left-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ r s → mul-Ring R r (inclusion-subset-Ring R S s)) + + is-left-linear-combination-subset-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-left-linear-combination-subset-Ring = + is-left-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ r s → mul-Ring R r (inclusion-subset-Ring R S s)) +``` + +### Left linear combinations of families of elements in a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) {I : UU l2} (a : I → type-Ring R) + where + + left-linear-combination-family-of-elements-Ring : + UU (l1 ⊔ l2) + left-linear-combination-family-of-elements-Ring = + left-linear-combination-subset-Ring R (trunc-Prop ∘ fiber a) + + ev-left-linear-combination-family-of-elements-Ring : + left-linear-combination-family-of-elements-Ring → type-Ring R + ev-left-linear-combination-family-of-elements-Ring = + ev-left-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) + + is-left-linear-combination-family-of-elements-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-left-linear-combination-family-of-elements-Ring = + is-left-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) +``` diff --git a/src/ring-theory/left-linear-combinations-of-elements-semirings.lagda.md b/src/ring-theory/left-linear-combinations-of-elements-semirings.lagda.md new file mode 100644 index 0000000000..2bb7561a9c --- /dev/null +++ b/src/ring-theory/left-linear-combinations-of-elements-semirings.lagda.md @@ -0,0 +1,306 @@ +# Left linear combinations with respect to semirings + +```agda +module ring-theory.left-linear-combinations-of-elements-semirings where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-binary-functions +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.identity-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.universe-levels + +open import group-theory.monoids + +open import ring-theory.monoids-with-semiring-actions +open import ring-theory.semirings +open import ring-theory.subsets-semirings + +open import lists.concatenation-lists +open import lists.functoriality-lists +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [semiring](ring-theory.semirings.md) $R$ and a type $A$. A {#concept "left linear combination"}} of elements of $R$ is a [list](lists.lists.md) of pairs $(r,a)$ consisting of an element $r:R$ and an element $a:A$. + +Furthermore, if we are given an action $\mu : R \to A \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a left linear combination $((r_0,a_0),\ldots,(r_{n-1},a_{n-1}))$ by defining + +$$ + ev((r_0,a_0),\ldots,(r_{n-1},a_{n-1})) := \sum_{i=0}^{n-1} \mu(r_i,a_i). +$$ + +To be explicit, left linear combinations of elements of a type $A$ have the semiring coefficients on the left. + + +## Definitions + +### The type of left linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (A : UU l2) + where + + left-linear-combination-Semiring : + UU (l1 ⊔ l2) + left-linear-combination-Semiring = + list (type-Semiring R × A) +``` + +### Multiplying linear combinations by a scalar + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) {A : UU l2} + where + + mul-left-linear-combination-Semiring : + type-Semiring R → + left-linear-combination-Semiring R A → + left-linear-combination-Semiring R A + mul-left-linear-combination-Semiring r = + map-list (λ (s , a) → (mul-Semiring R r s , a)) +``` + +### Evaluating left linear combinations of elements + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} + where + + ev-unital-magma-left-linear-combination-Semiring : + (M : Unital-Magma l3) (μ : type-Semiring R → A → type-Unital-Magma M) → + left-linear-combination-Semiring R A → type-Unital-Magma M + ev-unital-magma-left-linear-combination-Semiring M μ nil = + unit-Unital-Magma M + ev-unital-magma-left-linear-combination-Semiring M μ (cons (r , a) l) = + mul-Unital-Magma M + ( μ r a) + ( ev-unital-magma-left-linear-combination-Semiring M μ l) + + ev-monoid-left-linear-combination-Semiring : + (M : Monoid l3) (μ : type-Semiring R → A → type-Monoid M) → + left-linear-combination-Semiring R A → type-Monoid M + ev-monoid-left-linear-combination-Semiring M = + ev-unital-magma-left-linear-combination-Semiring (unital-magma-Monoid M) +``` + +### The predicate of being a left linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} + where + + is-left-linear-combination-Semiring : + (M : Unital-Magma l3) (μ : type-Semiring R → A → type-Unital-Magma M) → + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-left-linear-combination-Semiring M μ = + fiber (ev-unital-magma-left-linear-combination-Semiring R M μ) + + is-left-linear-combination-monoid-Semiring : + (M : Monoid l3) (μ : type-Semiring R → A → type-Monoid M) → + type-Monoid M → UU (l1 ⊔ l2 ⊔ l3) + is-left-linear-combination-monoid-Semiring M = + is-left-linear-combination-Semiring (unital-magma-Monoid M) +``` + +### The predicate of being a mere left linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Semiring R → A → type-Unital-Magma M) + where + + is-mere-left-linear-combination-prop-Semiring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-left-linear-combination-prop-Semiring x = + trunc-Prop (is-left-linear-combination-Semiring R M μ x) + + is-mere-left-linear-combination-Semiring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-left-linear-combination-Semiring x = + type-trunc-Prop (is-left-linear-combination-Semiring R M μ x) + + is-prop-is-mere-left-linear-combination-Semiring : + (x : type-Unital-Magma M) → + is-prop (is-mere-left-linear-combination-Semiring x) + is-prop-is-mere-left-linear-combination-Semiring x = + is-prop-type-trunc-Prop +``` + +### Left linear combinations of subsets of a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (S : subset-Semiring l2 R) + where + + left-linear-combination-subset-Semiring : + UU (l1 ⊔ l2) + left-linear-combination-subset-Semiring = + left-linear-combination-Semiring R (type-subset-Semiring R S) + + ev-left-linear-combination-subset-Semiring : + left-linear-combination-subset-Semiring → type-Semiring R + ev-left-linear-combination-subset-Semiring = + ev-unital-magma-left-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ r s → mul-Semiring R r (inclusion-subset-Semiring R S s)) + + is-left-linear-combination-subset-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-left-linear-combination-subset-Semiring = + is-left-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ r s → mul-Semiring R r (inclusion-subset-Semiring R S s)) + + is-mere-left-linear-combination-prop-subset-Semiring : + type-Semiring R → Prop (l1 ⊔ l2) + is-mere-left-linear-combination-prop-subset-Semiring = + is-mere-left-linear-combination-prop-Semiring R + ( additive-unital-magma-Semiring R) + ( λ r s → mul-Semiring R r (inclusion-subset-Semiring R S s)) +``` + +### Left linear combinations of families of elements in a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) {I : UU l2} (a : I → type-Semiring R) + where + + left-linear-combination-family-of-elements-Semiring : + UU (l1 ⊔ l2) + left-linear-combination-family-of-elements-Semiring = + left-linear-combination-subset-Semiring R (trunc-Prop ∘ fiber a) + + ev-unital-magma-left-linear-combination-family-of-elements-Semiring : + left-linear-combination-family-of-elements-Semiring → type-Semiring R + ev-unital-magma-left-linear-combination-family-of-elements-Semiring = + ev-left-linear-combination-subset-Semiring R (trunc-Prop ∘ fiber a) + + is-left-linear-combination-family-of-elements-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-left-linear-combination-family-of-elements-Semiring = + is-left-linear-combination-subset-Semiring R (trunc-Prop ∘ fiber a) +``` + +## Properties + +### Given a left action of a semiring $R$ on a type $A$ with values in a monoid, the evaluation function preserves concatenation + +We assume a monoid here, because we need associativity for the multiplicative operation of the monoid. + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Monoid l3) + (μ : type-Semiring R → A → type-Monoid M) + where + + preserves-concat-ev-monoid-left-linear-combination-Semiring : + (u v : left-linear-combination-Semiring R A) → + ev-monoid-left-linear-combination-Semiring R M + ( μ) + ( concat-list u v) = + mul-Monoid M + ( ev-monoid-left-linear-combination-Semiring R M μ u) + ( ev-monoid-left-linear-combination-Semiring R M μ v) + preserves-concat-ev-monoid-left-linear-combination-Semiring nil v = + inv + ( left-unit-law-mul-Monoid M + ( ev-monoid-left-linear-combination-Semiring R M μ v)) + preserves-concat-ev-monoid-left-linear-combination-Semiring + ( cons (r , s) u) v = + ( ap + ( mul-Monoid M (μ r s)) + ( preserves-concat-ev-monoid-left-linear-combination-Semiring u v)) ∙ + ( inv + ( associative-mul-Monoid M + ( μ r s) + ( ev-monoid-left-linear-combination-Semiring R M μ u) + ( ev-monoid-left-linear-combination-Semiring R M μ v))) + + is-left-linear-combination-mul-monoid-Semiring : + (x y : type-Monoid M) → + is-left-linear-combination-monoid-Semiring R M μ x → + is-left-linear-combination-monoid-Semiring R M μ y → + is-left-linear-combination-monoid-Semiring R M μ (mul-Monoid M x y) + is-left-linear-combination-mul-monoid-Semiring x y (u , refl) (v , refl) = + ( concat-list u v , + preserves-concat-ev-monoid-left-linear-combination-Semiring u v) +``` + +### Evaluation of linear combinations preserves scalar multiplication + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Monoid-With-Left-Semiring-Action l3 R) + (μ : type-Semiring R → A → type-Monoid-With-Left-Semiring-Action R M) + (α : + (s r : type-Semiring R) (a : A) → + μ (mul-Semiring R s r) a = + action-Monoid-With-Left-Semiring-Action R M s (μ r a)) + where + + preserves-mul-ev-left-linear-combination-Semiring : + (r : type-Semiring R) (u : left-linear-combination-Semiring R A) → + ev-monoid-left-linear-combination-Semiring R + ( monoid-Monoid-With-Left-Semiring-Action R M) + ( μ) + ( mul-left-linear-combination-Semiring R r u) = + action-Monoid-With-Left-Semiring-Action R M r + ( ev-monoid-left-linear-combination-Semiring R + ( monoid-Monoid-With-Left-Semiring-Action R M) + ( μ) + ( u)) + preserves-mul-ev-left-linear-combination-Semiring r nil = + inv (right-absorption-law-action-Monoid-With-Left-Semiring-Action R M r) + preserves-mul-ev-left-linear-combination-Semiring r (cons (s , a) u) = + ap-binary + ( mul-Monoid-With-Left-Semiring-Action R M) + ( α r s a) + ( preserves-mul-ev-left-linear-combination-Semiring r u) ∙ + inv + ( left-distributive-action-mul-Monoid-With-Left-Semiring-Action R M r + ( μ s a) + ( ev-monoid-left-linear-combination-Semiring R + ( monoid-Monoid-With-Left-Semiring-Action R M) + ( μ) + ( u))) + + is-left-linear-combination-action-Semiring : + (r : type-Semiring R) (x : type-Monoid-With-Left-Semiring-Action R M) → + is-left-linear-combination-monoid-Semiring R + ( monoid-Monoid-With-Left-Semiring-Action R M) + ( μ) + ( x) → + is-left-linear-combination-monoid-Semiring R + ( monoid-Monoid-With-Left-Semiring-Action R M) + ( μ) + ( action-Monoid-With-Left-Semiring-Action R M r x) + is-left-linear-combination-action-Semiring r x (u , refl) = + ( mul-left-linear-combination-Semiring R r u , + preserves-mul-ev-left-linear-combination-Semiring r u) +``` diff --git a/src/ring-theory/linear-combinations-of-elements-rings.lagda.md b/src/ring-theory/linear-combinations-of-elements-rings.lagda.md new file mode 100644 index 0000000000..dd8dce8017 --- /dev/null +++ b/src/ring-theory/linear-combinations-of-elements-rings.lagda.md @@ -0,0 +1,200 @@ +# Linear combinations with respect to rings + +```agda +module ring-theory.linear-combinations-of-elements-rings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.universe-levels + +open import ring-theory.rings +open import ring-theory.subsets-rings + +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [ring](ring-theory.rings.md) $R$ and a type $A$. A {#concept "linear combination"}} of elements of $A$ is a [list](lists.lists.md) of pairs $(r,a,s)$ consisting of an element $r,s:R$ and an element $a:A$. + +Furthermore, if we are given an action $\mu : R \to A \to R \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a linear combination $((r_0,a_0,s_0),\ldots,(r_{n-1},a_{n-1},s_{n-1}))$ by defining + +$$ + ev((r_0,a_0,s_0),\ldots,(r_{n-1},a_{n-1},s_{n-1})) := + \sum_{i=0}^{n-1} \mu(r_i,a_i,s_i). +$$ + +To be explicit, linear combinations of elements of a type $A$ have the ring coefficients on both sides. + +## Definitions + +### The type of linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (A : UU l2) + where + + linear-combination-Ring : + UU (l1 ⊔ l2) + linear-combination-Ring = + list (type-Ring R × A × type-Ring R) +``` + +### Evaluating linear combinations of elements in unital magmas + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Ring R → type-Unital-Magma M) + where + + ev-linear-combination-Ring : + linear-combination-Ring R A → type-Unital-Magma M + ev-linear-combination-Ring nil = + unit-Unital-Magma M + ev-linear-combination-Ring (cons (r , a , s) l) = + mul-Unital-Magma M (μ r a s) (ev-linear-combination-Ring l) +``` + +### The predicate of being a linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Ring R → type-Unital-Magma M) + where + + is-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-linear-combination-Ring = + fiber (ev-linear-combination-Ring R M μ) +``` + +### The predicate of being a mere linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Ring R → A → type-Ring R → type-Unital-Magma M) + where + + is-mere-linear-combination-prop-Ring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-linear-combination-prop-Ring x = + trunc-Prop (is-linear-combination-Ring R M μ x) + + is-mere-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-linear-combination-Ring x = + type-trunc-Prop (is-linear-combination-Ring R M μ x) + + is-prop-is-mere-linear-combination-Ring : + (x : type-Unital-Magma M) → + is-prop (is-mere-linear-combination-Ring x) + is-prop-is-mere-linear-combination-Ring x = + is-prop-type-trunc-Prop +``` + +### Linear combinations of subsets of a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) + where + + linear-combination-subset-Ring : + UU (l1 ⊔ l2) + linear-combination-subset-Ring = + linear-combination-Ring R (type-subset-Ring R S) + + ev-linear-combination-subset-Ring : + linear-combination-subset-Ring → type-Ring R + ev-linear-combination-subset-Ring = + ev-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ r s → + mul-Ring R (mul-Ring R r (inclusion-subset-Ring R S s))) + + is-linear-combination-subset-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-linear-combination-subset-Ring = + is-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ r s → + mul-Ring R (mul-Ring R r (inclusion-subset-Ring R S s))) +``` + +### Linear combinations of families of elements in a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) {I : UU l2} (a : I → type-Ring R) + where + + linear-combination-family-of-elements-Ring : + UU (l1 ⊔ l2) + linear-combination-family-of-elements-Ring = + linear-combination-subset-Ring R (trunc-Prop ∘ fiber a) + + ev-linear-combination-family-of-elements-Ring : + linear-combination-family-of-elements-Ring → type-Ring R + ev-linear-combination-family-of-elements-Ring = + ev-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) + + is-linear-combination-family-of-elements-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-linear-combination-family-of-elements-Ring = + is-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) +``` + +### Linear combinations of a single element in a ring + +Even though left linear combinations and right linear combinations of a single element $a$ in a ring $R$ can always be written in the form $(r,a)$ or $(a,r)$, resepectively, i.e., any element of the form $r_0a + \cdots + r_{n-1}a$ is equal to an element of the form $ra$ and similar for right linear combinations, the two-sided case is a bit different in that there might be rings in which an element of the form + +$$ + r_0as_0 + \cdots + r_{n-1}as_{n-1} +$$ + +is not equal to an element of the form $ras$, because the distributivity laws don't apply in this more general case. + +```agda +module _ + {l1 : Level} (R : Ring l1) (a : type-Ring R) + where + + linear-combination-element-Ring : + UU l1 + linear-combination-element-Ring = + linear-combination-subset-Ring R (λ y → Id-Prop (set-Ring R) y a) + + ev-linear-combination-element-Ring : + linear-combination-element-Ring → type-Ring R + ev-linear-combination-element-Ring = + ev-linear-combination-subset-Ring R + ( λ y → Id-Prop (set-Ring R) y a) + + is-linear-combination-element-Ring : + type-Ring R → UU l1 + is-linear-combination-element-Ring = + is-linear-combination-subset-Ring R + ( λ y → Id-Prop (set-Ring R) y a) +``` diff --git a/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md b/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md new file mode 100644 index 0000000000..c2755085fd --- /dev/null +++ b/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md @@ -0,0 +1,200 @@ +# Linear combinations with respect to semirings + +```agda +module ring-theory.linear-combinations-of-elements-semirings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.universe-levels + +open import ring-theory.semirings +open import ring-theory.subsets-semirings + +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [semiring](ring-theory.semirings.md) $R$ and a type $A$. A {#concept "linear combination"}} of elements of $A$ is a [list](lists.lists.md) of pairs $(r,a,s)$ consisting of an element $r,s:R$ and an element $a:A$. + +Furthermore, if we are given an action $\mu : R \to A \to R \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a linear combination $((r_0,a_0,s_0),\ldots,(r_{n-1},a_{n-1},s_{n-1}))$ by defining + +$$ + ev((r_0,a_0,s_0),\ldots,(r_{n-1},a_{n-1},s_{n-1})) := + \sum_{i=0}^{n-1} \mu(r_i,a_i,s_i). +$$ + +To be explicit, linear combinations of elements of a type $A$ have the semiring coefficients on both sides. + +## Definitions + +### The type of linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (A : UU l2) + where + + linear-combination-Semiring : + UU (l1 ⊔ l2) + linear-combination-Semiring = + list (type-Semiring R × A × type-Semiring R) +``` + +### Evaluating linear combinations of elements in magmas + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Semiring R → A → type-Semiring R → type-Unital-Magma M) + where + + ev-linear-combination-Semiring : + linear-combination-Semiring R A → type-Unital-Magma M + ev-linear-combination-Semiring nil = + unit-Unital-Magma M + ev-linear-combination-Semiring (cons (r , a , s) l) = + mul-Unital-Magma M (μ r a s) (ev-linear-combination-Semiring l) +``` + +### The predicate of being a linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Semiring R → A → type-Semiring R → type-Unital-Magma M) + where + + is-linear-combination-Semiring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-linear-combination-Semiring = + fiber (ev-linear-combination-Semiring R M μ) +``` + +### The predicate of being a mere linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : type-Semiring R → A → type-Semiring R → type-Unital-Magma M) + where + + is-mere-linear-combination-prop-Semiring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-linear-combination-prop-Semiring x = + trunc-Prop (is-linear-combination-Semiring R M μ x) + + is-mere-linear-combination-Semiring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-linear-combination-Semiring x = + type-trunc-Prop (is-linear-combination-Semiring R M μ x) + + is-prop-is-mere-linear-combination-Semiring : + (x : type-Unital-Magma M) → + is-prop (is-mere-linear-combination-Semiring x) + is-prop-is-mere-linear-combination-Semiring x = + is-prop-type-trunc-Prop +``` + +### Linear combinations of subsets of a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (S : subset-Semiring l2 R) + where + + linear-combination-subset-Semiring : + UU (l1 ⊔ l2) + linear-combination-subset-Semiring = + linear-combination-Semiring R (type-subset-Semiring R S) + + ev-linear-combination-subset-Semiring : + linear-combination-subset-Semiring → type-Semiring R + ev-linear-combination-subset-Semiring = + ev-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ r s → + mul-Semiring R (mul-Semiring R r (inclusion-subset-Semiring R S s))) + + is-linear-combination-subset-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-linear-combination-subset-Semiring = + is-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ r s → + mul-Semiring R (mul-Semiring R r (inclusion-subset-Semiring R S s))) +``` + +### Linear combinations of families of elements in a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) {I : UU l2} (a : I → type-Semiring R) + where + + linear-combination-family-of-elements-Semiring : + UU (l1 ⊔ l2) + linear-combination-family-of-elements-Semiring = + linear-combination-subset-Semiring R (trunc-Prop ∘ fiber a) + + ev-linear-combination-family-of-elements-Semiring : + linear-combination-family-of-elements-Semiring → type-Semiring R + ev-linear-combination-family-of-elements-Semiring = + ev-linear-combination-subset-Semiring R + ( trunc-Prop ∘ fiber a) + + is-linear-combination-family-of-elements-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-linear-combination-family-of-elements-Semiring = + is-linear-combination-subset-Semiring R + ( trunc-Prop ∘ fiber a) +``` + +### Linear combinations of a single element in a semiring + +Even though left linear combinations and right linear combinations of a single element $a$ in a semiring $R$ can always be written in the form $(r,a)$ or $(a,r)$, resepectively, i.e., any element of the form $r_0a + \cdots + r_{n-1}a$ is equal to an element of the form $ra$ and similar for right linear combinations, the two-sided case is a bit different in that there might be semirings in which an element of the form + +$$ + r_0as_0 + \cdots + r_{n-1}as_{n-1} +$$ + +is not equal to an element of the form $ras$, because the distributivity laws don't apply in this more general case. + +```agda +module _ + {l1 : Level} (R : Semiring l1) (a : type-Semiring R) + where + + linear-combination-element-Semiring : + UU l1 + linear-combination-element-Semiring = + linear-combination-subset-Semiring R (λ y → Id-Prop (set-Semiring R) y a) + + ev-linear-combination-element-Semiring : + linear-combination-element-Semiring → type-Semiring R + ev-linear-combination-element-Semiring = + ev-linear-combination-subset-Semiring R + ( λ y → Id-Prop (set-Semiring R) y a) + + is-linear-combination-element-Semiring : + type-Semiring R → UU l1 + is-linear-combination-element-Semiring = + is-linear-combination-subset-Semiring R + ( λ y → Id-Prop (set-Semiring R) y a) +``` diff --git a/src/ring-theory/monoids-with-semiring-actions.lagda.md b/src/ring-theory/monoids-with-semiring-actions.lagda.md new file mode 100644 index 0000000000..16cc5fabee --- /dev/null +++ b/src/ring-theory/monoids-with-semiring-actions.lagda.md @@ -0,0 +1,207 @@ +# Monoids with semiring actions + +```agda +module ring-theory.monoids-with-semiring-actions where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.sets +open import foundation.universe-levels + +open import group-theory.monoids + +open import ring-theory.semirings +``` + +
+ +## Idea + +Consider a [semiring](ring-theory.semirings.md) $R$. +A {{#concept "monoid with a left semiring action" Agda=Monoid-With-Left-Semiring-Action}} +from $R$ consists of a [monoid](group-theory.monoids.md) $M$ and a binary operation $\mu : R \to M \to M$ satisfying the axioms of a ring action: + +1. The action distributes from the left over monoid multiplication: + $$ + r(xy) = (rx)(ry). + $$ +2. The action distributes from the right over addition in the semiring: + $$ + (r+s)x = (rx)(sx). + $$ +3. Associativity of the action + $$ + (sr)x = s(rx). + $$ +4. The unit element of the semiring acts as the identity + $$ + 1x = x. + $$ +5. The unit element of the monoid absorbs any action + $$ + r1 = 1. + $$ + +## Definitions + +### Left actions of semirings on monoids + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (M : Monoid l2) + where + + left-action-semiring-Monoid : UU (l1 ⊔ l2) + left-action-semiring-Monoid = + Σ ( type-Semiring R → type-Monoid M → type-Monoid M) + ( λ μ → + ( (r : type-Semiring R) (x y : type-Monoid M) → + μ r (mul-Monoid M x y) = mul-Monoid M (μ r x) (μ r y)) × + ( (r s : type-Semiring R) (x : type-Monoid M) → + μ (add-Semiring R r s) x = mul-Monoid M (μ r x) (μ s x)) × + ( (s r : type-Semiring R) (x : type-Monoid M) → + μ (mul-Semiring R s r) x = μ s (μ r x)) × + ( (x : type-Monoid M) → μ (one-Semiring R) x = x) × + ( (r : type-Semiring R) → μ r (unit-Monoid M) = unit-Monoid M)) +``` + +### Monoids with left semiring actions + +```agda +Monoid-With-Left-Semiring-Action : + {l1 : Level} (l2 : Level) (R : Semiring l1) → UU (l1 ⊔ lsuc l2) +Monoid-With-Left-Semiring-Action l2 R = + Σ ( Monoid l2) (left-action-semiring-Monoid R) + +module _ + {l1 l2 : Level} (R : Semiring l1) (M : Monoid-With-Left-Semiring-Action l2 R) + where + + monoid-Monoid-With-Left-Semiring-Action : + Monoid l2 + monoid-Monoid-With-Left-Semiring-Action = + pr1 M + + set-Monoid-With-Left-Semiring-Action : + Set l2 + set-Monoid-With-Left-Semiring-Action = + set-Monoid monoid-Monoid-With-Left-Semiring-Action + + type-Monoid-With-Left-Semiring-Action : + UU l2 + type-Monoid-With-Left-Semiring-Action = + type-Monoid monoid-Monoid-With-Left-Semiring-Action + + is-set-type-Monoid-With-Left-Semiring-Action : + is-set type-Monoid-With-Left-Semiring-Action + is-set-type-Monoid-With-Left-Semiring-Action = + is-set-type-Monoid monoid-Monoid-With-Left-Semiring-Action + + mul-Monoid-With-Left-Semiring-Action : + (x y : type-Monoid-With-Left-Semiring-Action) → type-Monoid-With-Left-Semiring-Action + mul-Monoid-With-Left-Semiring-Action = + mul-Monoid monoid-Monoid-With-Left-Semiring-Action + + associative-mul-Monoid-With-Left-Semiring-Action : + (x y z : type-Monoid-With-Left-Semiring-Action) → + mul-Monoid-With-Left-Semiring-Action (mul-Monoid-With-Left-Semiring-Action x y) z = + mul-Monoid-With-Left-Semiring-Action x (mul-Monoid-With-Left-Semiring-Action y z) + associative-mul-Monoid-With-Left-Semiring-Action = + associative-mul-Monoid monoid-Monoid-With-Left-Semiring-Action + + unit-Monoid-With-Left-Semiring-Action : + type-Monoid-With-Left-Semiring-Action + unit-Monoid-With-Left-Semiring-Action = + unit-Monoid monoid-Monoid-With-Left-Semiring-Action + + left-unit-law-mul-Monoid-With-Left-Semiring-Action : + (x : type-Monoid-With-Left-Semiring-Action) → + mul-Monoid-With-Left-Semiring-Action + ( unit-Monoid-With-Left-Semiring-Action) + ( x) = + x + left-unit-law-mul-Monoid-With-Left-Semiring-Action = + left-unit-law-mul-Monoid monoid-Monoid-With-Left-Semiring-Action + + right-unit-law-mul-Monoid-With-Left-Semiring-Action : + (x : type-Monoid-With-Left-Semiring-Action) → + mul-Monoid-With-Left-Semiring-Action + ( x) + ( unit-Monoid-With-Left-Semiring-Action) = + ( x) + right-unit-law-mul-Monoid-With-Left-Semiring-Action = + right-unit-law-mul-Monoid monoid-Monoid-With-Left-Semiring-Action + + action-Monoid-With-Left-Semiring-Action : + type-Semiring R → type-Monoid-With-Left-Semiring-Action → + type-Monoid-With-Left-Semiring-Action + action-Monoid-With-Left-Semiring-Action = + pr1 (pr2 M) + + left-distributive-action-mul-Monoid-With-Left-Semiring-Action : + (r : type-Semiring R) (x y : type-Monoid-With-Left-Semiring-Action) → + action-Monoid-With-Left-Semiring-Action r + ( mul-Monoid-With-Left-Semiring-Action x y) = + mul-Monoid-With-Left-Semiring-Action + ( action-Monoid-With-Left-Semiring-Action r x) + ( action-Monoid-With-Left-Semiring-Action r y) + left-distributive-action-mul-Monoid-With-Left-Semiring-Action = + pr1 (pr2 (pr2 M)) + + right-distributive-action-add-Monoid-With-Left-Semiring-Action : + (r s : type-Semiring R) (x : type-Monoid-With-Left-Semiring-Action) → + action-Monoid-With-Left-Semiring-Action (add-Semiring R r s) x = + mul-Monoid-With-Left-Semiring-Action + ( action-Monoid-With-Left-Semiring-Action r x) + ( action-Monoid-With-Left-Semiring-Action s x) + right-distributive-action-add-Monoid-With-Left-Semiring-Action = + pr1 (pr2 (pr2 (pr2 M))) + + associative-action-Monoid-With-Left-Semiring-Action : + (s r : type-Semiring R) (x : type-Monoid-With-Left-Semiring-Action) → + action-Monoid-With-Left-Semiring-Action (mul-Semiring R s r) x = + action-Monoid-With-Left-Semiring-Action s + ( action-Monoid-With-Left-Semiring-Action r x) + associative-action-Monoid-With-Left-Semiring-Action = + pr1 (pr2 (pr2 (pr2 (pr2 M)))) + + left-unit-law-action-Monoid-With-Left-Semiring-Action : + (x : type-Monoid-With-Left-Semiring-Action) → + action-Monoid-With-Left-Semiring-Action (one-Semiring R) x = x + left-unit-law-action-Monoid-With-Left-Semiring-Action = + pr1 (pr2 (pr2 (pr2 (pr2 (pr2 M))))) + + right-absorption-law-action-Monoid-With-Left-Semiring-Action : + (r : type-Semiring R) → + action-Monoid-With-Left-Semiring-Action r + ( unit-Monoid-With-Left-Semiring-Action) = + unit-Monoid-With-Left-Semiring-Action + right-absorption-law-action-Monoid-With-Left-Semiring-Action = + pr2 (pr2 (pr2 (pr2 (pr2 (pr2 M))))) +``` + +## Properties + +### The underlying additive monoid of a semiring is a monoid with a left semiring action + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + additive-monoid-with-left-semiring-action-Semiring : + Monoid-With-Left-Semiring-Action l1 R + additive-monoid-with-left-semiring-action-Semiring = + ( additive-monoid-Semiring R , + mul-Semiring R , + left-distributive-mul-add-Semiring R , + right-distributive-mul-add-Semiring R , + associative-mul-Semiring R , + left-unit-law-mul-Semiring R , + right-zero-law-mul-Semiring R) +``` diff --git a/src/ring-theory/poset-of-left-ideals-semirings.lagda.md b/src/ring-theory/poset-of-left-ideals-semirings.lagda.md new file mode 100644 index 0000000000..d1f2c3de0b --- /dev/null +++ b/src/ring-theory/poset-of-left-ideals-semirings.lagda.md @@ -0,0 +1,178 @@ +# The poset of left ideals of a semiring + +```agda +module ring-theory.poset-of-left-ideals-semirings where +``` + +
Imports + +```agda +open import foundation.binary-relations +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.powersets +open import foundation.propositions +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.large-posets +open import order-theory.large-preorders +open import order-theory.order-preserving-maps-large-posets +open import order-theory.order-preserving-maps-large-preorders +open import order-theory.similarity-of-elements-large-posets + +open import ring-theory.left-ideals-semirings +open import ring-theory.semirings +``` + +
+ +## Idea + +The [left ideals](ring-theory.left-ideals-semirings.md) of a +[semiring](ring-theory.semirings.md) form a [large poset](order-theory.large-posets.md) +ordered by inclusion. + +## Definition + +### The inclusion relation on left ideals + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + leq-prop-left-ideal-Semiring : + {l2 l3 : Level} → + left-ideal-Semiring l2 R → left-ideal-Semiring l3 R → Prop (l1 ⊔ l2 ⊔ l3) + leq-prop-left-ideal-Semiring I J = + leq-prop-subtype + ( subset-left-ideal-Semiring R I) + ( subset-left-ideal-Semiring R J) + + leq-left-ideal-Semiring : + {l2 l3 : Level} → + left-ideal-Semiring l2 R → left-ideal-Semiring l3 R → UU (l1 ⊔ l2 ⊔ l3) + leq-left-ideal-Semiring I J = + subset-left-ideal-Semiring R I ⊆ subset-left-ideal-Semiring R J + + is-prop-leq-left-ideal-Semiring : + {l2 l3 : Level} + (I : left-ideal-Semiring l2 R) (J : left-ideal-Semiring l3 R) → + is-prop (leq-left-ideal-Semiring I J) + is-prop-leq-left-ideal-Semiring I J = + is-prop-leq-subtype + ( subset-left-ideal-Semiring R I) + ( subset-left-ideal-Semiring R J) + + refl-leq-left-ideal-Semiring : + {l2 : Level} → is-reflexive (leq-left-ideal-Semiring {l2}) + refl-leq-left-ideal-Semiring I = + refl-leq-subtype (subset-left-ideal-Semiring R I) + + transitive-leq-left-ideal-Semiring : + {l2 l3 l4 : Level} + (I : left-ideal-Semiring l2 R) + (J : left-ideal-Semiring l3 R) + (K : left-ideal-Semiring l4 R) → + leq-left-ideal-Semiring J K → + leq-left-ideal-Semiring I J → + leq-left-ideal-Semiring I K + transitive-leq-left-ideal-Semiring I J K = + transitive-leq-subtype + ( subset-left-ideal-Semiring R I) + ( subset-left-ideal-Semiring R J) + ( subset-left-ideal-Semiring R K) + + antisymmetric-leq-left-ideal-Semiring : + {l2 : Level} → is-antisymmetric (leq-left-ideal-Semiring {l2}) + antisymmetric-leq-left-ideal-Semiring I J U V = + eq-has-same-elements-left-ideal-Semiring R I J (λ x → U x , V x) +``` + +### The large poset of left ideals + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + left-ideal-Semiring-Large-Preorder : + Large-Preorder (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3) + type-Large-Preorder left-ideal-Semiring-Large-Preorder l = + left-ideal-Semiring l R + leq-prop-Large-Preorder left-ideal-Semiring-Large-Preorder = + leq-prop-left-ideal-Semiring R + refl-leq-Large-Preorder left-ideal-Semiring-Large-Preorder = + refl-leq-left-ideal-Semiring R + transitive-leq-Large-Preorder left-ideal-Semiring-Large-Preorder = + transitive-leq-left-ideal-Semiring R + + left-ideal-Semiring-Large-Poset : + Large-Poset (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3) + large-preorder-Large-Poset left-ideal-Semiring-Large-Poset = + left-ideal-Semiring-Large-Preorder + antisymmetric-leq-Large-Poset left-ideal-Semiring-Large-Poset = + antisymmetric-leq-left-ideal-Semiring R +``` + +### The similarity relation on left ideals in a semiring + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + sim-prop-left-ideal-Semiring : + {l2 l3 : Level} + (I : left-ideal-Semiring l2 R) (J : left-ideal-Semiring l3 R) → + Prop (l1 ⊔ l2 ⊔ l3) + sim-prop-left-ideal-Semiring = + sim-prop-Large-Poset (left-ideal-Semiring-Large-Poset R) + + sim-left-ideal-Semiring : + {l2 l3 : Level} + (I : left-ideal-Semiring l2 R) (J : left-ideal-Semiring l3 R) → + UU (l1 ⊔ l2 ⊔ l3) + sim-left-ideal-Semiring = sim-Large-Poset (left-ideal-Semiring-Large-Poset R) + + is-prop-sim-left-ideal-Semiring : + {l2 l3 : Level} + (I : left-ideal-Semiring l2 R) (J : left-ideal-Semiring l3 R) → + is-prop (sim-left-ideal-Semiring I J) + is-prop-sim-left-ideal-Semiring = + is-prop-sim-Large-Poset (left-ideal-Semiring-Large-Poset R) + + eq-sim-left-ideal-Semiring : + {l2 : Level} + (I J : left-ideal-Semiring l2 R) → sim-left-ideal-Semiring I J → I = J + eq-sim-left-ideal-Semiring = eq-sim-Large-Poset (left-ideal-Semiring-Large-Poset R) +``` + +## Properties + +### The forgetful function from left ideals to subsets preserves inclusions + +```agda +module _ + {l : Level} (R : Semiring l) + where + + preserves-order-subset-left-ideal-Semiring : + {l1 l2 : Level} + (I : left-ideal-Semiring l1 R) (J : left-ideal-Semiring l2 R) → + leq-left-ideal-Semiring R I J → + subset-left-ideal-Semiring R I ⊆ subset-left-ideal-Semiring R J + preserves-order-subset-left-ideal-Semiring I J H = H + + subset-left-ideal-hom-large-poset-Semiring : + hom-Large-Poset + ( λ l → l) + ( left-ideal-Semiring-Large-Poset R) + ( powerset-Large-Poset (type-Semiring R)) + map-hom-Large-Preorder subset-left-ideal-hom-large-poset-Semiring = + subset-left-ideal-Semiring R + preserves-order-hom-Large-Preorder + subset-left-ideal-hom-large-poset-Semiring = + preserves-order-subset-left-ideal-Semiring +``` diff --git a/src/ring-theory/right-ideals-semirings.lagda.md b/src/ring-theory/right-ideals-semirings.lagda.md new file mode 100644 index 0000000000..48f392a530 --- /dev/null +++ b/src/ring-theory/right-ideals-semirings.lagda.md @@ -0,0 +1,195 @@ +# Right ideals of semirings + +```agda +module ring-theory.right-ideals-semirings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fundamental-theorem-of-identity-types +open import foundation.identity-types +open import foundation.propositions +open import foundation.subtype-identity-principle +open import foundation.subtypes +open import foundation.torsorial-type-families +open import foundation.universe-levels + +open import ring-theory.semirings +open import ring-theory.subsets-semirings +``` + +
+ +## Idea + +A {{#concept "right ideal" Disambiguation="semiring" Agda=right-ideal-Semiring}} in a [semiring](ring-theory.semirings.md) `R` is a right submodule of +`R`. + +## Definitions + +### Right ideals + +```agda +module _ + {l1 : Level} (R : Semiring l1) + where + + is-right-ideal-subset-Semiring : + {l2 : Level} → subset-Semiring l2 R → UU (l1 ⊔ l2) + is-right-ideal-subset-Semiring P = + is-additive-submonoid-subset-Semiring R P × + is-closed-under-right-multiplication-subset-Semiring R P + + is-prop-is-right-ideal-subset-Semiring : + {l2 : Level} (S : subset-Semiring l2 R) → + is-prop (is-right-ideal-subset-Semiring S) + is-prop-is-right-ideal-subset-Semiring S = + is-prop-product + ( is-prop-is-additive-submonoid-subset-Semiring R S) + ( is-prop-is-closed-under-right-multiplication-subset-Semiring R S) + +right-ideal-Semiring : + (l : Level) {l1 : Level} (R : Semiring l1) → UU (lsuc l ⊔ l1) +right-ideal-Semiring l R = + Σ (subset-Semiring l R) (is-right-ideal-subset-Semiring R) + +module _ + {l1 l2 : Level} (R : Semiring l1) (I : right-ideal-Semiring l2 R) + where + + subset-right-ideal-Semiring : subset-Semiring l2 R + subset-right-ideal-Semiring = pr1 I + + is-in-right-ideal-Semiring : type-Semiring R → UU l2 + is-in-right-ideal-Semiring x = type-Prop (subset-right-ideal-Semiring x) + + type-right-ideal-Semiring : UU (l1 ⊔ l2) + type-right-ideal-Semiring = type-subset-Semiring R subset-right-ideal-Semiring + + inclusion-right-ideal-Semiring : + type-right-ideal-Semiring → type-Semiring R + inclusion-right-ideal-Semiring = + inclusion-subset-Semiring R subset-right-ideal-Semiring + + ap-inclusion-right-ideal-Semiring : + (x y : type-right-ideal-Semiring) → x = y → + inclusion-right-ideal-Semiring x = inclusion-right-ideal-Semiring y + ap-inclusion-right-ideal-Semiring = + ap-inclusion-subset-Semiring R subset-right-ideal-Semiring + + is-in-subset-inclusion-right-ideal-Semiring : + (x : type-right-ideal-Semiring) → + is-in-right-ideal-Semiring (inclusion-right-ideal-Semiring x) + is-in-subset-inclusion-right-ideal-Semiring = + is-in-subset-inclusion-subset-Semiring R subset-right-ideal-Semiring + + is-closed-under-eq-right-ideal-Semiring : + {x y : type-Semiring R} → is-in-right-ideal-Semiring x → + (x = y) → is-in-right-ideal-Semiring y + is-closed-under-eq-right-ideal-Semiring = + is-closed-under-eq-subset-Semiring R subset-right-ideal-Semiring + + is-closed-under-eq-right-ideal-Semiring' : + {x y : type-Semiring R} → is-in-right-ideal-Semiring y → + (x = y) → is-in-right-ideal-Semiring x + is-closed-under-eq-right-ideal-Semiring' = + is-closed-under-eq-subset-Semiring' R subset-right-ideal-Semiring + + is-right-ideal-subset-right-ideal-Semiring : + is-right-ideal-subset-Semiring R subset-right-ideal-Semiring + is-right-ideal-subset-right-ideal-Semiring = pr2 I + + is-additive-submonoid-subset-right-ideal-Semiring : + is-additive-submonoid-subset-Semiring R subset-right-ideal-Semiring + is-additive-submonoid-subset-right-ideal-Semiring = + pr1 is-right-ideal-subset-right-ideal-Semiring + + contains-zero-right-ideal-Semiring : + is-in-right-ideal-Semiring (zero-Semiring R) + contains-zero-right-ideal-Semiring = + pr1 is-additive-submonoid-subset-right-ideal-Semiring + + is-closed-under-addition-right-ideal-Semiring : + is-closed-under-addition-subset-Semiring R subset-right-ideal-Semiring + is-closed-under-addition-right-ideal-Semiring = + pr2 is-additive-submonoid-subset-right-ideal-Semiring + + is-closed-under-right-multiplication-right-ideal-Semiring : + is-closed-under-right-multiplication-subset-Semiring R + subset-right-ideal-Semiring + is-closed-under-right-multiplication-right-ideal-Semiring = + pr2 is-right-ideal-subset-right-ideal-Semiring + + is-right-ideal-right-ideal-Semiring : + is-right-ideal-subset-Semiring R subset-right-ideal-Semiring + is-right-ideal-right-ideal-Semiring = pr2 I +``` + +## Properties + +### Characterizing equality of right ideals in semirings + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) (I : right-ideal-Semiring l2 R) + where + + has-same-elements-right-ideal-Semiring : + (J : right-ideal-Semiring l3 R) → UU (l1 ⊔ l2 ⊔ l3) + has-same-elements-right-ideal-Semiring J = + has-same-elements-subtype + ( subset-right-ideal-Semiring R I) + ( subset-right-ideal-Semiring R J) + +module _ + {l1 l2 : Level} (R : Semiring l1) (I : right-ideal-Semiring l2 R) + where + + refl-has-same-elements-right-ideal-Semiring : + has-same-elements-right-ideal-Semiring R I I + refl-has-same-elements-right-ideal-Semiring = + refl-has-same-elements-subtype (subset-right-ideal-Semiring R I) + + is-torsorial-has-same-elements-right-ideal-Semiring : + is-torsorial (has-same-elements-right-ideal-Semiring R I) + is-torsorial-has-same-elements-right-ideal-Semiring = + is-torsorial-Eq-subtype + ( is-torsorial-has-same-elements-subtype + ( subset-right-ideal-Semiring R I)) + ( is-prop-is-right-ideal-subset-Semiring R) + ( subset-right-ideal-Semiring R I) + ( refl-has-same-elements-right-ideal-Semiring) + ( is-right-ideal-right-ideal-Semiring R I) + + has-same-elements-eq-right-ideal-Semiring : + (J : right-ideal-Semiring l2 R) → + (I = J) → has-same-elements-right-ideal-Semiring R I J + has-same-elements-eq-right-ideal-Semiring .I refl = + refl-has-same-elements-right-ideal-Semiring + + is-equiv-has-same-elements-eq-right-ideal-Semiring : + (J : right-ideal-Semiring l2 R) → + is-equiv (has-same-elements-eq-right-ideal-Semiring J) + is-equiv-has-same-elements-eq-right-ideal-Semiring = + fundamental-theorem-id + is-torsorial-has-same-elements-right-ideal-Semiring + has-same-elements-eq-right-ideal-Semiring + + extensionality-right-ideal-Semiring : + (J : right-ideal-Semiring l2 R) → + (I = J) ≃ has-same-elements-right-ideal-Semiring R I J + pr1 (extensionality-right-ideal-Semiring J) = + has-same-elements-eq-right-ideal-Semiring J + pr2 (extensionality-right-ideal-Semiring J) = + is-equiv-has-same-elements-eq-right-ideal-Semiring J + + eq-has-same-elements-right-ideal-Semiring : + (J : right-ideal-Semiring l2 R) → + has-same-elements-right-ideal-Semiring R I J → I = J + eq-has-same-elements-right-ideal-Semiring J = + map-inv-equiv (extensionality-right-ideal-Semiring J) +``` diff --git a/src/ring-theory/right-linear-combinations-of-elements-rings.lagda.md b/src/ring-theory/right-linear-combinations-of-elements-rings.lagda.md new file mode 100644 index 0000000000..1932248150 --- /dev/null +++ b/src/ring-theory/right-linear-combinations-of-elements-rings.lagda.md @@ -0,0 +1,163 @@ +# Right linear combinations with respect to rings + +```agda +module ring-theory.right-linear-combinations-of-elements-rings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.universe-levels + +open import ring-theory.rings +open import ring-theory.subsets-rings + +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [ring](ring-theory.rings.md) $R$ and a type $A$. A {#concept "right linear combination"}} of elements of $R$ is a [list](lists.lists.md) of pairs $(a,r)$ consisting of an element $a:A$ and an element $r:R$. + +Furthermore, if we are given an action $\mu : A \to R \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a right linear combination $((a_0,r_0),\ldots,(a_{n-1},r_{n-1}))$ by defining + +$$ + ev((a_0,r_0),\ldots,(a_{n-1},r_{n-1})) := \sum_{i=0}^{n-1} \mu(a_i,r_i). +$$ + +To be explicit, right linear combinations of elements of a type $A$ have the ring coefficients on the right. + +## Definitions + +### The type of right linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (A : UU l2) + where + + right-linear-combination-Ring : + UU (l1 ⊔ l2) + right-linear-combination-Ring = + list (A × type-Ring R) +``` + +### Evaluating right linear combinations of elements in unital magmas + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Ring R → type-Unital-Magma M) + where + + ev-right-linear-combination-Ring : + right-linear-combination-Ring R A → type-Unital-Magma M + ev-right-linear-combination-Ring nil = + unit-Unital-Magma M + ev-right-linear-combination-Ring (cons (a , r) l) = + mul-Unital-Magma M (μ a r) (ev-right-linear-combination-Ring l) +``` + +### The predicate of being a right linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Ring R → type-Unital-Magma M) + where + + is-right-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-right-linear-combination-Ring = + fiber (ev-right-linear-combination-Ring R M μ) +``` + +### The predicate of being a mere right linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Ring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Ring R → type-Unital-Magma M) + where + + is-mere-right-linear-combination-prop-Ring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-right-linear-combination-prop-Ring x = + trunc-Prop (is-right-linear-combination-Ring R M μ x) + + is-mere-right-linear-combination-Ring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-right-linear-combination-Ring x = + type-trunc-Prop (is-right-linear-combination-Ring R M μ x) + + is-prop-is-mere-right-linear-combination-Ring : + (x : type-Unital-Magma M) → + is-prop (is-mere-right-linear-combination-Ring x) + is-prop-is-mere-right-linear-combination-Ring x = + is-prop-type-trunc-Prop +``` + +### Right linear combinations of subsets of a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) + where + + right-linear-combination-subset-Ring : + UU (l1 ⊔ l2) + right-linear-combination-subset-Ring = + right-linear-combination-Ring R (type-subset-Ring R S) + + ev-right-linear-combination-subset-Ring : + right-linear-combination-subset-Ring → type-Ring R + ev-right-linear-combination-subset-Ring = + ev-right-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ s → mul-Ring R (inclusion-subset-Ring R S s)) + + is-right-linear-combination-subset-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-right-linear-combination-subset-Ring = + is-right-linear-combination-Ring R + ( additive-unital-magma-Ring R) + ( λ s → mul-Ring R (inclusion-subset-Ring R S s)) +``` + +### Right linear combinations of families of elements in a ring + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) {I : UU l2} (a : I → type-Ring R) + where + + right-linear-combination-family-of-elements-Ring : + UU (l1 ⊔ l2) + right-linear-combination-family-of-elements-Ring = + right-linear-combination-subset-Ring R (trunc-Prop ∘ fiber a) + + ev-right-linear-combination-family-of-elements-Ring : + right-linear-combination-family-of-elements-Ring → type-Ring R + ev-right-linear-combination-family-of-elements-Ring = + ev-right-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) + + is-right-linear-combination-family-of-elements-Ring : + type-Ring R → UU (l1 ⊔ l2) + is-right-linear-combination-family-of-elements-Ring = + is-right-linear-combination-subset-Ring R + ( trunc-Prop ∘ fiber a) +``` diff --git a/src/ring-theory/right-linear-combinations-of-elements-semirings.lagda.md b/src/ring-theory/right-linear-combinations-of-elements-semirings.lagda.md new file mode 100644 index 0000000000..3571008e50 --- /dev/null +++ b/src/ring-theory/right-linear-combinations-of-elements-semirings.lagda.md @@ -0,0 +1,163 @@ +# Right linear combinations with respect to semirings + +```agda +module ring-theory.right-linear-combinations-of-elements-semirings where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.universe-levels + +open import ring-theory.semirings +open import ring-theory.subsets-semirings + +open import lists.lists + +open import structured-types.magmas +``` + +
+ +## Idea + +Consider a [semiring](ring-theory.semirings.md) $R$ and a type $A$. A {#concept "right linear combination"}} of elements of $R$ is a [list](lists.lists.md) of pairs $(a,r)$ consisting of an element $a:A$ and an element $r:R$. + +Furthermore, if we are given an action $\mu : A \to R \to M$ taking values in a [unital magma](structured-types.magmas.md) $(M,+,0)$, then we can evaluate a right linear combination $((a_0,r_0),\ldots,(a_{n-1},r_{n-1}))$ by defining + +$$ + ev((a_0,r_0),\ldots,(a_{n-1},r_{n-1})) := \sum_{i=0}^{n-1} \mu(a_i,r_i). +$$ + +To be explicit, right linear combinations of elements of a type $A$ have the semiring coefficients on the right. + +## Definitions + +### The type of right linear combinations + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (A : UU l2) + where + + right-linear-combination-Semiring : + UU (l1 ⊔ l2) + right-linear-combination-Semiring = + list (A × type-Semiring R) +``` + +### Evaluating right linear combinations of elements in unital magmas + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Semiring R → type-Unital-Magma M) + where + + ev-right-linear-combination-Semiring : + right-linear-combination-Semiring R A → type-Unital-Magma M + ev-right-linear-combination-Semiring nil = + unit-Unital-Magma M + ev-right-linear-combination-Semiring (cons (a , r) l) = + mul-Unital-Magma M (μ a r) (ev-right-linear-combination-Semiring l) +``` + +### The predicate of being a right linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Semiring R → type-Unital-Magma M) + where + + is-right-linear-combination-Semiring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-right-linear-combination-Semiring = + fiber (ev-right-linear-combination-Semiring R M μ) +``` + +### The predicate of being a mere right linear combination + +```agda +module _ + {l1 l2 l3 : Level} (R : Semiring l1) + {A : UU l2} (M : Unital-Magma l3) + (μ : A → type-Semiring R → type-Unital-Magma M) + where + + is-mere-right-linear-combination-prop-Semiring : + type-Unital-Magma M → Prop (l1 ⊔ l2 ⊔ l3) + is-mere-right-linear-combination-prop-Semiring x = + trunc-Prop (is-right-linear-combination-Semiring R M μ x) + + is-mere-right-linear-combination-Semiring : + type-Unital-Magma M → UU (l1 ⊔ l2 ⊔ l3) + is-mere-right-linear-combination-Semiring x = + type-trunc-Prop (is-right-linear-combination-Semiring R M μ x) + + is-prop-is-mere-right-linear-combination-Semiring : + (x : type-Unital-Magma M) → + is-prop (is-mere-right-linear-combination-Semiring x) + is-prop-is-mere-right-linear-combination-Semiring x = + is-prop-type-trunc-Prop +``` + +### Right linear combinations of subsets of a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) (S : subset-Semiring l2 R) + where + + right-linear-combination-subset-Semiring : + UU (l1 ⊔ l2) + right-linear-combination-subset-Semiring = + right-linear-combination-Semiring R (type-subset-Semiring R S) + + ev-right-linear-combination-subset-Semiring : + right-linear-combination-subset-Semiring → type-Semiring R + ev-right-linear-combination-subset-Semiring = + ev-right-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ s → mul-Semiring R (inclusion-subset-Semiring R S s)) + + is-right-linear-combination-subset-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-right-linear-combination-subset-Semiring = + is-right-linear-combination-Semiring R + ( additive-unital-magma-Semiring R) + ( λ s → mul-Semiring R (inclusion-subset-Semiring R S s)) +``` + +### Right linear combinations of families of elements in a semiring + +```agda +module _ + {l1 l2 : Level} (R : Semiring l1) {I : UU l2} (a : I → type-Semiring R) + where + + right-linear-combination-family-of-elements-Semiring : + UU (l1 ⊔ l2) + right-linear-combination-family-of-elements-Semiring = + right-linear-combination-subset-Semiring R (trunc-Prop ∘ fiber a) + + ev-right-linear-combination-family-of-elements-Semiring : + right-linear-combination-family-of-elements-Semiring → type-Semiring R + ev-right-linear-combination-family-of-elements-Semiring = + ev-right-linear-combination-subset-Semiring R + ( trunc-Prop ∘ fiber a) + + is-right-linear-combination-family-of-elements-Semiring : + type-Semiring R → UU (l1 ⊔ l2) + is-right-linear-combination-family-of-elements-Semiring = + is-right-linear-combination-subset-Semiring R + ( trunc-Prop ∘ fiber a) +``` diff --git a/src/ring-theory/rings.lagda.md b/src/ring-theory/rings.lagda.md index 757f7994ee..e1acfb2833 100644 --- a/src/ring-theory/rings.lagda.md +++ b/src/ring-theory/rings.lagda.md @@ -39,6 +39,8 @@ open import lists.concatenation-lists open import lists.lists open import ring-theory.semirings + +open import structured-types.magmas ``` @@ -84,6 +86,9 @@ module _ additive-monoid-Ring : Monoid l additive-monoid-Ring = monoid-Ab ab-Ring + additive-unital-magma-Ring : Unital-Magma l + additive-unital-magma-Ring = unital-magma-Monoid additive-monoid-Ring + additive-semigroup-Ring : Semigroup l additive-semigroup-Ring = semigroup-Ab ab-Ring diff --git a/src/ring-theory/semirings.lagda.md b/src/ring-theory/semirings.lagda.md index 982869d8d3..7887825593 100644 --- a/src/ring-theory/semirings.lagda.md +++ b/src/ring-theory/semirings.lagda.md @@ -24,6 +24,8 @@ open import foundation.universe-levels open import group-theory.commutative-monoids open import group-theory.monoids open import group-theory.semigroups + +open import structured-types.magmas ``` @@ -81,6 +83,10 @@ module _ additive-monoid-Semiring = monoid-Commutative-Monoid additive-commutative-monoid-Semiring + additive-unital-magma-Semiring : Unital-Magma l + additive-unital-magma-Semiring = + unital-magma-Monoid additive-monoid-Semiring + additive-semigroup-Semiring : Semigroup l additive-semigroup-Semiring = semigroup-Commutative-Monoid additive-commutative-monoid-Semiring diff --git a/src/ring-theory/subsets-rings.lagda.md b/src/ring-theory/subsets-rings.lagda.md index df9a3f2c84..3c8e43061f 100644 --- a/src/ring-theory/subsets-rings.lagda.md +++ b/src/ring-theory/subsets-rings.lagda.md @@ -7,6 +7,8 @@ module ring-theory.subsets-rings where
Imports ```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.propositional-extensionality open import foundation.propositions @@ -17,6 +19,7 @@ open import foundation.universe-levels open import group-theory.subgroups-abelian-groups open import ring-theory.rings +open import ring-theory.subsets-semirings ```
@@ -32,33 +35,43 @@ A subset of a ring is a subtype of the underlying type of a ring ```agda subset-Ring : (l : Level) {l1 : Level} (R : Ring l1) → UU (lsuc l ⊔ l1) -subset-Ring l R = subtype l (type-Ring R) +subset-Ring l R = + subset-Semiring l (semiring-Ring R) is-set-subset-Ring : (l : Level) {l1 : Level} (R : Ring l1) → is-set (subset-Ring l R) is-set-subset-Ring l R = - is-set-function-type is-set-type-Prop + is-set-subset-Semiring l (semiring-Ring R) module _ {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) where - is-in-subset-Ring : type-Ring R → UU l2 - is-in-subset-Ring = is-in-subtype S + is-in-subset-Ring : + type-Ring R → UU l2 + is-in-subset-Ring = + is-in-subset-Semiring (semiring-Ring R) S - is-prop-is-in-subset-Ring : (x : type-Ring R) → is-prop (is-in-subset-Ring x) - is-prop-is-in-subset-Ring = is-prop-is-in-subtype S + is-prop-is-in-subset-Ring : + (x : type-Ring R) → is-prop (is-in-subset-Ring x) + is-prop-is-in-subset-Ring = + is-prop-is-in-subset-Semiring (semiring-Ring R) S - type-subset-Ring : UU (l1 ⊔ l2) - type-subset-Ring = type-subtype S + type-subset-Ring : + UU (l1 ⊔ l2) + type-subset-Ring = + type-subset-Semiring (semiring-Ring R) S - inclusion-subset-Ring : type-subset-Ring → type-Ring R - inclusion-subset-Ring = inclusion-subtype S + inclusion-subset-Ring : + type-subset-Ring → type-Ring R + inclusion-subset-Ring = + inclusion-subset-Semiring (semiring-Ring R) S ap-inclusion-subset-Ring : (x y : type-subset-Ring) → x = y → (inclusion-subset-Ring x = inclusion-subset-Ring y) - ap-inclusion-subset-Ring = ap-inclusion-subtype S + ap-inclusion-subset-Ring = + ap-inclusion-subset-Semiring (semiring-Ring R) S is-in-subset-inclusion-subset-Ring : (x : type-subset-Ring) → is-in-subset-Ring (inclusion-subset-Ring x) @@ -68,12 +81,12 @@ module _ is-closed-under-eq-subset-Ring : {x y : type-Ring R} → is-in-subset-Ring x → (x = y) → is-in-subset-Ring y is-closed-under-eq-subset-Ring = - is-closed-under-eq-subtype S + is-closed-under-eq-subset-Semiring (semiring-Ring R) S is-closed-under-eq-subset-Ring' : {x y : type-Ring R} → is-in-subset-Ring y → (x = y) → is-in-subset-Ring x is-closed-under-eq-subset-Ring' = - is-closed-under-eq-subtype' S + is-closed-under-eq-subset-Semiring' (semiring-Ring R) S ``` ### The condition that a subset contains zero @@ -84,14 +97,16 @@ module _ where contains-zero-subset-Ring : UU l2 - contains-zero-subset-Ring = is-in-subset-Ring R S (zero-Ring R) + contains-zero-subset-Ring = + contains-zero-subset-Semiring (semiring-Ring R) S ``` ### The condition that a subset contains one ```agda contains-one-subset-Ring : UU l2 - contains-one-subset-Ring = is-in-subset-Ring R S (one-Ring R) + contains-one-subset-Ring = + contains-one-subset-Semiring (semiring-Ring R) S ``` ### The condition that a subset is closed under addition @@ -99,9 +114,7 @@ module _ ```agda is-closed-under-addition-subset-Ring : UU (l1 ⊔ l2) is-closed-under-addition-subset-Ring = - {x y : type-Ring R} → - is-in-subset-Ring R S x → is-in-subset-Ring R S y → - is-in-subset-Ring R S (add-Ring R x y) + is-closed-under-addition-subset-Semiring (semiring-Ring R) S ``` ### The condition that a subset is closed under negatives @@ -118,58 +131,57 @@ module _ ```agda is-closed-under-multiplication-subset-Ring : UU (l1 ⊔ l2) is-closed-under-multiplication-subset-Ring = - (x y : type-Ring R) → is-in-subset-Ring R S x → is-in-subset-Ring R S y → - is-in-subset-Ring R S (mul-Ring R x y) + is-closed-under-multiplication-subset-Semiring (semiring-Ring R) S ``` ### The condition that a subset is closed under multiplication from the left by an arbitrary element ```agda - is-closed-under-left-multiplication-subset-Ring-Prop : Prop (l1 ⊔ l2) - is-closed-under-left-multiplication-subset-Ring-Prop = - Π-Prop - ( type-Ring R) - ( λ x → - Π-Prop - ( type-Ring R) - ( λ y → - function-Prop - ( is-in-subset-Ring R S y) - ( S (mul-Ring R x y)))) + is-closed-under-left-multiplication-prop-subset-Ring : Prop (l1 ⊔ l2) + is-closed-under-left-multiplication-prop-subset-Ring = + is-closed-under-left-multiplication-prop-subset-Semiring + ( semiring-Ring R) + ( S) is-closed-under-left-multiplication-subset-Ring : UU (l1 ⊔ l2) is-closed-under-left-multiplication-subset-Ring = - type-Prop is-closed-under-left-multiplication-subset-Ring-Prop + is-closed-under-left-multiplication-subset-Semiring (semiring-Ring R) S is-prop-is-closed-under-left-multiplication-subset-Ring : is-prop is-closed-under-left-multiplication-subset-Ring is-prop-is-closed-under-left-multiplication-subset-Ring = - is-prop-type-Prop is-closed-under-left-multiplication-subset-Ring-Prop + is-prop-is-closed-under-left-multiplication-subset-Semiring + ( semiring-Ring R) + ( S) ``` ### The condition that a subset is closed-under-multiplication from the right by an arbitrary element ```agda - is-closed-under-right-multiplication-subset-Ring-Prop : Prop (l1 ⊔ l2) - is-closed-under-right-multiplication-subset-Ring-Prop = - Π-Prop - ( type-Ring R) - ( λ x → - Π-Prop - ( type-Ring R) - ( λ y → - function-Prop - ( is-in-subset-Ring R S x) - ( S (mul-Ring R x y)))) + is-closed-under-right-multiplication-prop-subset-Ring : Prop (l1 ⊔ l2) + is-closed-under-right-multiplication-prop-subset-Ring = + is-closed-under-right-multiplication-prop-subset-Semiring + ( semiring-Ring R) + ( S) is-closed-under-right-multiplication-subset-Ring : UU (l1 ⊔ l2) is-closed-under-right-multiplication-subset-Ring = - type-Prop is-closed-under-right-multiplication-subset-Ring-Prop + is-closed-under-right-multiplication-subset-Semiring (semiring-Ring R) S is-prop-is-closed-under-right-multiplication-subset-Ring : is-prop is-closed-under-right-multiplication-subset-Ring is-prop-is-closed-under-right-multiplication-subset-Ring = - is-prop-type-Prop is-closed-under-right-multiplication-subset-Ring-Prop + is-prop-is-closed-under-right-multiplication-subset-Semiring + ( semiring-Ring R) + ( S) +``` + +### The condition that a subset is an additive submonoid + +```agda + is-additive-submonoid-subset-Ring : UU (l1 ⊔ l2) + is-additive-submonoid-subset-Ring = + contains-zero-subset-Ring × is-closed-under-addition-subset-Ring ``` ### The condition that a subset is an additive subgroup @@ -181,7 +193,8 @@ module _ is-additive-subgroup-subset-Ring : {l2 : Level} → subset-Ring l2 R → UU (l1 ⊔ l2) - is-additive-subgroup-subset-Ring = is-subgroup-Ab (ab-Ring R) + is-additive-subgroup-subset-Ring = + is-subgroup-Ab (ab-Ring R) is-prop-is-additive-subgroup-subset-Ring : {l2 : Level} (A : subset-Ring l2 R) → @@ -189,3 +202,61 @@ module _ is-prop-is-additive-subgroup-subset-Ring = is-prop-is-subgroup-Ab (ab-Ring R) ``` + +## Properties + +### Any subset that is closed under left or right multiplication is closed under negatives + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) + where + + is-closed-under-negatives-is-closed-under-left-multiplication-subset-Ring : + is-closed-under-left-multiplication-subset-Ring R S → + is-closed-under-negatives-subset-Ring R S + is-closed-under-negatives-is-closed-under-left-multiplication-subset-Ring + H {x} K = + is-closed-under-eq-subset-Ring R S + ( H (neg-one-Ring R) x K) + ( mul-neg-one-Ring R x) + + is-closed-under-negatives-is-closed-under-right-multiplication-subset-Ring : + is-closed-under-right-multiplication-subset-Ring R S → + is-closed-under-negatives-subset-Ring R S + is-closed-under-negatives-is-closed-under-right-multiplication-subset-Ring + H {x} K = + is-closed-under-eq-subset-Ring R S + ( H x (neg-one-Ring R) K) + ( mul-neg-one-Ring' R x) +``` + +### Any additive submonoid of a ring that is closed under left or right multiplication is an additive subgroup + +```agda +module _ + {l1 l2 : Level} (R : Ring l1) (S : subset-Ring l2 R) + where + + is-additive-subgroup-is-closed-under-left-multiplication-subset-Ring : + is-additive-submonoid-subset-Ring R S → + is-closed-under-left-multiplication-subset-Ring R S → + is-additive-subgroup-subset-Ring R S + is-additive-subgroup-is-closed-under-left-multiplication-subset-Ring + (H , K) L = + ( H , + K , + is-closed-under-negatives-is-closed-under-left-multiplication-subset-Ring + R S L) + + is-additive-subgroup-is-closed-under-right-multiplication-subset-Ring : + is-additive-submonoid-subset-Ring R S → + is-closed-under-right-multiplication-subset-Ring R S → + is-additive-subgroup-subset-Ring R S + is-additive-subgroup-is-closed-under-right-multiplication-subset-Ring + (H , K) L = + ( H , + K , + is-closed-under-negatives-is-closed-under-right-multiplication-subset-Ring + R S L) +``` diff --git a/src/ring-theory/subsets-semirings.lagda.md b/src/ring-theory/subsets-semirings.lagda.md index 0236a79809..9c80ebe955 100644 --- a/src/ring-theory/subsets-semirings.lagda.md +++ b/src/ring-theory/subsets-semirings.lagda.md @@ -7,6 +7,8 @@ module ring-theory.subsets-semirings where
Imports ```agda +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types open import foundation.identity-types open import foundation.propositional-extensionality open import foundation.propositions @@ -87,6 +89,11 @@ module _ contains-zero-subset-Semiring : UU l2 contains-zero-subset-Semiring = is-in-subtype S (zero-Semiring R) + + is-prop-contains-zero-subset-Semiring : + is-prop contains-zero-subset-Semiring + is-prop-contains-zero-subset-Semiring = + is-prop-is-in-subtype S (zero-Semiring R) ``` ### The condition that a subset contains one @@ -94,6 +101,11 @@ module _ ```agda contains-one-subset-Semiring : UU l2 contains-one-subset-Semiring = is-in-subtype S (one-Semiring R) + + is-prop-contains-one-subset-Semiring : + is-prop contains-one-subset-Semiring + is-prop-contains-one-subset-Semiring = + is-prop-is-in-subtype S (one-Semiring R) ``` ### The condition that a subset is closed under addition @@ -101,9 +113,19 @@ module _ ```agda is-closed-under-addition-subset-Semiring : UU (l1 ⊔ l2) is-closed-under-addition-subset-Semiring = - (x y : type-Semiring R) → + {x y : type-Semiring R} → is-in-subtype S x → is-in-subtype S y → is-in-subtype S (add-Semiring R x y) + + is-prop-is-closed-under-addition-subset-Semiring : + is-prop is-closed-under-addition-subset-Semiring + is-prop-is-closed-under-addition-subset-Semiring = + is-prop-implicit-Π + ( λ x → + is-prop-implicit-Π + ( λ y → + is-prop-function-type + ( is-prop-function-type (is-prop-is-in-subtype S _)))) ``` ### The condition that a subset is closed under multiplication @@ -113,6 +135,16 @@ module _ is-closed-under-multiplication-subset-Semiring = (x y : type-Semiring R) → is-in-subtype S x → is-in-subtype S y → is-in-subtype S (mul-Semiring R x y) + + is-prop-is-closed-under-multiplication-subset-Semiring : + is-prop is-closed-under-multiplication-subset-Semiring + is-prop-is-closed-under-multiplication-subset-Semiring = + is-prop-Π + ( λ x → + is-prop-Π + ( λ y → + is-prop-function-type + ( is-prop-function-type (is-prop-is-in-subtype S _)))) ``` ### The condition that a subset is closed under multiplication from the left by an arbitrary element @@ -122,13 +154,55 @@ module _ is-closed-under-left-multiplication-subset-Semiring = (x y : type-Semiring R) → is-in-subtype S y → is-in-subtype S (mul-Semiring R x y) + + is-prop-is-closed-under-left-multiplication-subset-Semiring : + is-prop is-closed-under-left-multiplication-subset-Semiring + is-prop-is-closed-under-left-multiplication-subset-Semiring = + is-prop-Π + ( λ x → + is-prop-Π (λ y → is-prop-function-type (is-prop-is-in-subtype S _))) + + is-closed-under-left-multiplication-prop-subset-Semiring : Prop (l1 ⊔ l2) + pr1 is-closed-under-left-multiplication-prop-subset-Semiring = + is-closed-under-left-multiplication-subset-Semiring + pr2 is-closed-under-left-multiplication-prop-subset-Semiring = + is-prop-is-closed-under-left-multiplication-subset-Semiring ``` -### The condition that a subset is closed-under-multiplication from the right by an arbitrary element +### The condition that a subset is closed under multiplication from the right by an arbitrary element ```agda is-closed-under-right-multiplication-subset-Semiring : UU (l1 ⊔ l2) is-closed-under-right-multiplication-subset-Semiring = (x y : type-Semiring R) → is-in-subtype S x → is-in-subtype S (mul-Semiring R x y) + + is-prop-is-closed-under-right-multiplication-subset-Semiring : + is-prop is-closed-under-right-multiplication-subset-Semiring + is-prop-is-closed-under-right-multiplication-subset-Semiring = + is-prop-Π + ( λ x → + is-prop-Π (λ y → is-prop-function-type (is-prop-is-in-subtype S _))) + + is-closed-under-right-multiplication-prop-subset-Semiring : Prop (l1 ⊔ l2) + pr1 is-closed-under-right-multiplication-prop-subset-Semiring = + is-closed-under-right-multiplication-subset-Semiring + pr2 is-closed-under-right-multiplication-prop-subset-Semiring = + is-prop-is-closed-under-right-multiplication-subset-Semiring +``` + +### The condition that a subset of a semiring is an additive submonoid + +```agda + is-additive-submonoid-subset-Semiring : UU (l1 ⊔ l2) + is-additive-submonoid-subset-Semiring = + contains-zero-subset-Semiring × + is-closed-under-addition-subset-Semiring + + is-prop-is-additive-submonoid-subset-Semiring : + is-prop is-additive-submonoid-subset-Semiring + is-prop-is-additive-submonoid-subset-Semiring = + is-prop-product + ( is-prop-contains-zero-subset-Semiring) + ( is-prop-is-closed-under-addition-subset-Semiring) ``` diff --git a/src/structured-types/magmas.lagda.md b/src/structured-types/magmas.lagda.md index 13372059eb..f61e86aaaa 100644 --- a/src/structured-types/magmas.lagda.md +++ b/src/structured-types/magmas.lagda.md @@ -52,13 +52,41 @@ is-unital-Magma M = is-unital (mul-Magma M) Unital-Magma : (l : Level) → UU (lsuc l) Unital-Magma l = Σ (Magma l) is-unital-Magma -magma-Unital-Magma : - {l : Level} → Unital-Magma l → Magma l -magma-Unital-Magma M = pr1 M - -is-unital-magma-Unital-Magma : - {l : Level} (M : Unital-Magma l) → is-unital-Magma (magma-Unital-Magma M) -is-unital-magma-Unital-Magma M = pr2 M +module _ + {l : Level} (M : Unital-Magma l) + where + + magma-Unital-Magma : Magma l + magma-Unital-Magma = pr1 M + + type-Unital-Magma : + UU l + type-Unital-Magma = + type-Magma magma-Unital-Magma + + mul-Unital-Magma : + (x y : type-Unital-Magma) → type-Unital-Magma + mul-Unital-Magma = + mul-Magma magma-Unital-Magma + + is-unital-magma-Unital-Magma : + is-unital-Magma magma-Unital-Magma + is-unital-magma-Unital-Magma = pr2 M + + unit-Unital-Magma : + type-Unital-Magma + unit-Unital-Magma = + pr1 is-unital-magma-Unital-Magma + + left-unit-law-mul-Unital-Magma : + (x : type-Unital-Magma) → mul-Unital-Magma unit-Unital-Magma x = x + left-unit-law-mul-Unital-Magma = + pr1 (pr2 is-unital-magma-Unital-Magma) + + right-unit-law-mul-Unital-Magma : + (x : type-Unital-Magma) → mul-Unital-Magma x unit-Unital-Magma = x + right-unit-law-mul-Unital-Magma = + pr2 (pr2 is-unital-magma-Unital-Magma) ``` ### Semigroups From c839ebe513f7d61aa4c72b94179c872a04e76c24 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 9 Feb 2025 13:56:46 -0500 Subject: [PATCH 68/72] edit --- .../linear-combinations-of-elements-semirings.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md b/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md index c2755085fd..5bcf5562eb 100644 --- a/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md +++ b/src/ring-theory/linear-combinations-of-elements-semirings.lagda.md @@ -54,7 +54,7 @@ module _ list (type-Semiring R × A × type-Semiring R) ``` -### Evaluating linear combinations of elements in magmas +### Evaluating linear combinations of elements in unital magmas ```agda module _ From f2ba05d29d733a13f2e42b24ed0825b171dc71be Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Sun, 9 Feb 2025 20:44:58 -0500 Subject: [PATCH 69/72] a few more merge conflicts --- .../deflationary-maps-preorders.lagda.md | 15 -- ...onary-maps-strictly-ordered-types.lagda.md | 162 ------------------ 2 files changed, 177 deletions(-) delete mode 100644 src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md diff --git a/src/order-theory/deflationary-maps-preorders.lagda.md b/src/order-theory/deflationary-maps-preorders.lagda.md index 7227fe30a8..508b702878 100644 --- a/src/order-theory/deflationary-maps-preorders.lagda.md +++ b/src/order-theory/deflationary-maps-preorders.lagda.md @@ -20,20 +20,6 @@ open import order-theory.preorders ## Idea -<<<<<<< HEAD -A map $f : P → P$ on a [preorder](order-theory.preorders.md) $P$ is said to be an -{{#concept "deflationary map" Disambiguation="preorder" Agda=deflationary-map-Preorder}} if the inequality - -$$ - f(x) \leq x -$$ - -holds for any element $x : P$. If $f$ is also [order preserving](order-theory.order-preserving-maps-preorders.md) we say that $f$ is an {{#concept "deflationary morphism" Disambiguation="preorder" Agda=deflationary-hom-Preorder}}. - -## Definitions - -### The predicate of being an deflationary map -======= A map $f : P → P$ on a [preorder](order-theory.preorders.md) $P$ is said to be a {{#concept "deflationary map" Disambiguation="preorder" Agda=deflationary-map-Preorder}} if the inequality @@ -50,7 +36,6 @@ $f$ is a ## Definitions ### The predicate of being a deflationary map ->>>>>>> a9ee07030da6be910bd6b703862ac69aaeab205f ```agda module _ diff --git a/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md b/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md deleted file mode 100644 index 069d80d2c3..0000000000 --- a/src/order-theory/inflationary-maps-strictly-ordered-types.lagda.md +++ /dev/null @@ -1,162 +0,0 @@ -# Inflationary maps on a strictly ordered type - -```agda -module order-theory.inflationary-maps-strictly-ordered-types where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.propositions -open import foundation.subtypes -open import foundation.universe-levels - -open import order-theory.strict-order-preserving-maps -open import order-theory.strictly-ordered-types -``` - -
- -## Idea - -A map $f : P → P$ on a [strictly ordered type](order-theory.strictly-ordered-types.md) $P$ is said to be an -{{#concept "inflationary map" Disambiguation="strictly ordered type" Agda=inflationary-map-Strictly-Ordered-Type}} if the inequality - -$$ - x < f(x) -$$ - -holds for any element $x : P$. If $f$ is also [order preserving](order-theory.order-preserving-maps-strictly-ordered-types.md) we say that $f$ is an {{#concept "inflationary morphism" Disambiguation="strictly ordered type" Agda=inflationary-hom-Strictly-Ordered-Type}}. - -## Definitions - -### The predicate of being an inflationary map - -```agda -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - (f : type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P) - where - - is-inflationary-prop-map-Strictly-Ordered-Type : - Prop (l1 ⊔ l2) - is-inflationary-prop-map-Strictly-Ordered-Type = - Π-Prop - ( type-Strictly-Ordered-Type P) - ( λ x → le-prop-Strictly-Ordered-Type P x (f x)) - - is-inflationary-map-Strictly-Ordered-Type : - UU (l1 ⊔ l2) - is-inflationary-map-Strictly-Ordered-Type = - type-Prop is-inflationary-prop-map-Strictly-Ordered-Type - - is-prop-is-inflationary-map-Strictly-Ordered-Type : - is-prop is-inflationary-map-Strictly-Ordered-Type - is-prop-is-inflationary-map-Strictly-Ordered-Type = - is-prop-type-Prop is-inflationary-prop-map-Strictly-Ordered-Type -``` - -### The type of inflationary maps on a strictly ordered type - -```agda -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - where - - inflationary-map-Strictly-Ordered-Type : - UU (l1 ⊔ l2) - inflationary-map-Strictly-Ordered-Type = - type-subtype (is-inflationary-prop-map-Strictly-Ordered-Type P) - -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - (f : inflationary-map-Strictly-Ordered-Type P) - where - - map-inflationary-map-Strictly-Ordered-Type : - type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P - map-inflationary-map-Strictly-Ordered-Type = - pr1 f - - is-inflationary-inflationary-map-Strictly-Ordered-Type : - is-inflationary-map-Strictly-Ordered-Type P - ( map-inflationary-map-Strictly-Ordered-Type) - is-inflationary-inflationary-map-Strictly-Ordered-Type = - pr2 f -``` - -### The predicate on order preserving maps of being inflationary - -```agda -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - (f : hom-Strictly-Ordered-Type P P) - where - - is-inflationary-prop-hom-Strictly-Ordered-Type : - Prop (l1 ⊔ l2) - is-inflationary-prop-hom-Strictly-Ordered-Type = - is-inflationary-prop-map-Strictly-Ordered-Type P - ( map-hom-Strictly-Ordered-Type P P f) - - is-inflationary-hom-Strictly-Ordered-Type : - UU (l1 ⊔ l2) - is-inflationary-hom-Strictly-Ordered-Type = - is-inflationary-map-Strictly-Ordered-Type P - ( map-hom-Strictly-Ordered-Type P P f) - - is-prop-is-inflationary-hom-Strictly-Ordered-Type : - is-prop is-inflationary-hom-Strictly-Ordered-Type - is-prop-is-inflationary-hom-Strictly-Ordered-Type = - is-prop-is-inflationary-map-Strictly-Ordered-Type P - ( map-hom-Strictly-Ordered-Type P P f) -``` - -### The type of inflationary morphisms on a strictly ordered type - -```agda -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - where - - inflationary-hom-Strictly-Ordered-Type : - UU (l1 ⊔ l2) - inflationary-hom-Strictly-Ordered-Type = - type-subtype (is-inflationary-prop-hom-Strictly-Ordered-Type P) - -module _ - {l1 l2 : Level} (P : Strictly-Ordered-Type l1 l2) - (f : inflationary-hom-Strictly-Ordered-Type P) - where - - hom-inflationary-hom-Strictly-Ordered-Type : - hom-Strictly-Ordered-Type P P - hom-inflationary-hom-Strictly-Ordered-Type = - pr1 f - - map-inflationary-hom-Strictly-Ordered-Type : - type-Strictly-Ordered-Type P → type-Strictly-Ordered-Type P - map-inflationary-hom-Strictly-Ordered-Type = - map-hom-Strictly-Ordered-Type P P hom-inflationary-hom-Strictly-Ordered-Type - - preserves-order-inflationary-hom-Strictly-Ordered-Type : - preserves-strict-order-map-Strictly-Ordered-Type P P - ( map-inflationary-hom-Strictly-Ordered-Type) - preserves-order-inflationary-hom-Strictly-Ordered-Type = - preserves-strict-order-hom-Strictly-Ordered-Type P P - ( hom-inflationary-hom-Strictly-Ordered-Type) - - is-inflationary-inflationary-hom-Strictly-Ordered-Type : - is-inflationary-map-Strictly-Ordered-Type P - ( map-inflationary-hom-Strictly-Ordered-Type) - is-inflationary-inflationary-hom-Strictly-Ordered-Type = - pr2 f - - inflationary-map-inflationary-hom-Strictly-Ordered-Type : - inflationary-map-Strictly-Ordered-Type P - pr1 inflationary-map-inflationary-hom-Strictly-Ordered-Type = - map-inflationary-hom-Strictly-Ordered-Type - pr2 inflationary-map-inflationary-hom-Strictly-Ordered-Type = - is-inflationary-inflationary-hom-Strictly-Ordered-Type -``` From 7c9b3cc938dfcccef899cf621599164e9a561279 Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 11 Feb 2025 11:23:02 -0500 Subject: [PATCH 70/72] edit? --- .../strict-inequality-natural-numbers.lagda.md | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md index 9dfa404a25..dd3b79d820 100644 --- a/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strict-inequality-natural-numbers.lagda.md @@ -30,7 +30,7 @@ open import foundation.transport-along-identifications open import foundation.unit-type open import foundation.universe-levels -open import order-theory.strictly-ordered-types +open import order-theory.strict-preorders ```
@@ -202,14 +202,14 @@ decide-le-neq-ℕ (succ-ℕ m) (succ-ℕ n) H = decide-le-neq-ℕ m n (λ p → H (ap succ-ℕ p)) ``` -### The strictly ordered type of natural numbers +### The strict preorder of natural numbers ```agda -ℕ-Strictly-Ordered-Type : Strictly-Ordered-Type lzero lzero -pr1 ℕ-Strictly-Ordered-Type = ℕ -pr1 (pr2 ℕ-Strictly-Ordered-Type) = le-ℕ-Prop -pr1 (pr2 (pr2 ℕ-Strictly-Ordered-Type)) = irreflexive-le-ℕ -pr2 (pr2 (pr2 ℕ-Strictly-Ordered-Type)) = transitive-le-ℕ +ℕ-Strict-Preorder : Strict-Preorder lzero lzero +pr1 ℕ-Strict-Preorder = ℕ +pr1 (pr2 ℕ-Strict-Preorder) = le-ℕ-Prop +pr1 (pr2 (pr2 ℕ-Strict-Preorder)) = irreflexive-le-ℕ +pr2 (pr2 (pr2 ℕ-Strict-Preorder)) = transitive-le-ℕ ``` ### `n < m` if and only if there exists a nonzero natural number `l` such that `l + n = m` From c83c8f3b9eb109edf2970969b7758dc48fb6b2bb Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 11 Feb 2025 12:13:02 -0500 Subject: [PATCH 71/72] products of elements of lists --- src/elementary-number-theory.lagda.md | 2 +- .../exponentiation-natural-numbers.lagda.md | 6 +-- ...fundamental-theorem-of-arithmetic.lagda.md | 8 +-- ...ve-decompositions-natural-numbers.lagda.md | 9 ++-- ...roducts-lists-of-natural-numbers.lagda.md} | 40 +++++++-------- .../products-of-natural-numbers.lagda.md | 18 ++++--- ...tionary-functions-natural-numbers.lagda.md | 11 ++-- ...ermutations-standard-finite-types.lagda.md | 50 +++++++++---------- ...nspositions-standard-finite-types.lagda.md | 8 +-- .../iterated-cartesian-product-types.lagda.md | 18 +++---- src/lists/permutation-lists.lagda.md | 36 ++++++------- src/lists/permutation-vectors.lagda.md | 28 +++++------ src/lists/sort-by-insertion-lists.lagda.md | 2 +- src/lists/sort-by-insertion-vectors.lagda.md | 4 +- src/lists/sorted-vectors.lagda.md | 4 +- src/lists/sorting-algorithms-lists.lagda.md | 2 +- src/lists/sorting-algorithms-vectors.lagda.md | 2 +- .../universal-quantification-lists.lagda.md | 24 ++++++--- ...nflationary-maps-strict-preorders.lagda.md | 4 +- ...ime-decomposition-natural-numbers.lagda.md | 2 +- 20 files changed, 144 insertions(+), 134 deletions(-) rename src/elementary-number-theory/{multiplication-lists-of-natural-numbers.lagda.md => products-lists-of-natural-numbers.lagda.md} (60%) diff --git a/src/elementary-number-theory.lagda.md b/src/elementary-number-theory.lagda.md index 1cf5ab00e6..43a8d7dd47 100644 --- a/src/elementary-number-theory.lagda.md +++ b/src/elementary-number-theory.lagda.md @@ -106,7 +106,7 @@ open import elementary-number-theory.monoid-of-natural-numbers-with-addition pub open import elementary-number-theory.monoid-of-natural-numbers-with-maximum public open import elementary-number-theory.multiplication-integer-fractions public open import elementary-number-theory.multiplication-integers public -open import elementary-number-theory.multiplication-lists-of-natural-numbers public +open import elementary-number-theory.products-lists-of-natural-numbers public open import elementary-number-theory.multiplication-natural-numbers public open import elementary-number-theory.multiplication-positive-and-negative-integers public open import elementary-number-theory.multiplication-rational-numbers public diff --git a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md index a7aaba635d..25650a3b7a 100644 --- a/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md +++ b/src/elementary-number-theory/exponentiation-natural-numbers.lagda.md @@ -345,9 +345,9 @@ preserves-order-exponent-exp-ℕ m H (succ-ℕ n) (succ-ℕ k) K = ```agda preserves-strict-order-exponent-exp-ℕ : (m : ℕ) → 1 <-ℕ m → - preserves-strict-order-map-Strictly-Ordered-Type - ( ℕ-Strictly-Ordered-Type) - ( ℕ-Strictly-Ordered-Type) + preserves-strict-order-map-Strict-Preorder + ( ℕ-Strict-Preorder) + ( ℕ-Strict-Preorder) ( m ^ℕ_) preserves-strict-order-exponent-exp-ℕ m H zero-ℕ (succ-ℕ k) K = le-one-exp-ℕ m (succ-ℕ k) H (is-nonzero-succ-ℕ k) diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index b2b7578234..6bd84555bc 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -17,7 +17,7 @@ open import elementary-number-theory.lists-of-prime-numbers open import elementary-number-theory.lower-bounds-natural-numbers open import elementary-number-theory.minimal-structured-natural-numbers open import elementary-number-theory.modular-arithmetic-standard-finite-types -open import elementary-number-theory.multiplication-lists-of-natural-numbers +open import elementary-number-theory.products-lists-of-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.multiplicative-decompositions-natural-numbers open import elementary-number-theory.natural-numbers @@ -778,7 +778,7 @@ is-prime-list-is-prime-is-in-list-ℕ (cons x p) P = ( is-prime-is-in-list-tail-ℕ p x P) is-prime-list-permute-list-ℕ : - (p : list ℕ) (t : Permutation (length-list p)) → is-prime-list-ℕ p → + (p : list ℕ) (t : permutation (length-list p)) → is-prime-list-ℕ p → is-prime-list-ℕ (permute-list p t) is-prime-list-permute-list-ℕ p t P = is-prime-list-is-prime-is-in-list-ℕ @@ -794,11 +794,11 @@ is-prime-list-permute-list-ℕ p t P = ( I))) is-decomposition-list-permute-list-ℕ : - (n : ℕ) (p : list ℕ) (t : Permutation (length-list p)) → + (n : ℕ) (p : list ℕ) (t : permutation (length-list p)) → is-multiplicative-decomposition-list-ℕ n p → is-multiplicative-decomposition-list-ℕ n (permute-list p t) is-decomposition-list-permute-list-ℕ n p t D = - {!!} -- inv (invariant-permutation-mul-list-ℕ p t) ∙ D + {!!} -- inv (permutation-invariant-mul-list-ℕ p t) ∙ D is-prime-decomposition-list-sort-concatenation-ℕ : (x y : ℕ) (H : 1 ≤-ℕ x) (I : 1 ≤-ℕ y) (p q : list ℕ) → diff --git a/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md index 487e0c5afe..c42eb9f057 100644 --- a/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/multiplicative-decompositions-natural-numbers.lagda.md @@ -11,7 +11,7 @@ module ```agda open import elementary-number-theory.divisibility-natural-numbers open import elementary-number-theory.equality-natural-numbers -open import elementary-number-theory.multiplication-lists-of-natural-numbers +open import elementary-number-theory.products-lists-of-natural-numbers open import elementary-number-theory.multiplication-natural-numbers open import elementary-number-theory.natural-numbers open import elementary-number-theory.nontrivial-divisors-natural-numbers @@ -40,9 +40,8 @@ of a [natural number](elementary-number-theory.natural-numbers.md) `n` is a [list](lists.lists.md) `l` of natural numbers [strictly greater than](elementary-number-theory.strict-inequality-natural-numbers.md) `1`, such that the -[product](elementary-number-theory.multiplication-lists-of-natural-numbers.md) -of the numbers in this list is [equal to](foundation-core.identity-types.md) -`n`. +[product](elementary-number-theory.products-lists-of-natural-numbers.md) of the +numbers in this list is [equal to](foundation-core.identity-types.md) `n`. ## Definitions @@ -184,7 +183,7 @@ pr1 (is-multiplicative-decomposition-concat-list-ℕ n m l k H K) = ( le-one-list-is-multiplicative-decomposition-list-ℕ n l H) ( le-one-list-is-multiplicative-decomposition-list-ℕ m k K) pr2 (is-multiplicative-decomposition-concat-list-ℕ n m l k H K) = - eq-mul-list-concat-list-ℕ l k ∙ + compute-mul-concat-list-ℕ l k ∙ ap-mul-ℕ ( eq-is-multiplicative-decomposition-list-ℕ n l H) ( eq-is-multiplicative-decomposition-list-ℕ m k K) diff --git a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md b/src/elementary-number-theory/products-lists-of-natural-numbers.lagda.md similarity index 60% rename from src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md rename to src/elementary-number-theory/products-lists-of-natural-numbers.lagda.md index 6ccbe97b88..2f4e7f96e6 100644 --- a/src/elementary-number-theory/multiplication-lists-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/products-lists-of-natural-numbers.lagda.md @@ -1,7 +1,7 @@ -# Multiplication of the elements of a list of natural numbers +# Products of lists of natural numbers ```agda -module elementary-number-theory.multiplication-lists-of-natural-numbers where +module elementary-number-theory.products-lists-of-natural-numbers where ```
Imports @@ -28,10 +28,14 @@ open import lists.universal-quantification-lists ## Idea -Given a list of natural number `l`, we define the product of the element of the -list. +Given a [list](lists.lists.md) of +[natural numbers](elementary-number-theory.natural-numbers.md) $l$, we define +the {{#concept "product" Disambiguation="list of natural numbers"}} of the +elements of $l$ by recursively multiplying all its elements. -## Definition +## Definitions + +### The product of the elements of a list ```agda mul-list-ℕ : @@ -51,30 +55,24 @@ leq-one-mul-list-ℕ (cons n l) (H , K) = preserves-order-mul-ℕ 1 n 1 (mul-list-ℕ l) H (leq-one-mul-list-ℕ l K) ``` -### `mul-list-ℕ` is invariant by permutation +### `mul-list-ℕ` is permutation invariant ```agda -invariant-permutation-mul-list-ℕ : - (l : list ℕ) (t : Permutation (length-list l)) → +permutation-invariant-mul-list-ℕ : + (l : list ℕ) (t : permutation (length-list l)) → mul-list-ℕ l = mul-list-ℕ (permute-list l t) -invariant-permutation-mul-list-ℕ = - invariant-permutation-fold-list - ( 1) - ( mul-ℕ) - ( λ a1 a2 b → - ( inv (associative-mul-ℕ a1 a2 b) ∙ - ( ap (λ n → n *ℕ b) (commutative-mul-ℕ a1 a2) ∙ - ( associative-mul-ℕ a2 a1 b)))) +permutation-invariant-mul-list-ℕ = + permutation-invariant-fold-list 1 mul-ℕ left-swap-mul-ℕ ``` ### `mul-list-ℕ` of a concatenation of lists ```agda -eq-mul-list-concat-list-ℕ : +compute-mul-concat-list-ℕ : (p q : list ℕ) → - (mul-list-ℕ (concat-list p q)) = (mul-list-ℕ p) *ℕ (mul-list-ℕ q) -eq-mul-list-concat-list-ℕ nil q = inv (left-unit-law-add-ℕ (mul-list-ℕ q)) -eq-mul-list-concat-list-ℕ (cons x p) q = - ap (mul-ℕ x) (eq-mul-list-concat-list-ℕ p q) ∙ + mul-list-ℕ (concat-list p q) = mul-list-ℕ p *ℕ mul-list-ℕ q +compute-mul-concat-list-ℕ nil q = inv (left-unit-law-add-ℕ (mul-list-ℕ q)) +compute-mul-concat-list-ℕ (cons x p) q = + ap (mul-ℕ x) (compute-mul-concat-list-ℕ p q) ∙ inv (associative-mul-ℕ x (mul-list-ℕ p) (mul-list-ℕ q)) ``` diff --git a/src/elementary-number-theory/products-of-natural-numbers.lagda.md b/src/elementary-number-theory/products-of-natural-numbers.lagda.md index 145ab2e7e7..9254b56aa4 100644 --- a/src/elementary-number-theory/products-of-natural-numbers.lagda.md +++ b/src/elementary-number-theory/products-of-natural-numbers.lagda.md @@ -31,17 +31,15 @@ open import univalent-combinatorics.standard-finite-types ## Idea -The product of a list of natural numbers is defined by iterated multiplication. +The +{{#concept "product" Disambiguation="standard finite family of natural numbers"}} +of a [standard finite](univalent-combinatorics.standard-finite-types.md) family +of [natural numbers](elementary-number-theory.natural-numbers.md) is defined by +iterated +[multiplication](elementary-number-theory.multiplication-natural-numbers.md). ## Definitions -### Products of lists of natural numbers - -```agda -product-list-ℕ : list ℕ → ℕ -product-list-ℕ = fold-list 1 mul-ℕ -``` - ### Products of families of natural numbers indexed by standard finite types ```agda @@ -117,3 +115,7 @@ pr1 (div-factor-Π-ℕ k f i) = pr2 (div-factor-Π-ℕ k f i) = eq-quotient-div-factor-Π-ℕ k f i ``` + +## See also + +- [Products of lists of natural numbers](elementary-number-theory.products-lists-of-natural-numbers.md) diff --git a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md index 729df17246..e84aabf481 100644 --- a/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md +++ b/src/elementary-number-theory/strictly-inflationary-functions-natural-numbers.lagda.md @@ -22,7 +22,7 @@ open import foundation.fibers-of-maps open import foundation.identity-types open import foundation.universe-levels -open import order-theory.inflationary-maps-strictly-ordered-types +open import order-theory.strictly-inflationary-maps-strict-preorders open import univalent-combinatorics.dependent-pair-types ``` @@ -50,7 +50,7 @@ there is a maximal natural number $k$ such that $f(k) \leq b$. ```agda is-strictly-inflationary-map-ℕ : (ℕ → ℕ) → UU lzero is-strictly-inflationary-map-ℕ = - is-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + is-strictly-inflationary-map-Strict-Preorder ℕ-Strict-Preorder ``` ### Strictly inflationary maps on the natural numbers @@ -58,18 +58,19 @@ is-strictly-inflationary-map-ℕ = ```agda strictly-inflationary-map-ℕ : UU lzero strictly-inflationary-map-ℕ = - inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + strictly-inflationary-map-Strict-Preorder ℕ-Strict-Preorder map-strictly-inflationary-map-ℕ : strictly-inflationary-map-ℕ → ℕ → ℕ map-strictly-inflationary-map-ℕ = - map-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + map-strictly-inflationary-map-Strict-Preorder ℕ-Strict-Preorder is-strictly-inflationary-strictly-inflationary-map-ℕ : (f : strictly-inflationary-map-ℕ) → is-strictly-inflationary-map-ℕ (map-strictly-inflationary-map-ℕ f) is-strictly-inflationary-strictly-inflationary-map-ℕ = - is-inflationary-inflationary-map-Strictly-Ordered-Type ℕ-Strictly-Ordered-Type + is-strictly-inflationary-strictly-inflationary-map-Strict-Preorder + ℕ-Strict-Preorder fiber-strictly-inflationary-map-ℕ : (f : strictly-inflationary-map-ℕ) → ℕ → UU lzero diff --git a/src/finite-group-theory/permutations-standard-finite-types.lagda.md b/src/finite-group-theory/permutations-standard-finite-types.lagda.md index 38918a9857..c2d29584ed 100644 --- a/src/finite-group-theory/permutations-standard-finite-types.lagda.md +++ b/src/finite-group-theory/permutations-standard-finite-types.lagda.md @@ -48,13 +48,13 @@ open import univalent-combinatorics.standard-finite-types ## Idea -A permutation of `Fin n` is an automorphism of `Fin n`. +A {{#concept "permutation" Disambiguation="standard finite type"}} on a [standard finite type](univalent-combinatorics.standard-finite-types.md) `Fin n` is an [automorphism](foundation.automorphisms.md) of `Fin n`. ## Definitions ```agda -Permutation : (n : ℕ) → UU lzero -Permutation n = Aut (Fin n) +permutation : (n : ℕ) → UU lzero +permutation n = Aut (Fin n) ``` ## Properties @@ -63,14 +63,12 @@ Permutation n = Aut (Fin n) ```agda list-transpositions-permutation-Fin' : - (n : ℕ) (f : Permutation (succ-ℕ n)) → + (n : ℕ) (f : permutation (succ-ℕ n)) → (x : Fin (succ-ℕ n)) → Id (map-equiv f (inr star)) x → - ( list - ( Σ - ( Fin (succ-ℕ n) → Decidable-Prop lzero) - ( λ P → - has-cardinality 2 - ( Σ (Fin (succ-ℕ n)) (type-Decidable-Prop ∘ P))))) + list + ( Σ ( Fin (succ-ℕ n) → Decidable-Prop lzero) + ( λ P → + has-cardinality 2 (Σ (Fin (succ-ℕ n)) (type-Decidable-Prop ∘ P)))) list-transpositions-permutation-Fin' zero-ℕ f x p = nil list-transpositions-permutation-Fin' (succ-ℕ n) f (inl x) p = cons @@ -93,7 +91,7 @@ list-transpositions-permutation-Fin' (succ-ℕ n) f (inl x) p = { inr star} { inl x} ( neq-inr-inl) - f' : (Permutation (succ-ℕ n)) + f' : permutation (succ-ℕ n) f' = map-inv-equiv ( extend-equiv-Maybe (Fin-Set (succ-ℕ n))) @@ -110,11 +108,11 @@ list-transpositions-permutation-Fin' (succ-ℕ n) f (inr star) p = ( Fin-succ-Fin-transposition (succ-ℕ n)) ( list-transpositions-permutation-Fin' n f' (map-equiv f' (inr star)) refl) where - f' : (Permutation (succ-ℕ n)) + f' : permutation (succ-ℕ n) f' = map-inv-equiv (extend-equiv-Maybe (Fin-Set (succ-ℕ n))) (pair f p) list-transpositions-permutation-Fin : - (n : ℕ) (f : Permutation n) → + (n : ℕ) (f : permutation n) → ( list ( Σ ( Fin n → Decidable-Prop lzero) @@ -125,7 +123,7 @@ list-transpositions-permutation-Fin (succ-ℕ n) f = abstract retraction-permutation-list-transpositions-Fin' : - (n : ℕ) (f : Permutation (succ-ℕ n)) → + (n : ℕ) (f : permutation (succ-ℕ n)) → (x : Fin (succ-ℕ n)) → Id (map-equiv f (inr star)) x → (y z : Fin (succ-ℕ n)) → Id (map-equiv f y) z → Id @@ -189,7 +187,7 @@ abstract { inl x} ( neq-inr-inl) P : - Σ ( Permutation (succ-ℕ (succ-ℕ n))) + Σ ( permutation (succ-ℕ (succ-ℕ n))) ( λ g → Id (map-equiv g (inr star)) (inr star)) P = pair @@ -200,9 +198,9 @@ abstract { inr star} { inl x} ( neq-inr-inl))) - F' : (Permutation (succ-ℕ n)) + F' : permutation (succ-ℕ n) F' = map-inv-equiv (extend-equiv-Maybe (Fin-Set (succ-ℕ n))) P - lemma2 : (map-equiv (transposition t) (inl z)) = (inl z) + lemma2 : map-equiv (transposition t) (inl z) = inl z lemma2 = is-fixed-point-standard-transposition ( has-decidable-equality-Fin (succ-ℕ (succ-ℕ n))) @@ -283,7 +281,7 @@ abstract { inl x} ( neq-inr-inl) P : - Σ ( Permutation (succ-ℕ (succ-ℕ n))) + Σ ( permutation (succ-ℕ (succ-ℕ n))) ( λ g → Id (map-equiv g (inr star)) (inr star)) P = pair ( transposition t ∘e f) @@ -293,7 +291,7 @@ abstract { inr star} { inl x} ( neq-inr-inl)) - F' : (Permutation (succ-ℕ n)) + F' : permutation (succ-ℕ n) F' = map-inv-equiv (extend-equiv-Maybe (Fin-Set (succ-ℕ n))) P lemma : Id @@ -358,7 +356,7 @@ abstract { inr star} { inl x} ( neq-inr-inl) - F' : (Permutation (succ-ℕ n)) + F' : permutation (succ-ℕ n) F' = map-inv-equiv ( extend-equiv-Maybe (Fin-Set (succ-ℕ n))) @@ -398,7 +396,7 @@ abstract n f' (map-equiv f' (inr star)) refl y (map-equiv f' y) refl)) ∙ ( computation-inv-extend-equiv-Maybe (Fin-Set (succ-ℕ n)) f p y))) where - f' : (Permutation (succ-ℕ n)) + f' : permutation (succ-ℕ n) f' = map-inv-equiv (extend-equiv-Maybe (Fin-Set (succ-ℕ n))) (pair f p) retraction-permutation-list-transpositions-Fin' (succ-ℕ n) f (inr star) p (inl y) (inr star) q = @@ -428,11 +426,11 @@ abstract ( inr star)) ∙ ( inv p)) where - f' : (Permutation (succ-ℕ n)) + f' : permutation (succ-ℕ n) f' = map-inv-equiv (extend-equiv-Maybe (Fin-Set (succ-ℕ n))) (pair f p) retraction-permutation-list-transpositions-Fin : - (n : ℕ) (f : Permutation n) → + (n : ℕ) (f : permutation n) → htpy-equiv ( permutation-list-transpositions ( list-transpositions-permutation-Fin n f)) @@ -447,7 +445,7 @@ abstract permutation-list-standard-transpositions-Fin : (n : ℕ) → list (Σ (Fin n × Fin n) (λ (i , j) → i ≠ j)) → - Permutation n + permutation n permutation-list-standard-transpositions-Fin n = fold-list ( id-equiv) @@ -455,7 +453,7 @@ permutation-list-standard-transpositions-Fin n = standard-transposition (has-decidable-equality-Fin n) neq ∘e p) list-standard-transpositions-permutation-Fin : - (n : ℕ) (f : Permutation n) → + (n : ℕ) (f : permutation n) → list (Σ (Fin n × Fin n) (λ (i , j) → i ≠ j)) list-standard-transpositions-permutation-Fin n f = map-list @@ -496,7 +494,7 @@ private htpy-permutation-list n l) retraction-permutation-list-standard-transpositions-Fin : - (n : ℕ) (f : Permutation n) → + (n : ℕ) (f : permutation n) → htpy-equiv ( permutation-list-standard-transpositions-Fin ( n) diff --git a/src/finite-group-theory/transpositions-standard-finite-types.lagda.md b/src/finite-group-theory/transpositions-standard-finite-types.lagda.md index bd345f2692..cbe388472a 100644 --- a/src/finite-group-theory/transpositions-standard-finite-types.lagda.md +++ b/src/finite-group-theory/transpositions-standard-finite-types.lagda.md @@ -57,7 +57,7 @@ module _ (n : ℕ) (i j : Fin n) (neq : i ≠ j) where - transposition-Fin : Permutation n + transposition-Fin : permutation n transposition-Fin = standard-transposition (has-decidable-equality-Fin n) neq map-transposition-Fin : Fin n → Fin n @@ -109,7 +109,7 @@ module _ refl is-involution-map-swap-two-last-elements-transposition-Fin (inr star) = refl - swap-two-last-elements-transposition-Fin : Permutation (succ-ℕ (succ-ℕ n)) + swap-two-last-elements-transposition-Fin : permutation (succ-ℕ (succ-ℕ n)) pr1 swap-two-last-elements-transposition-Fin = map-swap-two-last-elements-transposition-Fin pr2 swap-two-last-elements-transposition-Fin = @@ -197,7 +197,7 @@ We show that this definiton is an instance of the previous one. ```agda adjacent-transposition-Fin : (n : ℕ) → (k : Fin n) → - Permutation (succ-ℕ n) + permutation (succ-ℕ n) adjacent-transposition-Fin (succ-ℕ n) (inl x) = equiv-coproduct (adjacent-transposition-Fin n x) id-equiv adjacent-transposition-Fin (succ-ℕ n) (inr x) = @@ -567,7 +567,7 @@ list-adjacent-transpositions-transposition-Fin ( inl (inr star)) = cons (inr star) nil permutation-list-adjacent-transpositions : - (n : ℕ) → list (Fin n) → Permutation (succ-ℕ n) + (n : ℕ) → list (Fin n) → permutation (succ-ℕ n) permutation-list-adjacent-transpositions n nil = id-equiv permutation-list-adjacent-transpositions n (cons x l) = adjacent-transposition-Fin n x ∘e diff --git a/src/foundation/iterated-cartesian-product-types.lagda.md b/src/foundation/iterated-cartesian-product-types.lagda.md index ab1849056b..9b5542ae28 100644 --- a/src/foundation/iterated-cartesian-product-types.lagda.md +++ b/src/foundation/iterated-cartesian-product-types.lagda.md @@ -144,29 +144,29 @@ equiv-product-iterated-product-lists (cons x p) q = ```agda permutation-iterated-product-Fin-Π : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → UU l + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → UU l permutation-iterated-product-Fin-Π n A t = iterated-product-Fin-Π n (A ∘ map-equiv t) equiv-permutation-iterated-product-Fin-Π : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → permutation-iterated-product-Fin-Π n A t ≃ iterated-product-Fin-Π n A equiv-permutation-iterated-product-Fin-Π n A t = equiv-Π (λ z → A z) t (λ a → id-equiv) eq-permutation-iterated-product-Fin-Π : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → permutation-iterated-product-Fin-Π n A t = iterated-product-Fin-Π n A eq-permutation-iterated-product-Fin-Π n A t = eq-equiv (equiv-permutation-iterated-product-Fin-Π n A t) permutation-iterated-product-Fin-recursive : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → UU l + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → UU l permutation-iterated-product-Fin-recursive n A t = iterated-product-Fin-recursive n (A ∘ map-equiv t) equiv-permutation-iterated-product-Fin-recursive : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → permutation-iterated-product-Fin-recursive n A t ≃ iterated-product-Fin-recursive n A equiv-permutation-iterated-product-Fin-recursive n A t = @@ -175,19 +175,19 @@ equiv-permutation-iterated-product-Fin-recursive n A t = ( equiv-iterated-product-Fin-recursive-Π n (A ∘ map-equiv t)) eq-permutation-iterated-product-Fin-recursive : - {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : Permutation n) → + {l : Level} (n : ℕ) (A : (Fin n → UU l)) (t : permutation n) → permutation-iterated-product-Fin-recursive n A t = iterated-product-Fin-recursive n A eq-permutation-iterated-product-Fin-recursive n A t = eq-equiv (equiv-permutation-iterated-product-Fin-recursive n A t) permutation-iterated-product-lists : - {l : Level} (L : list (UU l)) (t : Permutation (length-list L)) → UU l + {l : Level} (L : list (UU l)) (t : permutation (length-list L)) → UU l permutation-iterated-product-lists L t = iterated-product-lists (permute-list L t) equiv-permutation-iterated-product-lists : - {l : Level} (L : list (UU l)) (t : Permutation (length-list L)) → + {l : Level} (L : list (UU l)) (t : permutation (length-list L)) → permutation-iterated-product-lists L t ≃ iterated-product-lists L equiv-permutation-iterated-product-lists L t = @@ -209,7 +209,7 @@ equiv-permutation-iterated-product-lists L t = ( equiv-iterated-product-Fin-recursive-lists (permute-list L t)))))) eq-permutation-iterated-product-lists : - {l : Level} (L : list (UU l)) (t : Permutation (length-list L)) → + {l : Level} (L : list (UU l)) (t : permutation (length-list L)) → permutation-iterated-product-lists L t = iterated-product-lists L eq-permutation-iterated-product-lists L t = diff --git a/src/lists/permutation-lists.lagda.md b/src/lists/permutation-lists.lagda.md index 857ca511c1..5c7812e703 100644 --- a/src/lists/permutation-lists.lagda.md +++ b/src/lists/permutation-lists.lagda.md @@ -46,7 +46,7 @@ module _ {l : Level} {A : UU l} where - permute-list : (l : list A) → Permutation (length-list l) → list A + permute-list : (l : list A) → permutation (length-list l) → list A permute-list l s = list-vec ( length-list l) @@ -59,12 +59,12 @@ module _ is-permutation-list : (list A → list A) → UU l is-permutation-list f = (l : list A) → - Σ ( Permutation (length-list l)) + Σ ( permutation (length-list l)) ( λ t → f l = permute-list l t) permutation-is-permutation-list : (f : list A → list A) → is-permutation-list f → - ((l : list A) → Permutation (length-list l)) + ((l : list A) → permutation (length-list l)) permutation-is-permutation-list f P l = pr1 (P l) eq-permute-list-permutation-is-permutation-list : @@ -80,7 +80,7 @@ module _ ```agda length-permute-list : - (l : list A) (t : Permutation (length-list l)) → + (l : list A) (t : permutation (length-list l)) → length-list (permute-list l t) = (length-list l) length-permute-list l t = compute-length-list-list-vec @@ -92,7 +92,7 @@ module _ ```agda eq-vec-list-permute-list : - (l : list A) (f : Permutation (length-list l)) → + (l : list A) (f : permutation (length-list l)) → permute-vec (length-list l) (vec-list l) f = tr ( vec A) @@ -124,7 +124,7 @@ module _ ```agda is-in-list-is-in-permute-list : - (l : list A) (t : Permutation (length-list l)) (x : A) → + (l : list A) (t : permutation (length-list l)) (x : A) → x ∈-list (permute-list l t) → x ∈-list l is-in-list-is-in-permute-list l t x I = is-in-list-is-in-vec-list @@ -148,7 +148,7 @@ module _ ( I)))) is-in-permute-list-is-in-list : - (l : list A) (t : Permutation (length-list l)) (x : A) → + (l : list A) (t : permutation (length-list l)) (x : A) → x ∈-list l → x ∈-list (permute-list l t) is-in-permute-list-is-in-list l t x I = is-in-list-is-in-vec-list @@ -182,12 +182,12 @@ module _ fold-vec b μ (tr (vec A) eq v) = fold-vec b μ v invariant-fold-vec-tr v refl = refl - invariant-permutation-fold-list : - (l : list A) → (f : Permutation (length-list l)) → + permutation-invariant-fold-list : + (l : list A) → (f : permutation (length-list l)) → fold-list b μ l = fold-list b μ (permute-list l f) - invariant-permutation-fold-list l f = + permutation-invariant-fold-list l f = ( inv (htpy-fold-list-fold-vec b μ l) ∙ - ( invariant-permutation-fold-vec b μ C (vec-list l) f ∙ + ( permutation-invariant-fold-vec b μ C (vec-list l) f ∙ ( ap (fold-vec b μ) (eq-vec-list-permute-list l f) ∙ ( ( invariant-fold-vec-tr { m = length-list l} @@ -201,14 +201,14 @@ module _ ```agda compute-tr-permute-vec : {l : Level} {A : UU l} {n m : ℕ} - (e : n = m) (v : vec A n) (t : Permutation m) → + (e : n = m) (v : vec A n) (t : permutation m) → tr ( vec A) ( e) ( permute-vec ( n) ( v) - ( tr Permutation (inv e) t)) = + ( tr permutation (inv e) t)) = permute-vec ( m) ( tr (vec A) e v) @@ -223,7 +223,7 @@ compute-tr-map-vec f refl v = refl helper-compute-list-vec-map-vec-permute-vec-vec-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} - (f : A → B) (p : list A) (t : Permutation (length-list p)) → + (f : A → B) (p : list A) (t : permutation (length-list p)) → tr ( vec B) ( inv (length-permute-list p t)) @@ -251,7 +251,7 @@ helper-compute-list-vec-map-vec-permute-vec-vec-list f p t = compute-list-vec-map-vec-permute-vec-vec-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} - (f : A → B) (p : list A) (t : Permutation (length-list p)) → + (f : A → B) (p : list A) (t : permutation (length-list p)) → list-vec ( length-list p) ( map-vec f (permute-vec (length-list p) (vec-list p) t)) = @@ -267,8 +267,8 @@ compute-list-vec-map-vec-permute-vec-vec-list f p t = eq-map-list-permute-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} - (f : A → B) (p : list A) (t : Permutation (length-list p)) → - permute-list (map-list f p) (tr Permutation (inv (length-map-list f p)) t) = + (f : A → B) (p : list A) (t : permutation (length-list p)) → + permute-list (map-list f p) (tr permutation (inv (length-map-list f p)) t) = map-list f (permute-list p t) eq-map-list-permute-list {B = B} f p t = ( ( ap @@ -283,7 +283,7 @@ eq-map-list-permute-list {B = B} f p t = ( permute-vec ( length-list (map-list f p)) ( x) - ( tr Permutation (inv (length-map-list f p)) t))) + ( tr permutation (inv (length-map-list f p)) t))) ( vec-list-map-list-map-vec-list' f p)) ∙ ( ( compute-tr-permute-vec ( length-map-list f p) diff --git a/src/lists/permutation-vectors.lagda.md b/src/lists/permutation-vectors.lagda.md index 4099f1bfe0..f77bcc3a13 100644 --- a/src/lists/permutation-vectors.lagda.md +++ b/src/lists/permutation-vectors.lagda.md @@ -50,7 +50,7 @@ module _ {l : Level} {A : UU l} where - permute-vec : (n : ℕ) → vec A n → Permutation n → vec A n + permute-vec : (n : ℕ) → vec A n → permutation n → vec A n permute-vec n v s = listed-vec-functional-vec n (functional-vec-vec n v ∘ (map-equiv s)) ``` @@ -61,12 +61,12 @@ module _ is-permutation-vec : (n : ℕ) → (vec A n → vec A n) → UU l is-permutation-vec n f = (v : vec A n) → - Σ ( Permutation n) + Σ ( permutation n) ( λ t → f v = permute-vec n v t) permutation-is-permutation-vec : (n : ℕ) (f : vec A n → vec A n) → is-permutation-vec n f → - (v : vec A n) → Permutation n + (v : vec A n) → permutation n permutation-is-permutation-vec n f P v = pr1 (P v) eq-permute-vec-permutation-is-permutation-vec : @@ -91,7 +91,7 @@ module _ compute-composition-permute-vec : (n : ℕ) (v : vec A n) → - (a b : Permutation n) → + (a b : permutation n) → permute-vec n v (a ∘e b) = permute-vec n (permute-vec n v a) b compute-composition-permute-vec n v a b = ap @@ -129,7 +129,7 @@ module _ (n : ℕ) (v : vec A n) (x : A) - (t : Permutation n) → + (t : permutation n) → permute-vec (succ-ℕ n) (x ∷ v) (equiv-coproduct t id-equiv) = (x ∷ permute-vec n v t) compute-equiv-coproduct-permutation-id-equiv-permute-vec n v x t = @@ -146,7 +146,7 @@ module _ ap-permute-vec : {n : ℕ} - (a : Permutation n) + (a : permutation n) {v w : vec A n} → v = w → permute-vec n v a = permute-vec n w a @@ -157,13 +157,13 @@ module _ ```agda is-in-functional-vec-is-in-permute-functional-vec : - (n : ℕ) (v : Fin n → A) (t : Permutation n) (x : A) → + (n : ℕ) (v : Fin n → A) (t : permutation n) (x : A) → in-functional-vec n x (v ∘ map-equiv t) → in-functional-vec n x v is-in-functional-vec-is-in-permute-functional-vec n v t x (k , refl) = map-equiv t k , refl is-in-vec-is-in-permute-vec : - (n : ℕ) (v : vec A n) (t : Permutation n) (x : A) → + (n : ℕ) (v : vec A n) (t : permutation n) (x : A) → x ∈-vec (permute-vec n v t) → x ∈-vec v is-in-vec-is-in-permute-vec n v t x I = is-in-vec-is-in-functional-vec @@ -182,13 +182,13 @@ module _ ( is-in-functional-vec-is-in-vec n (permute-vec n v t) x I))) is-in-permute-functional-vec-is-in-functional-vec : - (n : ℕ) (v : Fin n → A) (t : Permutation n) (x : A) → + (n : ℕ) (v : Fin n → A) (t : permutation n) (x : A) → in-functional-vec n x v → in-functional-vec n x (v ∘ map-equiv t) is-in-permute-functional-vec-is-in-functional-vec n v t x (k , refl) = map-inv-equiv t k , ap v (inv (is-section-map-inv-equiv t k)) is-in-permute-vec-is-in-vec : - (n : ℕ) (v : vec A n) (t : Permutation n) (x : A) → + (n : ℕ) (v : vec A n) (t : permutation n) (x : A) → x ∈-vec v → x ∈-vec (permute-vec n v t) is-in-permute-vec-is-in-vec n v t x I = is-in-vec-is-in-functional-vec @@ -346,10 +346,10 @@ module _ ( transposition-Fin (succ-ℕ n) i j neq) ( permutation-list-standard-transpositions-Fin (succ-ℕ n) l)))))) - invariant-permutation-fold-vec : - {n : ℕ} → (v : vec A n) → (f : Permutation n) → + permutation-invariant-fold-vec : + {n : ℕ} → (v : vec A n) → (f : permutation n) → fold-vec b μ v = fold-vec b μ (permute-vec n v f) - invariant-permutation-fold-vec {n} v f = + permutation-invariant-fold-vec {n} v f = ( ( invariant-list-transpositions-fold-vec ( v) ( list-standard-transpositions-permutation-Fin n f)) ∙ @@ -369,7 +369,7 @@ module _ ```agda eq-map-vec-permute-vec : {l1 l2 : Level} {A : UU l1} {B : UU l2} - (f : A → B) {n : ℕ} (v : vec A n) (t : Permutation n) → + (f : A → B) {n : ℕ} (v : vec A n) (t : permutation n) → permute-vec n (map-vec f v) t = map-vec f (permute-vec n v t) eq-map-vec-permute-vec f {n} v t = diff --git a/src/lists/sort-by-insertion-lists.lagda.md b/src/lists/sort-by-insertion-lists.lagda.md index deca8a7272..ffc3c2d59a 100644 --- a/src/lists/sort-by-insertion-lists.lagda.md +++ b/src/lists/sort-by-insertion-lists.lagda.md @@ -65,7 +65,7 @@ module _ permutation-insertion-sort-list : (l : list (type-Decidable-Total-Order X)) → - Permutation (length-list l) + permutation (length-list l) permutation-insertion-sort-list = permutation-list-is-sort-list X diff --git a/src/lists/sort-by-insertion-vectors.lagda.md b/src/lists/sort-by-insertion-vectors.lagda.md index 56d1870c3a..04bc0e900d 100644 --- a/src/lists/sort-by-insertion-vectors.lagda.md +++ b/src/lists/sort-by-insertion-vectors.lagda.md @@ -89,7 +89,7 @@ module _ (x y : type-Decidable-Total-Order X) (l : vec (type-Decidable-Total-Order X) n) → leq-or-strict-greater-Decidable-Poset X x y → - Permutation (succ-ℕ (succ-ℕ (n))) + permutation (succ-ℕ (succ-ℕ (n))) helper-permutation-insertion-sort-vec x y l (inl _) = id-equiv helper-permutation-insertion-sort-vec {0} x y empty-vec (inr _) = swap-two-last-elements-transposition-Fin 0 @@ -106,7 +106,7 @@ module _ permutation-insertion-sort-vec : {n : ℕ} (v : vec (type-Decidable-Total-Order X) n) → - Permutation n + permutation n permutation-insertion-sort-vec {zero-ℕ} empty-vec = id-equiv permutation-insertion-sort-vec {1} l = id-equiv permutation-insertion-sort-vec {succ-ℕ (succ-ℕ n)} (x ∷ y ∷ l) = diff --git a/src/lists/sorted-vectors.lagda.md b/src/lists/sorted-vectors.lagda.md index dbb2353d4f..3e9a7117a4 100644 --- a/src/lists/sorted-vectors.lagda.md +++ b/src/lists/sorted-vectors.lagda.md @@ -201,7 +201,7 @@ module _ (n : ℕ) (x : type-Decidable-Total-Order X) (fv : functional-vec (type-Decidable-Total-Order X) n) - (a : Permutation n) → + (a : permutation n) → is-least-element-functional-vec n x fv → is-least-element-functional-vec n x (fv ∘ map-equiv a) is-least-element-permute-functional-vec n x fv a p k = @@ -247,7 +247,7 @@ module _ {n : ℕ} (x : type-Decidable-Total-Order X) (v : vec (type-Decidable-Total-Order X) n) - (a : Permutation n) → + (a : permutation n) → is-least-element-vec x v → is-least-element-vec x (permute-vec n v a) is-least-element-permute-vec {n} x v a p = diff --git a/src/lists/sorting-algorithms-lists.lagda.md b/src/lists/sorting-algorithms-lists.lagda.md index 14b5f6541e..2b99d85fcd 100644 --- a/src/lists/sorting-algorithms-lists.lagda.md +++ b/src/lists/sorting-algorithms-lists.lagda.md @@ -58,7 +58,7 @@ module _ permutation-list-is-sort-list : is-sort-list → (l : list (type-Decidable-Total-Order X)) → - Permutation (length-list l) + permutation (length-list l) permutation-list-is-sort-list S l = permutation-is-permutation-list f (is-permutation-list-is-sort-list S) l diff --git a/src/lists/sorting-algorithms-vectors.lagda.md b/src/lists/sorting-algorithms-vectors.lagda.md index 029591d5fb..a6c9cc1a18 100644 --- a/src/lists/sorting-algorithms-vectors.lagda.md +++ b/src/lists/sorting-algorithms-vectors.lagda.md @@ -55,7 +55,7 @@ module _ is-permutation-vec-is-sort-vec S n = pr1 (S n) permutation-vec-is-sort-vec : - is-sort-vec → (n : ℕ) → vec (type-Decidable-Total-Order X) n → Permutation n + is-sort-vec → (n : ℕ) → vec (type-Decidable-Total-Order X) n → permutation n permutation-vec-is-sort-vec S n v = permutation-is-permutation-vec n f (is-permutation-vec-is-sort-vec S n) v diff --git a/src/lists/universal-quantification-lists.lagda.md b/src/lists/universal-quantification-lists.lagda.md index 6a27b991eb..a57bf8e722 100644 --- a/src/lists/universal-quantification-lists.lagda.md +++ b/src/lists/universal-quantification-lists.lagda.md @@ -35,22 +35,34 @@ open import finite-group-theory.permutations-standard-finite-types ## Idea -Consider a predicate `P` on a type `X`, and consider a [list](lists.lists.md) `l` of elements of `X`. The {{#concept "universal quantification" Disambiguation="elements of a list" Agda=for-all-list}} of `P` over the elements of `l` is the type of choices of elements `P x` for each element `x ∈ l`. More specifically, the universal quantification `∀ l P` of `P` over `l` is defined inductively by +Consider a predicate `P` on a type `X`, and consider a [list](lists.lists.md) +`l` of elements of `X`. The +{{#concept "universal quantification" Disambiguation="elements of a list" Agda=for-all-list}} +of `P` over the elements of `l` is the type of choices of elements `P x` for +each element `x ∈ l`. More specifically, the universal quantification `∀ l P` of +`P` over `l` is defined inductively by ```text ∀ nil P := unit ∀ (cons x l) P := (P x) × (∀ l P) ``` -Alternatively, the universal quantification over the elements of a list can be defined directly by +Alternatively, the universal quantification over the elements of a list can be +defined directly by ```text ∀ l P := (x : X) → x ∈ l → P x. ``` -These definitions are [equivalent](foundation-core.equivalences.md). However, note that the inductive definition of the universal quantification has the same universe level as `P`, while the direct definition is of universe level `l1 ⊔ l2`, where `l1` is the universe level of `X` and `l2` is the universe level of `P`. +These definitions are [equivalent](foundation-core.equivalences.md). However, +note that the inductive definition of the universal quantification has the same +universe level as `P`, while the direct definition is of universe level +`l1 ⊔ l2`, where `l1` is the universe level of `X` and `l2` is the universe +level of `P`. -Note that the universal quantification over the elements of a list `l` is equivalent to the type of [dependent lists](lists.dependent-lists.md) of elements of `P` over `l`. +Note that the universal quantification over the elements of a list `l` is +equivalent to the type of [dependent lists](lists.dependent-lists.md) of +elements of `P` over `l`. ## Definitions @@ -157,7 +169,7 @@ module _ module _ (x : X) (l : list X) (P : X → UU l2) where - + map-compute-for-all-elements-cons-list : P x × for-all-elements-list l P → for-all-elements-list (cons x l) P map-compute-for-all-elements-cons-list (a , b) = @@ -255,7 +267,7 @@ pr2 (for-all-concat-list (cons x l) k P H K) = ```text for-all-permute-list : {l1 l2 : Level} {X : UU l1} (l : list X) (P : X → UU l2) - (e : Permutation (length-list l)) → + (e : permutation (length-list l)) → for-all-list l P → for-all-list (permute-list l e) P for-all-permute-list l P e H = {!!} ``` diff --git a/src/order-theory/strictly-inflationary-maps-strict-preorders.lagda.md b/src/order-theory/strictly-inflationary-maps-strict-preorders.lagda.md index d8a4672caa..ea8d15df33 100644 --- a/src/order-theory/strictly-inflationary-maps-strict-preorders.lagda.md +++ b/src/order-theory/strictly-inflationary-maps-strict-preorders.lagda.md @@ -84,10 +84,10 @@ module _ map-strictly-inflationary-map-Strict-Preorder = pr1 f - is-inflationary-strictly-inflationary-map-Strict-Preorder : + is-strictly-inflationary-strictly-inflationary-map-Strict-Preorder : is-strictly-inflationary-map-Strict-Preorder P ( map-strictly-inflationary-map-Strict-Preorder) - is-inflationary-strictly-inflationary-map-Strict-Preorder = + is-strictly-inflationary-strictly-inflationary-map-Strict-Preorder = pr2 f ``` diff --git a/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md b/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md index b30db00541..23eaec6c21 100644 --- a/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md +++ b/src/univalent-combinatorics/cycle-prime-decomposition-natural-numbers.lagda.md @@ -136,7 +136,7 @@ equiv-product-cycle-prime-decomposition-ℕ n m H I = ( list-fundamental-theorem-arithmetic-ℕ n H) ( list-fundamental-theorem-arithmetic-ℕ m I))) ( tr - ( Permutation) + ( permutation) ( inv ( length-map-list ( Cyclic-Type lzero) From df6c71994dc304c6106ade8aa5dda95dffc4823d Mon Sep 17 00:00:00 2001 From: Egbert Rijke Date: Tue, 11 Feb 2025 12:54:11 -0500 Subject: [PATCH 72/72] edits? --- .../fundamental-theorem-of-arithmetic.lagda.md | 12 +++--------- 1 file changed, 3 insertions(+), 9 deletions(-) diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index 6bd84555bc..5ffdf08795 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -783,15 +783,9 @@ is-prime-list-permute-list-ℕ : is-prime-list-permute-list-ℕ p t P = is-prime-list-is-prime-is-in-list-ℕ ( permute-list p t) - ( λ x I → is-prime-is-in-list-is-prime-list-ℕ - ( p) - ( P) - ( x) - ( is-in-list-is-in-permute-list - ( p) - ( t) - ( x) - ( I))) + ( λ x I → + is-prime-is-in-list-is-prime-list-ℕ p P x + ( is-in-list-is-in-permute-list p t x I)) is-decomposition-list-permute-list-ℕ : (n : ℕ) (p : list ℕ) (t : permutation (length-list p)) →