From 4ee67ce8d5f8828a89c8b39ec85c6156ccac456f Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 15 Jul 2023 17:38:07 +0100 Subject: [PATCH 01/47] Add Ring.divide --- .../Algebra/GroupWithZero/Divisibility.lean | 67 ++++++++++++++++++- 1 file changed, 66 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index 04994b03b2e232..f3a1bc9ed62e12 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -9,7 +9,7 @@ Neil Strickland, Aaron Anderson ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathlib.Algebra.GroupWithZero.Basic +import Mathlib.Algebra.GroupWithZero.Units.Basic import Mathlib.Algebra.Divisibility.Units /-! @@ -17,6 +17,8 @@ import Mathlib.Algebra.Divisibility.Units Lemmas about divisibility in groups and monoids with zero. +We also define `Ring.divide`, a globally defined function on any ring +(in fact any `MonoidWithZero`), which returns division whenever divisible and zero otherwise. -/ @@ -141,3 +143,66 @@ theorem eq_of_forall_dvd' (h : ∀ c, c ∣ a ↔ c ∣ b) : a = b := #align eq_of_forall_dvd' eq_of_forall_dvd' end CancelCommMonoidWithZero + +namespace Ring + +open Classical + +variable {M₀ : Type u} [MonoidWithZero M₀] + +/-- Introduce a binary function `divide` on monoids with zero `M₀`, which sends `x` and `y` to +`x / y` if `y` is non-zero and divides `x`, and to `0` otherwise. This definition is somewhat +ad hoc, but one needs a fully (rather than partially) defined division function for some purposes. + +Note that while this is in the `Ring` namespace for brevity, it requires the weaker assumption +`MonoidWithZero M₀` instead of `Ring M₀`. -/ +noncomputable def divide (x y : M₀) : M₀ := + if h : y ≠ 0 ∧ y ∣ x then h.right.choose else 0 + +lemma divide_dvd {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : divide x y = hx.choose := by + rw [divide, dif_pos ⟨hy, hx⟩] + +lemma divide_not_dvd {x y : M₀} (hx : ¬y ∣ x) : divide x y = 0 := by + simp only [divide, dif_neg, hx, and_false] + +lemma divide_zero (x : M₀) : divide x 0 = 0 := by + simp only [divide, dif_neg, ne_eq, false_and] + +lemma divide_one [NeZero (1 : M₀)] (x : M₀) : divide x 1 = x := by + rw [divide_dvd one_ne_zero <| one_dvd x, ← one_mul <| Exists.choose _] + exact (one_dvd x).choose_spec.symm + +lemma zero_divide [NoZeroDivisors M₀] (y : M₀) : divide 0 y = 0 := by + by_cases hy : y = 0 + · rw [hy, divide_zero] + · rw [divide_dvd hy <| dvd_zero y] + exact (eq_zero_or_eq_zero_of_mul_eq_zero (dvd_zero y).choose_spec.symm).resolve_left hy + +lemma one_divide {M₀ : Type u} [CommMonoidWithZero M₀] (y : M₀) : divide 1 y = inverse y := by + by_cases hy : y = 0 + · rw [hy, divide_zero, inverse_zero] + · by_cases hy' : y ∣ 1 + · have hy'' : IsUnit y := isUnit_of_dvd_one hy' + rw [divide_dvd hy hy', ← (inverse_mul_eq_iff_eq_mul y 1 _ hy'').mpr hy'.choose_spec, mul_one] + · rw [divide_not_dvd hy', inverse_non_unit] + exact hy' ∘ isUnit_iff_dvd_one.mp + +lemma mul_divide_cancel {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : y * divide x y = x := by + simp only [divide_dvd hy hx, hx.choose_spec.symm] + +lemma mul_divide_cancel_left [IsLeftCancelMulZero M₀] {x y : M₀} (hx : x ≠ 0) : + divide (x * y) x = y := by + rw [divide_dvd hx <| dvd_mul_right x y] + exact mul_left_cancel₀ hx (dvd_mul_right x y).choose_spec.symm + +variable {M₀ : Type u} [CommMonoidWithZero M₀] + +lemma divide_mul_cancel {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : divide x y * y = x := by + rw [mul_comm, mul_divide_cancel hy hx] + +lemma mul_divide_cancel_right [IsRightCancelMulZero M₀] {x y : M₀} (hy : y ≠ 0) : + divide (x * y) y = x := by + rw [divide_dvd hy <| dvd_mul_left y x] + exact mul_right_cancel₀ hy <| mul_comm _ y ▸ (dvd_mul_left y x).choose_spec.symm + +end Ring From eab1469a0720330eb05e12ac51605b09734c5513 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 18 Jul 2023 08:12:41 +0100 Subject: [PATCH 02/47] Add IsUnit.divide_eq_mul_inverse --- Mathlib/Algebra/GroupWithZero/Divisibility.lean | 11 ++++++++--- 1 file changed, 8 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index f3a1bc9ed62e12..1ef6516dd6984c 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -148,7 +148,7 @@ namespace Ring open Classical -variable {M₀ : Type u} [MonoidWithZero M₀] +variable {M₀ : Type _} [MonoidWithZero M₀] /-- Introduce a binary function `divide` on monoids with zero `M₀`, which sends `x` and `y` to `x / y` if `y` is non-zero and divides `x`, and to `0` otherwise. This definition is somewhat @@ -178,7 +178,7 @@ lemma zero_divide [NoZeroDivisors M₀] (y : M₀) : divide 0 y = 0 := by · rw [divide_dvd hy <| dvd_zero y] exact (eq_zero_or_eq_zero_of_mul_eq_zero (dvd_zero y).choose_spec.symm).resolve_left hy -lemma one_divide {M₀ : Type u} [CommMonoidWithZero M₀] (y : M₀) : divide 1 y = inverse y := by +lemma one_divide {M₀ : Type _} [CommMonoidWithZero M₀] (y : M₀) : divide 1 y = inverse y := by by_cases hy : y = 0 · rw [hy, divide_zero, inverse_zero] · by_cases hy' : y ∣ 1 @@ -195,7 +195,7 @@ lemma mul_divide_cancel_left [IsLeftCancelMulZero M₀] {x y : M₀} (hx : x ≠ rw [divide_dvd hx <| dvd_mul_right x y] exact mul_left_cancel₀ hx (dvd_mul_right x y).choose_spec.symm -variable {M₀ : Type u} [CommMonoidWithZero M₀] +variable {M₀ : Type _} [CommMonoidWithZero M₀] lemma divide_mul_cancel {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : divide x y * y = x := by rw [mul_comm, mul_divide_cancel hy hx] @@ -206,3 +206,8 @@ lemma mul_divide_cancel_right [IsRightCancelMulZero M₀] {x y : M₀} (hy : y exact mul_right_cancel₀ hy <| mul_comm _ y ▸ (dvd_mul_left y x).choose_spec.symm end Ring + +lemma IsUnit.divide_eq_mul_inverse {M₀ : Type _} [Nontrivial M₀] [CommMonoidWithZero M₀] + [IsLeftCancelMul M₀] {x y : M₀} (hy : IsUnit y) : Ring.divide x y = x * Ring.inverse y := by + rw [Ring.divide_dvd hy.ne_zero hy.dvd, ← hy.mul_right_inj, ← hy.dvd.choose_spec, mul_comm, + Ring.inverse_mul_cancel_right _ _ hy] From 50f07f5c33aa831de016891d31814ce3221abdf9 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 18 Jul 2023 23:16:46 +0100 Subject: [PATCH 03/47] Remove redundant assumption --- Mathlib/Algebra/GroupWithZero/Divisibility.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index 1ef6516dd6984c..f607523cc1d725 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -208,6 +208,6 @@ lemma mul_divide_cancel_right [IsRightCancelMulZero M₀] {x y : M₀} (hy : y end Ring lemma IsUnit.divide_eq_mul_inverse {M₀ : Type _} [Nontrivial M₀] [CommMonoidWithZero M₀] - [IsLeftCancelMul M₀] {x y : M₀} (hy : IsUnit y) : Ring.divide x y = x * Ring.inverse y := by + {x y : M₀} (hy : IsUnit y) : Ring.divide x y = x * Ring.inverse y := by rw [Ring.divide_dvd hy.ne_zero hy.dvd, ← hy.mul_right_inj, ← hy.dvd.choose_spec, mul_comm, Ring.inverse_mul_cancel_right _ _ hy] From 7cd0323999a6b48012857de9b212ed5cb498289c Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 8 Aug 2023 23:48:07 +0100 Subject: [PATCH 04/47] Replace NeZero with Nontrivial --- Mathlib/Algebra/GroupWithZero/Divisibility.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index f607523cc1d725..caf4dcc73a71ac 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -168,7 +168,7 @@ lemma divide_not_dvd {x y : M₀} (hx : ¬y ∣ x) : divide x y = 0 := by lemma divide_zero (x : M₀) : divide x 0 = 0 := by simp only [divide, dif_neg, ne_eq, false_and] -lemma divide_one [NeZero (1 : M₀)] (x : M₀) : divide x 1 = x := by +lemma divide_one [Nontrivial M₀] (x : M₀) : divide x 1 = x := by rw [divide_dvd one_ne_zero <| one_dvd x, ← one_mul <| Exists.choose _] exact (one_dvd x).choose_spec.symm From bd2f1116133ba08549d55a5f4e1e88c1914936f5 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 21 Aug 2023 12:45:43 +0100 Subject: [PATCH 05/47] Define division polynomials --- .../EllipticCurve/Torsion.lean | 185 ++++++++++++++++++ 1 file changed, 185 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean new file mode 100644 index 00000000000000..427b90e41486c8 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import Mathlib.AlgebraicGeometry.EllipticCurve.Point + +/-! +# Torsion points on Weierstrass curves +-/ + +universe u + +namespace WeierstrassCurve + +open Polynomial + +open scoped Polynomial PolynomialPolynomial + +variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) + +noncomputable def divisionPolynomial : ℕ → (R[X][Y]) +| 0 => 0 +| 1 => 1 +| 2 => Y - W.negPolynomial +| 3 => C <| 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ +| 4 => (Y - W.negPolynomial) + * C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) +| (n + 5) => +if hn : Even n then + let m := ((even_iff_exists_two_mul n).mp hn).choose + have h4 : m + 4 < n + 5 := + by linarith only [show n = 2 * m by exact ((even_iff_exists_two_mul n).mp hn).choose_spec] + have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 + have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 + have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 + divisionPolynomial (m + 4) * divisionPolynomial (m + 2) ^ 3 + - divisionPolynomial (m + 1) * divisionPolynomial (m + 3) ^ 3 +else + let m := (Nat.odd_iff_not_even.mpr hn).choose + have h : m < n := + by linarith only [show n = 2 * m + 1 by exact (Nat.odd_iff_not_even.mpr hn).choose_spec] + have h5 : m + 5 < n + 5 := add_lt_add_right h 5 + have h4 : m + 4 < n + 5 := (lt_add_one _).trans h5 + have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 + have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 + have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 + Ring.divide + (divisionPolynomial (m + 2) ^ 2 * divisionPolynomial (m + 3) * divisionPolynomial (m + 5) + - divisionPolynomial (m + 1) * divisionPolynomial (m + 3) * divisionPolynomial (m + 4) ^ 2) + (Y - W.negPolynomial) + +@[simp] +lemma divisionPolynomial_zero : W.divisionPolynomial 0 = 0 := + rfl + +@[simp] +lemma divisionPolynomial_one : W.divisionPolynomial 1 = 1 := + rfl + +@[simp] +lemma divisionPolynomial_two : W.divisionPolynomial 2 = Y - W.negPolynomial := + rfl + +@[simp] +lemma divisionPolynomial_three : + W.divisionPolynomial 3 = + C (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) := + rfl + +@[simp] +lemma divisionPolynomial_four : + W.divisionPolynomial 4 = + (Y - W.negPolynomial) + * C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) := + rfl + +@[simp] +lemma divisionPolynomial_odd (m : ℕ) : + W.divisionPolynomial (2 * (m + 2) + 1) = + W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 + - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 := by + erw [divisionPolynomial, dif_pos <| even_two_mul m] + simp only [← (mul_right_inj' two_ne_zero).mp (⟨m, rfl⟩ : ∃ n : ℕ, 2 * m = 2 * n).choose_spec] + +@[simp] +lemma divisionPolynomial_even (m : ℕ) : + W.divisionPolynomial (2 * (m + 3)) = + Ring.divide + (W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) + * W.divisionPolynomial (m + 5) + - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) + * W.divisionPolynomial (m + 4) ^ 2) + (Y - W.negPolynomial) := by + erw [divisionPolynomial, dif_neg <| Nat.odd_iff_not_even.mp <| odd_two_mul_add_one m] + simp only [← (mul_right_inj' two_ne_zero).mp <| + Nat.succ_injective (odd_two_mul_add_one m).choose_spec] + +private +lemma even_rw (y p1 p2 p3 p4 p5 : R) : + (y * p2) ^ 2 * p3 * p5 - p1 * p3 * (y * p4) ^ 2 = + y * (y * (p2 ^ 2 * p3 * p5 - p1 * p3 * p4 ^ 2)) := by + ring1 + +private +lemma odd_rw (y p1 p2 p3 p4 p5 : R) : + p2 ^ 2 * (y * p3) * (y * p5) - (y * p1) * (y * p3) * p4 ^ 2 = + y * (y * (p2 ^ 2 * p3 * p5 - p1 * p3 * p4 ^ 2)) := by + ring1 + +set_option maxHeartbeats 500000 in +lemma dvd_divisionPolynomial_even [IsDomain R] (m : ℕ) : + Y - W.negPolynomial ∣ W.divisionPolynomial (2 * m) := by + suffices ∀ {n : ℕ}, Even n → Y - W.negPolynomial ∣ W.divisionPolynomial n by + exact this <| even_two_mul _ + intro n hn + induction' n using Nat.strongRec' with n ih + rcases hn with ⟨n, rfl⟩ + rcases n with _ | _ | _ | n + · exact ⟨0, by rw [mul_zero, W.divisionPolynomial_zero]⟩ + · exact ⟨1, by rw [mul_one, W.divisionPolynomial_two]⟩ + · exact ⟨_, W.divisionPolynomial_four⟩ + · rw [show n + 3 + (n + 3) = 2 * n + 6 by exact (two_mul <| n + 3).symm] at ih ⊢ + rw [divisionPolynomial] + simp only [dif_neg <| Nat.odd_iff_not_even.mp <| odd_two_mul_add_one n, + ← (mul_right_inj' two_ne_zero).mp <| Nat.succ_injective (odd_two_mul_add_one n).choose_spec] + rcases n.even_or_odd' with ⟨n, rfl | rfl⟩ + · cases' ih (2 * n + 2) (by linarith only) (even_two_mul <| n + 1) with _ h2 + cases' ih (2 * n + 4) (by linarith only) (even_two_mul <| n + 2) with _ h4 + rw [h2, h4, even_rw] + by_cases hY : Y - W.negPolynomial = 0 + · rw [hY, zero_mul, Ring.divide_zero] + · simpa only [Ring.mul_divide_cancel_left hY] using dvd_mul_right _ _ + · cases' ih (2 * n + 1 + 1) (by linarith only) (even_two_mul <| n + 1) with _ h1 + cases' ih (2 * n + 1 + 3) (by linarith only) (even_two_mul <| n + 2) with _ h3 + cases' ih (2 * n + 1 + 5) (by linarith only) (even_two_mul <| n + 3) with _ h5 + rw [h1, h3, h5, odd_rw] + by_cases hY : Y - W.negPolynomial = 0 + · rw [hY, zero_mul, Ring.divide_zero] + · simpa only [Ring.mul_divide_cancel_left hY] using dvd_mul_right _ _ + +lemma mul_divisionPolynomial_even [IsDomain R] (m : ℕ) : + (Y - W.negPolynomial) * W.divisionPolynomial (2 * (m + 3)) = + W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * W.divisionPolynomial (m + 5) + - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) + * W.divisionPolynomial (m + 4) ^ 2 := by + rw [divisionPolynomial_even] + rcases m.even_or_odd' with ⟨m, rfl | rfl⟩ + · rw [show 2 * m + 2 = 2 * (m + 1) by rfl, show 2 * m + 4 = 2 * (m + 2) by rfl, + (W.dvd_divisionPolynomial_even $ m + 1).choose_spec, + (W.dvd_divisionPolynomial_even $ m + 2).choose_spec, even_rw] + by_cases hY : Y - W.negPolynomial = 0 + · simp only [hY, zero_mul] + · rw [Ring.mul_divide_cancel hY] + exact dvd_mul_right _ _ + · rw [show 2 * m + 1 + 1 = 2 * (m + 1) by rfl, show 2 * m + 1 + 3 = 2 * (m + 2) by rfl, + show 2 * m + 1 + 5 = 2 * (m + 3) by rfl, + (W.dvd_divisionPolynomial_even $ m + 1).choose_spec, + (W.dvd_divisionPolynomial_even $ m + 2).choose_spec, + (W.dvd_divisionPolynomial_even $ m + 3).choose_spec, odd_rw] + by_cases hY : Y - W.negPolynomial = 0 + · simp only [hY, zero_mul] + · rw [Ring.mul_divide_cancel hY] + exact dvd_mul_right _ _ + +noncomputable +def divisionPolynomial' : ℤ → (R[X][Y]) +| Int.ofNat n => W.divisionPolynomial n +| Int.negSucc n => -W.divisionPolynomial (n + 1) + +@[simp] +lemma divisionPolynomial'_nat (n : ℕ) : W.divisionPolynomial' n = W.divisionPolynomial n := + rfl + +@[simp] +lemma divisionPolynomial'_neg (n : ℕ) : W.divisionPolynomial' (-n) = -W.divisionPolynomial n := by + cases n + · erw [Int.neg_zero, divisionPolynomial'_nat, W.divisionPolynomial_zero, neg_zero] + · rfl + +end WeierstrassCurve From 24f6d6e531a5f7c65a4f9ecc72bfc9d67846eae3 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 27 Nov 2023 01:29:21 +0000 Subject: [PATCH 06/47] Remove divisibility changes --- .../Algebra/GroupWithZero/Divisibility.lean | 72 +------------------ 1 file changed, 1 insertion(+), 71 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index 5d31c129647650..7771d178878fe4 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Amelia Livingston, Yury Kudryashov, Neil Strickland, Aaron Anderson -/ -import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.Algebra.GroupWithZero.Basic import Mathlib.Algebra.Divisibility.Units #align_import algebra.group_with_zero.divisibility from "leanprover-community/mathlib"@"e8638a0fcaf73e4500469f368ef9494e495099b3" @@ -14,8 +14,6 @@ import Mathlib.Algebra.Divisibility.Units Lemmas about divisibility in groups and monoids with zero. -We also define `Ring.divide`, a globally defined function on any ring -(in fact any `MonoidWithZero`), which returns division whenever divisible and zero otherwise. -/ @@ -140,71 +138,3 @@ theorem eq_of_forall_dvd' (h : ∀ c, c ∣ a ↔ c ∣ b) : a = b := #align eq_of_forall_dvd' eq_of_forall_dvd' end CancelCommMonoidWithZero - -namespace Ring - -open Classical - -variable {M₀ : Type _} [MonoidWithZero M₀] - -/-- Introduce a binary function `divide` on monoids with zero `M₀`, which sends `x` and `y` to -`x / y` if `y` is non-zero and divides `x`, and to `0` otherwise. This definition is somewhat -ad hoc, but one needs a fully (rather than partially) defined division function for some purposes. - -Note that while this is in the `Ring` namespace for brevity, it requires the weaker assumption -`MonoidWithZero M₀` instead of `Ring M₀`. -/ -noncomputable def divide (x y : M₀) : M₀ := - if h : y ≠ 0 ∧ y ∣ x then h.right.choose else 0 - -lemma divide_dvd {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : divide x y = hx.choose := by - rw [divide, dif_pos ⟨hy, hx⟩] - -lemma divide_not_dvd {x y : M₀} (hx : ¬y ∣ x) : divide x y = 0 := by - simp only [divide, dif_neg, hx, and_false] - -lemma divide_zero (x : M₀) : divide x 0 = 0 := by - simp only [divide, dif_neg, ne_eq, false_and] - -lemma divide_one [Nontrivial M₀] (x : M₀) : divide x 1 = x := by - rw [divide_dvd one_ne_zero <| one_dvd x, ← one_mul <| Exists.choose _] - exact (one_dvd x).choose_spec.symm - -lemma zero_divide [NoZeroDivisors M₀] (y : M₀) : divide 0 y = 0 := by - by_cases hy : y = 0 - · rw [hy, divide_zero] - · rw [divide_dvd hy <| dvd_zero y] - exact (eq_zero_or_eq_zero_of_mul_eq_zero (dvd_zero y).choose_spec.symm).resolve_left hy - -lemma one_divide {M₀ : Type _} [CommMonoidWithZero M₀] (y : M₀) : divide 1 y = inverse y := by - by_cases hy : y = 0 - · rw [hy, divide_zero, inverse_zero] - · by_cases hy' : y ∣ 1 - · have hy'' : IsUnit y := isUnit_of_dvd_one hy' - rw [divide_dvd hy hy', ← (inverse_mul_eq_iff_eq_mul y 1 _ hy'').mpr hy'.choose_spec, mul_one] - · rw [divide_not_dvd hy', inverse_non_unit] - exact hy' ∘ isUnit_iff_dvd_one.mp - -lemma mul_divide_cancel {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : y * divide x y = x := by - simp only [divide_dvd hy hx, hx.choose_spec.symm] - -lemma mul_divide_cancel_left [IsLeftCancelMulZero M₀] {x y : M₀} (hx : x ≠ 0) : - divide (x * y) x = y := by - rw [divide_dvd hx <| dvd_mul_right x y] - exact mul_left_cancel₀ hx (dvd_mul_right x y).choose_spec.symm - -variable {M₀ : Type _} [CommMonoidWithZero M₀] - -lemma divide_mul_cancel {x y : M₀} (hy : y ≠ 0) (hx : y ∣ x) : divide x y * y = x := by - rw [mul_comm, mul_divide_cancel hy hx] - -lemma mul_divide_cancel_right [IsRightCancelMulZero M₀] {x y : M₀} (hy : y ≠ 0) : - divide (x * y) y = x := by - rw [divide_dvd hy <| dvd_mul_left y x] - exact mul_right_cancel₀ hy <| mul_comm _ y ▸ (dvd_mul_left y x).choose_spec.symm - -end Ring - -lemma IsUnit.divide_eq_mul_inverse {M₀ : Type _} [Nontrivial M₀] [CommMonoidWithZero M₀] - {x y : M₀} (hy : IsUnit y) : Ring.divide x y = x * Ring.inverse y := by - rw [Ring.divide_dvd hy.ne_zero hy.dvd, ← hy.mul_right_inj, ← hy.dvd.choose_spec, mul_comm, - Ring.inverse_mul_cancel_right _ _ hy] From 04200b4e58328cbf11aa963188e636c88dc1ddeb Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 1 Jan 2024 00:50:34 +0000 Subject: [PATCH 07/47] Define elliptic divisibility sequences --- Mathlib.lean | 1 + .../EllipticDivisibilitySequence.lean | 222 ++++++++++++++++++ 2 files changed, 223 insertions(+) create mode 100644 Mathlib/NumberTheory/EllipticDivisibilitySequence.lean diff --git a/Mathlib.lean b/Mathlib.lean index 72f713857128d4..b25c49263926ad 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2742,6 +2742,7 @@ import Mathlib.NumberTheory.DiophantineApproximation import Mathlib.NumberTheory.DirichletCharacter.Basic import Mathlib.NumberTheory.DirichletCharacter.Bounds import Mathlib.NumberTheory.Divisors +import Mathlib.NumberTheory.EllipticDivisibilitySequence import Mathlib.NumberTheory.EulerProduct.Basic import Mathlib.NumberTheory.EulerProduct.DirichletLSeries import Mathlib.NumberTheory.FLT.Basic diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean new file mode 100644 index 00000000000000..eeeff01b8bad61 --- /dev/null +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -0,0 +1,222 @@ +/- +Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import Mathlib.Data.Nat.Parity +import Mathlib.Tactic.Linarith + +/-! +# Elliptic divisibility sequences + +This file defines the type of elliptic divisibility sequences and a few examples. + +## Mathematical background + +Let $R$ be a commutative ring. An elliptic sequence is a sequence $h : \mathbb{Z} \to R$ satisfying +$$ h(m + n)h(m - n)h(r)^2 = h(m + r)h(m - r)h(n)^2 - h(n + r)h(n - r)h(m)^2, $$ +for any $m, n, r \in \mathbb{Z}$. A divisibility sequence is a sequence $h : \mathbb{Z} \to R$ +satisfying $h(m) \mid h(n)$ for any $m, n \in \mathbb{Z}$ such that $m \mid n$. + +Some examples of elliptic divisibility sequences include + * the integers $\mathbb{Z}$, + * certain terms of Lucas sequences, and + * division polynomials of elliptic curves. + +## Main definitions + + * `isEllSequence`: a sequence indexed by integers is an elliptic sequence. + * `isDivSequence`: a sequence indexed by integers is a divisibility sequence. + * `isEllDivSequence`: a sequence indexed by integers is an elliptic divisibility sequence. + * `EllDivSequence'`: a canonical example of an elliptic divisibility sequence indexed by naturals. + * `EllDivSequence`: a canonical example of an elliptic divisibility sequence indexed by integers. + +## Main statements + + * TODO: prove that `EllDivSequence` is an elliptic divisibility sequence. + * TODO: prove that a general elliptic divisibility sequence can be given by `EllDivSequence`. + +## References + +M Ward, *Memoir on Elliptic Divisibility Sequences* + +## Tags + +elliptic, divisibility, sequence +-/ + +universe u v w + +variable {R : Type u} [CommRing R] + +/-- The proposition that a sequence indexed by integers is an elliptic sequence. -/ +def isEllSequence (h : ℤ → R) : Prop := + ∀ m n r : ℤ, h (m + n) * h (m - n) * h r ^ 2 = + h (m + r) * h (m - r) * h n ^ 2 - h (n + r) * h (n - r) * h m ^ 2 + +/-- The proposition that a sequence indexed by integers is a divisibility sequence. -/ +def isDivSequence (h : ℤ → R) : Prop := + ∀ m n : ℕ, m ∣ n → h m ∣ h n + +/-- The proposition that a sequence indexed by integers is an elliptic divisibility sequence. -/ +def isEllDivSequence (h : ℤ → R) : Prop := + isEllSequence h ∧ isDivSequence h + +/-- The integers form an elliptic divisibility sequence. -/ +lemma Int.isEllDivSequence : isEllDivSequence id := + ⟨fun _ _ _ => by simp only [id_eq]; ring1, fun _ _ => Int.ofNat_dvd.mpr⟩ + +private def EllDivSequence'' (b c d : R) : ℕ → R + | 0 => 0 + | 1 => 1 + | 2 => 1 + | 3 => c + | 4 => d + | (n + 5) => + if hn : Even n then + let m := n / 2 + have h4 : m + 4 < n + 5 := + by linarith only [show n = 2 * m by exact (Nat.two_mul_div_two_of_even hn).symm] + have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 + have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 + have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 + EllDivSequence'' b c d (m + 4) * EllDivSequence'' b c d (m + 2) ^ 3 * + (if Even m then b ^ 4 else 1) + - EllDivSequence'' b c d (m + 1) * EllDivSequence'' b c d (m + 3) ^ 3 * + (if Even m then 1 else b ^ 4) + else + let m := n / 2 + have h5 : m + 5 < n + 5 := by + linarith only [show n = 2 * m + 1 + by exact (Nat.two_mul_div_two_add_one_of_odd <| Nat.odd_iff_not_even.mpr hn).symm] + have h4 : m + 4 < n + 5 := (lt_add_one _).trans h5 + have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 + have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 + have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 + EllDivSequence'' b c d (m + 2) ^ 2 * EllDivSequence'' b c d (m + 3) * + EllDivSequence'' b c d (m + 5) + - EllDivSequence'' b c d (m + 1) * EllDivSequence'' b c d (m + 3) * + EllDivSequence'' b c d (m + 4) ^ 2 + +variable (b c d : R) + +/-- The canonical example of an elliptic divisibility sequence `h : ℕ → R`, +with initial values `h(2) = b`, `h(3) = c`, and `h(4) = d`. + +This is defined in terms of a truncated sequence whose even terms differ by a factor of `b`. -/ +def EllDivSequence' (n : ℕ) : R := + EllDivSequence'' b c d n * if Even n then b else 1 + +@[simp] +lemma EllDivSequence'_zero : EllDivSequence' b c d 0 = 0 := by + rw [EllDivSequence', EllDivSequence'', zero_mul] + +@[simp] +lemma EllDivSequence'_one : EllDivSequence' b c d 1 = 1 := by + rw [EllDivSequence', EllDivSequence'', one_mul, if_neg Nat.not_even_one] + +@[simp] +lemma EllDivSequence'_two : EllDivSequence' b c d 2 = b := by + rw [EllDivSequence', EllDivSequence'', one_mul, if_pos even_two] + +@[simp] +lemma EllDivSequence'_three : EllDivSequence' b c d 3 = c := by + rw [EllDivSequence', EllDivSequence'', if_neg <| by decide, mul_one] + +@[simp] +lemma EllDivSequence'_four : EllDivSequence' b c d 4 = d * b := by + rw [EllDivSequence', EllDivSequence'', if_pos <| by decide] + +@[simp] +lemma EllDivSequence'_odd (m : ℕ) : EllDivSequence' b c d (2 * (m + 2) + 1) = + EllDivSequence' b c d (m + 4) * EllDivSequence' b c d (m + 2) ^ 3 + - EllDivSequence' b c d (m + 1) * EllDivSequence' b c d (m + 3) ^ 3 := by + rw [EllDivSequence', if_neg <| fun h => Nat.even_add_one.mp h <| even_two_mul _, + show 2 * (m + 2) + 1 = 2 * m + 5 by rfl, EllDivSequence'', dif_pos <| even_two_mul m] + simp only [EllDivSequence', Nat.mul_div_right _ zero_lt_two] + by_cases hm : Even m + · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm + have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 + have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 + have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 + rw [if_pos hm, if_pos hm, if_pos hm4, if_pos hm2, if_neg hm1, if_neg hm3] + ring1 + · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm + have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 + have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 + have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 + rw [if_neg hm, if_neg hm, if_neg hm4, if_neg hm2, if_pos hm1, if_pos hm3] + ring1 + +@[simp] +lemma EllDivSequence'_even (m : ℕ) : EllDivSequence' b c d (2 * (m + 3)) * b = + EllDivSequence' b c d (m + 2) ^ 2 * EllDivSequence' b c d (m + 3) * + EllDivSequence' b c d (m + 5) + - EllDivSequence' b c d (m + 1) * EllDivSequence' b c d (m + 3) * + EllDivSequence' b c d (m + 4) ^ 2 := by + rw [EllDivSequence', if_pos <| even_two_mul _, show 2 * (m + 3) = 2 * m + 1 + 5 by rfl, + EllDivSequence'', dif_neg <| fun h => Nat.even_add_one.mp h <| even_two_mul _] + simp only [EllDivSequence', Nat.mul_add_div two_pos, show 1 / 2 = 0 by rfl] + by_cases hm : Even m + · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm + have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 + have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 + have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 + have hm5 : ¬Even (m + 5) := fun h => Nat.even_add_one.mp h hm4 + rw [if_pos hm2, if_neg hm3, if_neg hm5, if_neg hm1, if_pos hm4] + ring1 + · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm + have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 + have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 + have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 + have hm5 : Even (m + 5) := Nat.even_add_one.mpr hm4 + rw [if_neg hm2, if_pos hm3, if_pos hm5, if_pos hm1, if_neg hm4] + ring1 + +/-- The canonical example of an elliptic divisibility sequence `h : ℤ → R`, +with initial values `h(2) = b`, `h(3) = c`, and `h(4) = d`. + +This extends `EllDivSequence'` by defining its values at negative integers. -/ +def EllDivSequence : ℤ → R + | Int.ofNat n => EllDivSequence' b c d n + | Int.negSucc n => -EllDivSequence' b c d (n + 1) + +@[simp] +lemma EllDivSequence_zero : EllDivSequence b c d 0 = 0 := + EllDivSequence'_zero b c d + +@[simp] +lemma EllDivSequence_one : EllDivSequence b c d 1 = 1 := + EllDivSequence'_one b c d + +@[simp] +lemma EllDivSequence_two : EllDivSequence b c d 2 = b := + EllDivSequence'_two b c d + +@[simp] +lemma EllDivSequence_three : EllDivSequence b c d 3 = c := + EllDivSequence'_three b c d + +@[simp] +lemma EllDivSequence_four : EllDivSequence b c d 4 = d * b := + EllDivSequence'_four b c d + +@[simp] +lemma EllDivSequence_odd (m : ℕ) : EllDivSequence b c d (2 * (m + 2) + 1) = + EllDivSequence b c d (m + 4) * EllDivSequence b c d (m + 2) ^ 3 - + EllDivSequence b c d (m + 1) * EllDivSequence b c d (m + 3) ^ 3 := + EllDivSequence'_odd b c d m + +@[simp] +lemma EllDivSequence_even (m : ℕ) : EllDivSequence b c d (2 * (m + 3)) * b = + EllDivSequence b c d (m + 2) ^ 2 * EllDivSequence b c d (m + 3) * EllDivSequence b c d (m + 5) - + EllDivSequence b c d (m + 1) * EllDivSequence b c d (m + 3) * + EllDivSequence b c d (m + 4) ^ 2 := + EllDivSequence'_even b c d m + +@[simp] +lemma EllDivSequence_neg (n : ℕ) : EllDivSequence b c d (-n) = -EllDivSequence b c d n := by + induction n + erw [EllDivSequence_zero, neg_zero] + rfl From ba72afb8d3311ab45d1d3da01d916cf3135d318d Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 1 Jan 2024 01:37:07 +0000 Subject: [PATCH 08/47] Use elliptic divisibility sequences --- Mathlib.lean | 1 + .../EllipticCurve/Torsion.lean | 189 ++++-------------- 2 files changed, 38 insertions(+), 152 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index b25c49263926ad..e2424bc91443f2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -492,6 +492,7 @@ import Mathlib.Algebra.Tropical.Lattice import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.Group +import Mathlib.AlgebraicGeometry.EllipticCurve.Torsion import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass import Mathlib.AlgebraicGeometry.FunctionField import Mathlib.AlgebraicGeometry.GammaSpecAdjunction diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean index 427b90e41486c8..ff6a5eae2bffd8 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean @@ -4,182 +4,67 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.AlgebraicGeometry.EllipticCurve.Point +import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.NumberTheory.EllipticDivisibilitySequence /-! # Torsion points on Weierstrass curves -/ +open Polynomial PolynomialPolynomial + universe u -namespace WeierstrassCurve - -open Polynomial - -open scoped Polynomial PolynomialPolynomial - -variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) - -noncomputable def divisionPolynomial : ℕ → (R[X][Y]) -| 0 => 0 -| 1 => 1 -| 2 => Y - W.negPolynomial -| 3 => C <| 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ -| 4 => (Y - W.negPolynomial) - * C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 - + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) -| (n + 5) => -if hn : Even n then - let m := ((even_iff_exists_two_mul n).mp hn).choose - have h4 : m + 4 < n + 5 := - by linarith only [show n = 2 * m by exact ((even_iff_exists_two_mul n).mp hn).choose_spec] - have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 - have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 - have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 - divisionPolynomial (m + 4) * divisionPolynomial (m + 2) ^ 3 - - divisionPolynomial (m + 1) * divisionPolynomial (m + 3) ^ 3 -else - let m := (Nat.odd_iff_not_even.mpr hn).choose - have h : m < n := - by linarith only [show n = 2 * m + 1 by exact (Nat.odd_iff_not_even.mpr hn).choose_spec] - have h5 : m + 5 < n + 5 := add_lt_add_right h 5 - have h4 : m + 4 < n + 5 := (lt_add_one _).trans h5 - have h3 : m + 3 < n + 5 := (lt_add_one _).trans h4 - have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 - have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 - Ring.divide - (divisionPolynomial (m + 2) ^ 2 * divisionPolynomial (m + 3) * divisionPolynomial (m + 5) - - divisionPolynomial (m + 1) * divisionPolynomial (m + 3) * divisionPolynomial (m + 4) ^ 2) - (Y - W.negPolynomial) +namespace WeierstrassCurve.Affine + +variable {R : Type u} [CommRing R] (W : Affine R) + +noncomputable def divisionPolynomial (n : ℤ) : R[X][Y] := + EllDivSequence (Y - W.negPolynomial) + (C <| 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) + (C <| 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) + n @[simp] lemma divisionPolynomial_zero : W.divisionPolynomial 0 = 0 := - rfl + EllDivSequence_zero _ _ _ @[simp] lemma divisionPolynomial_one : W.divisionPolynomial 1 = 1 := - rfl + EllDivSequence_one _ _ _ @[simp] lemma divisionPolynomial_two : W.divisionPolynomial 2 = Y - W.negPolynomial := - rfl - -@[simp] -lemma divisionPolynomial_three : - W.divisionPolynomial 3 = - C (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) := - rfl + EllDivSequence_two _ _ _ @[simp] -lemma divisionPolynomial_four : - W.divisionPolynomial 4 = - (Y - W.negPolynomial) - * C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 - + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X - + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) := - rfl +lemma divisionPolynomial_three : W.divisionPolynomial 3 = + C (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) := + EllDivSequence_three _ _ _ @[simp] -lemma divisionPolynomial_odd (m : ℕ) : - W.divisionPolynomial (2 * (m + 2) + 1) = - W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 := by - erw [divisionPolynomial, dif_pos <| even_two_mul m] - simp only [← (mul_right_inj' two_ne_zero).mp (⟨m, rfl⟩ : ∃ n : ℕ, 2 * m = 2 * n).choose_spec] +lemma divisionPolynomial_four : W.divisionPolynomial 4 = + C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) * + (Y - W.negPolynomial) := + EllDivSequence_four _ _ _ @[simp] -lemma divisionPolynomial_even (m : ℕ) : - W.divisionPolynomial (2 * (m + 3)) = - Ring.divide - (W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) - * W.divisionPolynomial (m + 5) - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) - * W.divisionPolynomial (m + 4) ^ 2) - (Y - W.negPolynomial) := by - erw [divisionPolynomial, dif_neg <| Nat.odd_iff_not_even.mp <| odd_two_mul_add_one m] - simp only [← (mul_right_inj' two_ne_zero).mp <| - Nat.succ_injective (odd_two_mul_add_one m).choose_spec] - -private -lemma even_rw (y p1 p2 p3 p4 p5 : R) : - (y * p2) ^ 2 * p3 * p5 - p1 * p3 * (y * p4) ^ 2 = - y * (y * (p2 ^ 2 * p3 * p5 - p1 * p3 * p4 ^ 2)) := by - ring1 - -private -lemma odd_rw (y p1 p2 p3 p4 p5 : R) : - p2 ^ 2 * (y * p3) * (y * p5) - (y * p1) * (y * p3) * p4 ^ 2 = - y * (y * (p2 ^ 2 * p3 * p5 - p1 * p3 * p4 ^ 2)) := by - ring1 - -set_option maxHeartbeats 500000 in -lemma dvd_divisionPolynomial_even [IsDomain R] (m : ℕ) : - Y - W.negPolynomial ∣ W.divisionPolynomial (2 * m) := by - suffices ∀ {n : ℕ}, Even n → Y - W.negPolynomial ∣ W.divisionPolynomial n by - exact this <| even_two_mul _ - intro n hn - induction' n using Nat.strongRec' with n ih - rcases hn with ⟨n, rfl⟩ - rcases n with _ | _ | _ | n - · exact ⟨0, by rw [mul_zero, W.divisionPolynomial_zero]⟩ - · exact ⟨1, by rw [mul_one, W.divisionPolynomial_two]⟩ - · exact ⟨_, W.divisionPolynomial_four⟩ - · rw [show n + 3 + (n + 3) = 2 * n + 6 by exact (two_mul <| n + 3).symm] at ih ⊢ - rw [divisionPolynomial] - simp only [dif_neg <| Nat.odd_iff_not_even.mp <| odd_two_mul_add_one n, - ← (mul_right_inj' two_ne_zero).mp <| Nat.succ_injective (odd_two_mul_add_one n).choose_spec] - rcases n.even_or_odd' with ⟨n, rfl | rfl⟩ - · cases' ih (2 * n + 2) (by linarith only) (even_two_mul <| n + 1) with _ h2 - cases' ih (2 * n + 4) (by linarith only) (even_two_mul <| n + 2) with _ h4 - rw [h2, h4, even_rw] - by_cases hY : Y - W.negPolynomial = 0 - · rw [hY, zero_mul, Ring.divide_zero] - · simpa only [Ring.mul_divide_cancel_left hY] using dvd_mul_right _ _ - · cases' ih (2 * n + 1 + 1) (by linarith only) (even_two_mul <| n + 1) with _ h1 - cases' ih (2 * n + 1 + 3) (by linarith only) (even_two_mul <| n + 2) with _ h3 - cases' ih (2 * n + 1 + 5) (by linarith only) (even_two_mul <| n + 3) with _ h5 - rw [h1, h3, h5, odd_rw] - by_cases hY : Y - W.negPolynomial = 0 - · rw [hY, zero_mul, Ring.divide_zero] - · simpa only [Ring.mul_divide_cancel_left hY] using dvd_mul_right _ _ - -lemma mul_divisionPolynomial_even [IsDomain R] (m : ℕ) : - (Y - W.negPolynomial) * W.divisionPolynomial (2 * (m + 3)) = - W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * W.divisionPolynomial (m + 5) - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) - * W.divisionPolynomial (m + 4) ^ 2 := by - rw [divisionPolynomial_even] - rcases m.even_or_odd' with ⟨m, rfl | rfl⟩ - · rw [show 2 * m + 2 = 2 * (m + 1) by rfl, show 2 * m + 4 = 2 * (m + 2) by rfl, - (W.dvd_divisionPolynomial_even $ m + 1).choose_spec, - (W.dvd_divisionPolynomial_even $ m + 2).choose_spec, even_rw] - by_cases hY : Y - W.negPolynomial = 0 - · simp only [hY, zero_mul] - · rw [Ring.mul_divide_cancel hY] - exact dvd_mul_right _ _ - · rw [show 2 * m + 1 + 1 = 2 * (m + 1) by rfl, show 2 * m + 1 + 3 = 2 * (m + 2) by rfl, - show 2 * m + 1 + 5 = 2 * (m + 3) by rfl, - (W.dvd_divisionPolynomial_even $ m + 1).choose_spec, - (W.dvd_divisionPolynomial_even $ m + 2).choose_spec, - (W.dvd_divisionPolynomial_even $ m + 3).choose_spec, odd_rw] - by_cases hY : Y - W.negPolynomial = 0 - · simp only [hY, zero_mul] - · rw [Ring.mul_divide_cancel hY] - exact dvd_mul_right _ _ - -noncomputable -def divisionPolynomial' : ℤ → (R[X][Y]) -| Int.ofNat n => W.divisionPolynomial n -| Int.negSucc n => -W.divisionPolynomial (n + 1) +lemma divisionPolynomial_odd (m : ℕ) : W.divisionPolynomial (2 * (m + 2) + 1) = + W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 - + W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 := + EllDivSequence_odd _ _ _ m @[simp] -lemma divisionPolynomial'_nat (n : ℕ) : W.divisionPolynomial' n = W.divisionPolynomial n := - rfl +lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) * (Y - W.negPolynomial) = + W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * W.divisionPolynomial (m + 5) - + W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) * + W.divisionPolynomial (m + 4) ^ 2 := + EllDivSequence_even _ _ _ m @[simp] -lemma divisionPolynomial'_neg (n : ℕ) : W.divisionPolynomial' (-n) = -W.divisionPolynomial n := by - cases n - · erw [Int.neg_zero, divisionPolynomial'_nat, W.divisionPolynomial_zero, neg_zero] - · rfl +lemma divisionPolynomial_neg (n : ℕ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := + EllDivSequence_neg _ _ _ n -end WeierstrassCurve +end WeierstrassCurve.Affine From bb1bf7e95654637a90ba11ed2113cc9181275116 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Tue, 2 Jan 2024 22:51:43 +0000 Subject: [PATCH 09/47] Implement Jacobian coordinates --- Mathlib.lean | 1 + .../EllipticCurve/Jacobian.lean | 982 ++++++++++++++++++ 2 files changed, 983 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean diff --git a/Mathlib.lean b/Mathlib.lean index 72f713857128d4..65cf3ba1d377be 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -492,6 +492,7 @@ import Mathlib.Algebra.Tropical.Lattice import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.Group +import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass import Mathlib.AlgebraicGeometry.FunctionField import Mathlib.AlgebraicGeometry.GammaSpecAdjunction diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean new file mode 100644 index 00000000000000..563ee8c86a817f --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -0,0 +1,982 @@ +/- +Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ +import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.Data.MvPolynomial.CommRing + +/-! +# Jacobian coordinates for Weierstrass curves + +This file defines the type of points on a Weierstrass curve as a tuple, consisting of a equivalence +class of triples up to scaling by weights, satisfying a Weierstrass equation with a nonsingular +condition. This file also defines the negation and addition operations of the group law for this +type, and proves that they respect the Weierstrass equation and the nonsingular condition. The fact +that they form an abelian group is proven in `Mathlib.AlgebraicGeometry.EllipticCurve.Group`. + +## Mathematical background + +Let `W` be a Weierstrass curve over a field `F`. A point on the weighted projective plane with +weights $(2, 3, 1)$ is an equivalence class of triples $[x:y:z]$ with coordinates in `F` such that +$(x, y, z) \sim (x', y', z')$ precisely if there is some unit $u$ of `F` such that +$(x, y, z) = (u^2x', u^3y', uz')$, with an extra condition that $(x, y, z) \ne (0, 0, 0)$. +A rational point is a point on the $(2, 3, 1)$-projective plane satisfying a $(2, 3, 1)$-homogeneous +Weierstrass equation $Y^2 + a_1XYZ + a_3YZ^3 = X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6$, and being +nonsingular means the partial derivatives $W_X(X, Y, Z)$, $W_Y(X, Y, Z)$, and $W_Z(X, Y, Z)$ do not +vanish simultaneously. Note that the vanishing of the Weierstrass equation and its partial +derivatives are independent of the representative for $[x:y:z]$, and the nonsingularity condition +already implies that $(x, y, z) \ne (0, 0, 0)$, so a nonsingular rational point on `W` can simply be +given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any representative. +In cryptography, as well as in this file, this is often called the Jacobian coordinates of `W`. + +As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms +an abelian group under the same secant-and-tangent process, but the polynomials involved are +$(2, 3, 1)$-homogeneous, and any instances of division become multiplication in the $Z$-coordinate. +Note that most computational proofs follow from their analogous proofs for affine coordinates. + +## Main definitions + + * `WeierstrassCurve.Jacobian.PointClass`: the equivalence class of a point representative. + * `WeierstrassCurve.Jacobian.toAffine`: the Weierstrass curve in affine coordinates. + * `WeierstrassCurve.Jacobian.nonsingular`: the nonsingular condition on a point representative. + * `WeierstrassCurve.Jacobian.nonsingular_lift`: the nonsingular condition on a point class. + * `WeierstrassCurve.Jacobian.neg`: the negation operation on a point representative. + * `WeierstrassCurve.Jacobian.neg_map`: the negation operation on a point class. + * `WeierstrassCurve.Jacobian.add`: the addition operation on a point representative. + * `WeierstrassCurve.Jacobian.add_map`: the addition operation on a point class. + * `WeierstrassCurve.Jacobian.Point`: a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.neg`: the negation operation on a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.add`: the addition operation on a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.toAffine_addEquiv`: the equivalence between the nonsingular + rational points on a Jacobian Weierstrass curve with those on an affine Weierstrass curve. + +## Main statements + + * `WeierstrassCurve.Jacobian.nonsingular_neg`: negation preserves the nonsingular condition. + * `WeierstrassCurve.Jacobian.nonsingular_add`: addition preserves the nonsingular condition. + +## Implementation notes + +A point representative is implemented as a term `P` of type `Fin 3 → R`, which allows for the vector +notation `![x, y, z]`. However, `P` is not definitionally equivalent to the expanded vector +`![P x, P y, P z]`, so the auxiliary lemma `fin3_def` can be used to convert between the two forms. +The equivalence of two point representatives `P` and `Q` is implemented as an equivalence of orbits +of the action of `Rˣ`, or equivalently that there is some unit `u` of `R` such that `P = u • Q`. +However, `u • Q` is again not definitionally equal to `![u² * Q x, u³ * Q y, u * Q z]`, so the +auxiliary lemmas `smul_fin3` and `smul_fin3_ext` can be used to convert between the two forms. + +## References + +[J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, rational point, Jacobian coordinates +-/ + +local notation "x" => 0 + +local notation "y" => 1 + +local notation "z" => 2 + +local macro "matrix_simp" : tactic => + `(tactic| simp only [Matrix.head_cons, Matrix.tail_cons, Matrix.smul_empty, Matrix.smul_cons, + Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.cons_val_two]) + +universe u + +lemma fin3_def {R : Type u} (P : Fin 3 → R) : P = ![P x, P y, P z] := by + ext n; fin_cases n <;> rfl + +private instance {R : Type u} [CommRing R] : SMul Rˣ <| Fin 3 → R := + ⟨fun u P => ![u ^ 2 * P x, u ^ 3 * P y, u * P z]⟩ + +lemma smul_fin3 {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : + u • P = ![u ^ 2 * P x, u ^ 3 * P y, u * P z] := + rfl + +lemma smul_fin3_ext {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : + (u • P) 0 = u ^ 2 * P x ∧ (u • P) 1 = u ^ 3 * P y ∧ (u • P) 2 = u * P z := + ⟨rfl, rfl, rfl⟩ + +private instance {R : Type u} [CommRing R] : MulAction Rˣ <| Fin 3 → R where + one_smul := fun _ => by + simp only [smul_fin3, Units.val_one, one_pow, one_mul, ← fin3_def] + mul_smul := fun u v P => by + simp only [smul_fin3, Units.val_mul, mul_pow, mul_assoc] + matrix_simp + +/-! ## Weierstrass curves -/ + +/-- An abbreviation for a Weierstrass curve in Jacobian coordinates. -/ +abbrev WeierstrassCurve.Jacobian := + WeierstrassCurve + +namespace WeierstrassCurve.Jacobian + +open MvPolynomial + +local macro "eval_simp" : tactic => + `(tactic| simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, eval_pow]) + +variable (R : Type u) [CommRing R] + +/-- The equivalence setoid for a point representative. -/ +def PointSetoid : Setoid <| Fin 3 → R := + MulAction.orbitRel Rˣ <| Fin 3 → R + +attribute [local instance] PointSetoid + +/-- The equivalence class of a point representative. -/ +abbrev PointClass : Type u := + MulAction.orbitRel.Quotient Rˣ <| Fin 3 → R + +variable {R} (W : Jacobian R) + +/-- The coercion to a Weierstrass curve in affine coordinates. -/ +@[pp_dot] +abbrev toAffine : Affine R := + W + +section Equation + +/-! ### Equations and nonsingularity -/ + +/-- The polynomial $W(X, Y, Z) := Y^2 + a_1XYZ + a_3YZ^3 - (X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6)$ +associated to a Weierstrass curve `W` over `R`. This is represented as a term of type +`MvPolynomial (Fin 3) R`, where `X 0`, `X 1`, and `X 2` represent $X$, $Y$, and $Z$ respectively. -/ +@[pp_dot] +noncomputable def polynomial : MvPolynomial (Fin 3) R := + X 1 ^ 2 + C W.a₁ * X 0 * X 1 * X 2 + C W.a₃ * X 1 * X 2 ^ 3 + - (X 0 ^ 3 + C W.a₂ * X 0 ^ 2 * X 2 ^ 2 + C W.a₄ * X 0 * X 2 ^ 4 + C W.a₆ * X 2 ^ 6) + +lemma eval_polynomial (P : Fin 3 → R) : eval P W.polynomial = + P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 + - (P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6) := by + rw [polynomial] + eval_simp + +/-- The proposition that a point representative $(x, y, z)$ lies in `W`. +In other words, $W(x, y, z) = 0$. -/ +@[pp_dot] +def equation (P : Fin 3 → R) : Prop := + eval P W.polynomial = 0 + +lemma equation_iff (P : Fin 3 → R) : W.equation P ↔ + P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 + = P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6 := by + rw [equation, eval_polynomial, sub_eq_zero] + +lemma equation_zero : W.equation ![1, 1, 0] := + (W.equation_iff ![1, 1, 0]).mpr <| by matrix_simp; ring1 + +lemma equation_zero' (Y : R) : W.equation ![Y ^ 2, Y ^ 3, 0] := + (W.equation_iff ![Y ^ 2, Y ^ 3, 0]).mpr <| by matrix_simp; ring1 + +lemma equation_some (X Y : R) : W.equation ![X, Y, 1] ↔ W.toAffine.equation X Y := by + rw [equation_iff, W.toAffine.equation_iff] + congr! 1 <;> matrix_simp <;> ring1 + +lemma equation_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.equation (u • P) ↔ W.equation P := + have (u : Rˣ) {P : Fin 3 → R} (h : W.equation P) : W.equation <| u • P := by + rw [equation_iff] at h ⊢ + linear_combination (norm := (simp only [smul_fin3_ext]; ring1)) (u : R) ^ 6 * h + ⟨fun h => by convert this u⁻¹ h; rw [inv_smul_smul], this u⟩ + +/-- The partial derivative $W_X(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $X$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialX : MvPolynomial (Fin 3) R := + C W.a₁ * X 1 * X 2 - (C 3 * X 0 ^ 2 + C (2 * W.a₂) * X 0 * X 2 ^ 2 + C W.a₄ * X 2 ^ 4) + +lemma eval_polynomialX (P : Fin 3 → R) : eval P W.polynomialX = + W.a₁ * P y * P z - (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4) := by + rw [polynomialX] + eval_simp + +/-- The partial derivative $W_Y(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Y$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialY : MvPolynomial (Fin 3) R := + C 2 * X 1 + C W.a₁ * X 0 * X 2 + C W.a₃ * X 2 ^ 3 + +lemma eval_polynomialY (P : Fin 3 → R) : + eval P W.polynomialY = 2 * P y + W.a₁ * P x * P z + W.a₃ * P z ^ 3 := by + rw [polynomialY] + eval_simp + +/-- The partial derivative $W_Z(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Z$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialZ : MvPolynomial (Fin 3) R := + C W.a₁ * X 0 * X 1 + C (3 * W.a₃) * X 1 * X 2 ^ 2 + - (C (2 * W.a₂) * X 0 ^ 2 * X 2 + C (4 * W.a₄) * X 0 * X 2 ^ 3 + C (6 * W.a₆) * X 2 ^ 5) + +lemma eval_polynomialZ (P : Fin 3 → R) : eval P W.polynomialZ = + W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 + - (2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by + rw [polynomialZ] + eval_simp + +/-- The proposition that a point representative $(x, y, z)$ in `W` is nonsingular. +In other words, either $W_X(x, y, z) \ne 0$, $W_Y(x, y, z) \ne 0$, or $W_Z(x, y, z) \ne 0$. -/ +@[pp_dot] +def nonsingular (P : Fin 3 → R) : Prop := + W.equation P ∧ (eval P W.polynomialX ≠ 0 ∨ eval P W.polynomialY ≠ 0 ∨ eval P W.polynomialZ ≠ 0) + +lemma nonsingular_iff (P : Fin 3 → R) : W.nonsingular P ↔ W.equation P ∧ + (W.a₁ * P y * P z ≠ 3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 ∨ + P y ≠ -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 ∨ + W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 + ≠ 2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by + rw [nonsingular, eval_polynomialX, eval_polynomialY, eval_polynomialZ, sub_ne_zero, sub_ne_zero, + ← sub_ne_zero (a := P y)] + congr! 4 + ring1 + +lemma nonsingular_zero [Nontrivial R] : W.nonsingular ![1, 1, 0] := + (W.nonsingular_iff ![1, 1, 0]).mpr ⟨W.equation_zero, + by simp; by_contra! h; exact one_ne_zero <| by linear_combination -h.1 - h.2.1⟩ + +lemma nonsingular_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : + W.nonsingular ![Y ^ 2, Y ^ 3, 0] := + (W.nonsingular_iff ![Y ^ 2, Y ^ 3, 0]).mpr ⟨W.equation_zero' Y, + by simp [hy]; by_contra! h; exact pow_ne_zero 3 hy <| by linear_combination Y ^ 3 * h.1 - h.2.1⟩ + +lemma nonsingular_some (X Y : R) : W.nonsingular ![X, Y, 1] ↔ W.toAffine.nonsingular X Y := by + rw [nonsingular_iff] + matrix_simp + simp only [W.toAffine.nonsingular_iff, equation_some, and_congr_right_iff, + W.toAffine.equation_iff, ← not_and_or, not_iff_not, one_pow, mul_one, Iff.comm, iff_self_and] + intro h hX hY + linear_combination (norm := ring1) 6 * h - 2 * X * hX - 3 * Y * hY + +lemma nonsingular_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.nonsingular (u • P) ↔ W.nonsingular P := + have (u : Rˣ) {P : Fin 3 → R} (h : W.nonsingular <| u • P) : W.nonsingular P := by + rcases (W.nonsingular_iff _).mp h with ⟨h, h'⟩ + refine (W.nonsingular_iff P).mpr ⟨(W.equation_smul_iff P u).mp h, ?_⟩ + contrapose! h' + simp only [smul_fin3_ext] + exact ⟨by linear_combination (norm := ring1) (u : R) ^ 4 * h'.left, + by linear_combination (norm := ring1) (u : R) ^ 3 * h'.right.left, + by linear_combination (norm := ring1) (u : R) ^ 5 * h'.right.right⟩ + ⟨this u, fun h => this u⁻¹ <| by rwa [inv_smul_smul]⟩ + +lemma nonsingular_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.nonsingular P ↔ W.nonsingular Q := by + rcases h with ⟨u, rfl⟩ + exact W.nonsingular_smul_iff Q u + +/-- The proposition that a point class on `W` is nonsingular. If `P` is a point representative, +then `W.nonsingular_lift ⟦P⟧` is definitionally equivalent to `W.nonsingular P`. -/ +@[pp_dot] +def nonsingular_lift (P : PointClass R) : Prop := + P.lift W.nonsingular fun _ _ => propext ∘ W.nonsingular_of_equiv + +@[simp] +lemma nonsingular_lift_eq (P : Fin 3 → R) : W.nonsingular_lift ⟦P⟧ = W.nonsingular P := + rfl + +lemma nonsingular_lift_zero [Nontrivial R] : W.nonsingular_lift ⟦![1, 1, 0]⟧ := + W.nonsingular_zero + +lemma nonsingular_lift_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : + W.nonsingular_lift ⟦![Y ^ 2, Y ^ 3, 0]⟧ := + W.nonsingular_zero' hy + +lemma nonsingular_lift_some (X Y : R) : + W.nonsingular_lift ⟦![X, Y, 1]⟧ ↔ W.toAffine.nonsingular X Y := + W.nonsingular_some X Y + +variable {F : Type u} [Field F] {W : Jacobian F} + +lemma equiv_of_Zeq0 {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z = 0) + (hQz : Q z = 0) : P ≈ Q := by + rw [fin3_def P, hPz] at hP ⊢ + rw [fin3_def Q, hQz] at hQ ⊢ + simp [nonsingular_iff, equation_iff] at hP hQ + have hPx : P x ≠ 0 := fun h => by simp [h] at hP; simp [hP] at hP + have hQx : Q x ≠ 0 := fun h => by simp [h] at hQ; simp [hQ] at hQ + have hPy : P y ≠ 0 := fun h => by simp [h] at hP; exact hPx <| pow_eq_zero hP.left.symm + have hQy : Q y ≠ 0 := fun h => by simp [h] at hQ; exact hQx <| pow_eq_zero hQ.left.symm + use Units.mk0 _ <| mul_ne_zero (div_ne_zero hPy hPx) (div_ne_zero hQx hQy) + simp [smul_fin3, mul_pow, div_pow] + congr! 2 + · field_simp [hP.left, hQ.left] + ring1 + · field_simp [← hP.left, ← hQ.left] + ring1 + +lemma equiv_zero_of_Zeq0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z = 0) : P ≈ ![1, 1, 0] := + equiv_of_Zeq0 h W.nonsingular_zero hPz rfl + +lemma equiv_some_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : P ≈ ![P x / P z ^ 2, P y / P z ^ 3, 1] := + ⟨Units.mk0 _ hPz, by simp [smul_fin3, ← fin3_def P, mul_div_cancel' _ <| pow_ne_zero _ hPz]⟩ + +lemma nonsingular_iff_affine_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.nonsingular P ↔ W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := + (W.nonsingular_of_equiv <| equiv_some_of_Zne0 hPz).trans <| W.nonsingular_some .. + +lemma nonsingular_of_affine_of_Zne0 {P : Fin 3 → F} + (h : W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3)) (hPz : P z ≠ 0) : + W.nonsingular P := + (nonsingular_iff_affine_of_Zne0 hPz).mpr h + +lemma nonsingular_affine_of_Zne0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z ≠ 0) : + W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := + (nonsingular_iff_affine_of_Zne0 hPz).mp h + +end Equation + +section Polynomial + +/-! ### Group operation polynomials -/ + +/-- The $Y$-coordinate of the negation of a point representative. -/ +@[pp_dot] +def negY (P : Fin 3 → R) : R := + -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 + +lemma negY_smul (P : Fin 3 → R) (u : Rˣ) : W.negY (u • P) = u ^ 3 * W.negY P := by + simp only [negY, smul_fin3_ext] + ring1 + +/-- The $X$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addX_of_Xne (P Q : Fin 3 → R) : R := + (P y * Q z ^ 3 - P z ^ 3 * Q y) ^ 2 + + W.a₁ * P z * Q z * (P y * Q z ^ 3 - P z ^ 3 * Q y) * (P x * Q z ^ 2 - P z ^ 2 * Q x) + - (W.a₂ * P z ^ 2 * Q z ^ 2 + P x * Q z ^ 2 + P z ^ 2 * Q x) + * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2 + +lemma addX_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addX_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addX_of_Xne P Q := by + simp only [addX_of_Xne, smul_fin3_ext] + ring1 + +/-- The $X$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addX_of_Yne (P : Fin 3 → R) : R := + (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) ^ 2 + + W.a₁ * (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) + * (P y * P z - W.negY P * P z) + - (W.a₂ * P z ^ 2 + 2 * P x) * (P y - W.negY P) ^ 2 + +lemma addX_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addX_of_Yne (u • P) = (u : R) ^ 8 * W.addX_of_Yne P := by + simp only [addX_of_Yne, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the addition of two point representatives, before applying the final +negation that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where their $Z$-coordinates are non-zero and their +$X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addY'_of_Xne (P Q : Fin 3 → R) : R := + (P y * Q z ^ 3 - P z ^ 3 * Q y) + * (W.addX_of_Xne P Q - P x * Q z ^ 2 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) + + P y * Q z ^ 3 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 3 + +lemma addY'_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addY'_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY'_of_Xne P Q := by + simp only [addY'_of_Xne, addX_of_Xne_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the doubling of a point representative, before applying the final negation +that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where its $Z$-coordinate is non-zero and its $Y$-coordinate +is distinct from that of its negation. -/ +@[pp_dot] +def addY'_of_Yne (P : Fin 3 → R) : R := + (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) + * (W.addX_of_Yne P - P x * (P y - W.negY P) ^ 2) + + P y * (P y - W.negY P) ^ 3 + +lemma addY'_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addY'_of_Yne (u • P) = (u : R) ^ 12 * W.addY'_of_Yne P := by + simp only [addY'_of_Yne, addX_of_Yne_smul, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Z$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +def addZ_of_Xne (P Q : Fin 3 → R) : R := + P z * Q z * (P x * Q z ^ 2 - P z ^ 2 * Q x) + +lemma addZ_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + addZ_of_Xne (u • P) (v • Q) = (u : R) ^ 3 * (v : R) ^ 3 * addZ_of_Xne P Q := by + simp only [addZ_of_Xne, smul_fin3_ext] + ring1 + +/-- The $Z$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addZ_of_Yne (P : Fin 3 → R) : R := + P z * (P y - W.negY P) + +lemma addZ_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addZ_of_Yne (u • P) = (u : R) ^ 4 * W.addZ_of_Yne P := by + simp only [addZ_of_Yne, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addY_of_Xne (P Q : Fin 3 → R) : R := + W.negY ![W.addX_of_Xne P Q, W.addY'_of_Xne P Q, addZ_of_Xne P Q] + +lemma addY_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addY_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY_of_Xne P Q := by + simp only [addY_of_Xne, negY, addX_of_Xne_smul, addY'_of_Xne_smul, addZ_of_Xne_smul] + matrix_simp + ring1 + +/-- The $Y$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addY_of_Yne (P : Fin 3 → R) : R := + W.negY ![W.addX_of_Yne P, W.addY'_of_Yne P, W.addZ_of_Yne P] + +lemma addY_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addY_of_Yne (u • P) = (u : R) ^ 12 * W.addY_of_Yne P := by + simp only [addY_of_Yne, negY, addX_of_Yne_smul, addY'_of_Yne_smul, addZ_of_Yne_smul] + matrix_simp + ring1 + +variable {F : Type u} [Field F] {W : Jacobian F} + +lemma negY_divZ {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.negY P / P z ^ 3 = W.toAffine.negY (P x / P z ^ 2) (P y / P z ^ 3) := by + field_simp [negY, Affine.negY] + ring1 + +lemma Yne_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + P y ≠ W.negY P := by + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + have hx' : P x * (P z / P z ^ 3) = Q x * (Q z / Q z ^ 3) := by + simp_rw [pow_succ _ 2, div_mul_right _ hPz, div_mul_right _ hQz, mul_one_div, hx] + have hy' : P y / P z ^ 3 = Q y / Q z ^ 3 := + Affine.Yeq_of_Yne (nonsingular_affine_of_Zne0 hP hPz).left + (nonsingular_affine_of_Zne0 hQ hQz).left hx <| (negY_divZ hQz).symm ▸ hy + simp_rw [negY, sub_div, neg_div, mul_div_assoc, mul_assoc, ← hy', ← hx', + div_self <| pow_ne_zero 3 hQz, ← div_self <| pow_ne_zero 3 hPz, ← mul_assoc, ← mul_div_assoc, + ← neg_div, ← sub_div, div_left_inj' <| pow_ne_zero 3 hPz] at hy + exact hy + +lemma addX_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addX_of_Xne P Q / addZ_of_Xne P Q ^ 2 = + W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.slope_of_Xne <| + by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] + field_simp [sub_ne_zero_of_ne hx, addX_of_Xne, addZ_of_Xne] + ring1 + +lemma addX_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.addX_of_Yne P / W.addZ_of_Yne P ^ 2 = W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← hx, ← negY_divZ hPz] + field_simp [addX_of_Yne, addX_of_Yne, addZ_of_Yne] + ring1 + +lemma addY'_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY'_of_Xne P Q / addZ_of_Xne P Q ^ 3 = + W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY', ← addX_div_addZ_of_Xne hPz hQz hx, Affine.slope_of_Xne <| + by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] + field_simp [sub_ne_zero_of_ne hx, addY'_of_Xne, addX_of_Xne, addZ_of_Xne] + ring1 + +lemma addY'_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY'_of_Yne P / W.addZ_of_Yne P ^ 3 = + W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + rw [Affine.addY', ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy] + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← negY_divZ hPz] + field_simp [addY'_of_Yne, addX_of_Yne, addZ_of_Yne] + ring1 + +lemma addZ_ne_zero_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : addZ_of_Xne P Q ≠ 0 := + mul_ne_zero (mul_ne_zero hPz hQz) <| sub_ne_zero_of_ne hx + +lemma addZ_ne_zero_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addZ_of_Yne P ≠ 0 := + mul_ne_zero hPz <| sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + +lemma addY_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY_of_Xne P Q / addZ_of_Xne P Q ^ 3 = + W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY, ← addX_div_addZ_of_Xne hPz hQz hx, ← addY'_div_addZ_of_Xne hPz hQz hx] + exact negY_divZ <| addZ_ne_zero_of_Xne hPz hQz hx + +lemma addY_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY_of_Yne P / W.addZ_of_Yne P ^ 3 = + W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY, ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, + ← addY'_div_addZ_of_Yne hP hQ hPz hQz hx hy] + exact negY_divZ <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy + +end Polynomial + +section Representative + +/-! ### Group operations on point representatives -/ + +/-- The negation of a point representative. -/ +@[pp_dot] +def neg (P : Fin 3 → R) : Fin 3 → R := + ![P x, W.negY P, P z] + +@[simp] +lemma neg_zero : W.neg ![1, 1, 0] = ![1, -1, 0] := by + erw [neg, negY, mul_zero, zero_pow three_pos, mul_zero, sub_zero, sub_zero] + rfl + +@[simp] +lemma neg_some (X Y : R) : W.neg ![X, Y, 1] = ![X, -Y - W.a₁ * X - W.a₃, 1] := by + erw [neg, negY, mul_one, one_pow, mul_one] + rfl + +lemma neg_smul_equiv (P : Fin 3 → R) (u : Rˣ) : W.neg (u • P) ≈ W.neg P := + ⟨u, by simp_rw [neg, negY_smul, smul_fin3]; rfl⟩ + +lemma neg_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.neg P ≈ W.neg Q := by + rcases h with ⟨u, rfl⟩ + exact W.neg_smul_equiv Q u + +/-- The negation of a point class. If `P` is a point representative, +then `W.neg_map ⟦P⟧` is definitionally equivalent to `W.neg P`. -/ +@[pp_dot] +def neg_map (P : PointClass R) : PointClass R := + P.map W.neg fun _ _ => W.neg_equiv + +lemma neg_map_eq {P : Fin 3 → R} : W.neg_map ⟦P⟧ = ⟦W.neg P⟧ := + rfl + +@[simp] +lemma neg_map_zero : W.neg_map ⟦![1, 1, 0]⟧ = ⟦![1, 1, 0]⟧ := by + simpa only [neg_map_eq, neg_zero, Quotient.eq] using ⟨-1, by norm_num [smul_fin3]⟩ + +@[simp] +lemma neg_map_some (X Y : R) : W.neg_map ⟦![X, Y, 1]⟧ = ⟦![X, -Y - W.a₁ * X - W.a₃, 1]⟧ := by + rw [neg_map_eq, neg_some] + +open scoped Classical + +/-- The addition of two point representatives. -/ +@[pp_dot] +noncomputable def add (P Q : Fin 3 → R) : Fin 3 → R := + if P z = 0 then Q else if Q z = 0 then P else if P x * Q z ^ 2 = P z ^ 2 * Q x then + if P y * Q z ^ 3 = P z ^ 3 * W.negY Q then ![1, 1, 0] else + ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] + else ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] + +@[simp] +lemma add_of_Zeq0_left {P Q : Fin 3 → R} (hPz : P z = 0) : W.add P Q = Q := + if_pos hPz + +lemma add_zero_left (P : Fin 3 → R) : W.add ![1, 1, 0] P = P := + W.add_of_Zeq0_left rfl + +@[simp] +lemma add_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : W.add P Q = P := by + rw [add, if_neg hPz, if_pos hQz] + +lemma add_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add P ![1, 1, 0] = P := + W.add_of_Zeq0_right hPz rfl + +@[simp] +lemma add_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : + W.add P Q = ![1, 1, 0] := by + rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_pos hy] + +@[simp] +lemma add_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.add P Q = ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] := by + rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_neg hy] + +@[simp] +lemma add_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.add P Q = ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] := by + rw [add, if_neg hPz, if_neg hQz, if_neg hx] + +variable [IsDomain R] + +lemma add_smul_equiv (P Q : Fin 3 → R) (u v : Rˣ) : W.add (u • P) (v • Q) ≈ W.add P Q := by + have huv (n : ℕ) : (u ^ n * v ^ n : R) ≠ 0 := + mul_ne_zero (pow_ne_zero n u.ne_zero) (pow_ne_zero n v.ne_zero) + by_cases hPz : P z = 0 + · exact ⟨v, by rw [W.add_of_Zeq0_left hPz, + W.add_of_Zeq0_left <| by simp only [smul_fin3_ext, hPz, mul_zero]]⟩ + · have huz : u * P z ≠ 0 := mul_ne_zero u.ne_zero hPz + by_cases hQz : Q z = 0 + · rw [W.add_of_Zeq0_right hPz hQz, + W.add_of_Zeq0_right huz <| by simp only [smul_fin3_ext, hQz, mul_zero]] + exact ⟨u, rfl⟩ + · have hvz : v * Q z ≠ 0 := mul_ne_zero v.ne_zero hQz + by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · rw [W.add_of_Yeq huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by + simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm, hy], + W.add_of_Yeq hPz hQz hx hy] + · rw [W.add_of_Yne huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by + simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm] + exact hy ∘ mul_left_cancel₀ (huv 3), + addX_of_Yne_smul, addY_of_Yne_smul, addZ_of_Yne_smul, W.add_of_Yne hPz hQz hx hy] + exact ⟨u ^ 4, by simp only [smul_fin3, ← Units.val_pow_eq_pow_val, ← pow_mul]; rfl⟩ + · rw [W.add_of_Xne huz hvz <| by + simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm]; exact hx ∘ mul_left_cancel₀ (huv 2), + addX_of_Xne_smul, addY_of_Xne_smul, addZ_of_Xne_smul, W.add_of_Xne hPz hQz hx] + exact ⟨u ^ 3 * v ^ 3, + by simp_rw [smul_fin3, ← Units.val_pow_eq_pow_val, mul_pow, ← pow_mul]; rfl⟩ + +lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : W.add P Q ≈ W.add P' Q' := by + rcases hP, hQ with ⟨⟨u, rfl⟩, ⟨v, rfl⟩⟩ + exact W.add_smul_equiv P' Q' u v + +/-- The addition of two point classes. If `P` is a point representative, +then `W.add_map ⟦P⟧ ⟦Q⟧` is definitionally equivalent to `W.add P Q`. -/ +@[pp_dot] +noncomputable def add_map (P Q : PointClass R) : PointClass R := + Quotient.map₂ W.add (fun _ _ hP _ _ hQ => W.add_equiv hP hQ) P Q + +lemma add_map_eq (P Q : Fin 3 → R) : W.add_map ⟦P⟧ ⟦Q⟧ = ⟦W.add P Q⟧ := + rfl + +@[simp] +lemma add_map_of_Zeq0_left {P : Fin 3 → R} {Q : PointClass R} (hPz : P z = 0) : + W.add_map ⟦P⟧ Q = Q := by + rcases Q with ⟨Q⟩ + erw [add_map_eq, W.add_of_Zeq0_left hPz] + rfl + +lemma add_map_zero_left (P : PointClass R) : W.add_map ⟦![1, 1, 0]⟧ P = P := + W.add_map_of_Zeq0_left rfl + +@[simp] +lemma add_map_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦P⟧ := by + rw [add_map_eq, W.add_of_Zeq0_right hPz hQz] + +lemma add_map_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add_map ⟦P⟧ ⟦![1, 1, 0]⟧ = ⟦P⟧ := by + rw [add_map_eq, W.add_zero_right hPz] + +@[simp] +lemma add_map_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![1, 1, 0]⟧ := by + rw [add_map_eq, W.add_of_Yeq hPz hQz hx hy] + +@[simp] +lemma add_map_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P]⟧ := by + rw [add_map_eq, W.add_of_Yne hPz hQz hx hy] + +@[simp] +lemma add_map_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q]⟧ := by + rw [add_map_eq, W.add_of_Xne hPz hQz hx] + +variable {F : Type u} [Field F] {W : Jacobian F} + +@[simp] +lemma add_map_of_Zeq0_right' {P : PointClass F} {Q : Fin 3 → F} (hP : W.nonsingular_lift P) + (hQ : W.nonsingular Q) (hQz : Q z = 0) : W.add_map P ⟦Q⟧ = P := by + rcases P with ⟨P⟩ + by_cases hPz : P z = 0 + · erw [W.add_map_of_Zeq0_left hPz, Quotient.eq] + exact equiv_of_Zeq0 hQ hP hQz hPz + · exact W.add_map_of_Zeq0_right hPz hQz + +lemma add_map_zero_right' {P : PointClass F} (hP : W.nonsingular_lift P) : + W.add_map P ⟦![1, 1, 0]⟧ = P := + add_map_of_Zeq0_right' hP W.nonsingular_zero rfl + +variable {F : Type u} [Field F] {W : Jacobian F} + +/-- The negation of a nonsingular point representative in `W` lies in `W`. -/ +lemma nonsingular_neg {P : Fin 3 → F} (h : W.nonsingular P) : W.nonsingular <| W.neg P := by + by_cases hPz : P z = 0 + · rw [W.nonsingular_of_equiv <| W.neg_equiv <| equiv_zero_of_Zeq0 h hPz, neg_zero] + convert W.nonsingular_zero' <| neg_ne_zero.mpr one_ne_zero <;> norm_num1 + · rw [nonsingular_iff_affine_of_Zne0 <| by exact hPz] at h ⊢ + rwa [← Affine.nonsingular_neg_iff, ← negY_divZ hPz] at h + +lemma nonsingular_lift_neg_map {P : PointClass F} (h : W.nonsingular_lift P) : + W.nonsingular_lift <| W.neg_map P := by + rcases P with ⟨_⟩ + exact nonsingular_neg h + +/-- The addition of two nonsingular point representatives in `W` lies in `W`. -/ +lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : + W.nonsingular <| W.add P Q := by + by_cases hPz : P z = 0 + · rwa [W.nonsingular_of_equiv <| W.add_equiv (equiv_zero_of_Zeq0 hP hPz) <| Setoid.refl Q, + W.add_of_Zeq0_left <| by exact rfl] + · by_cases hQz : Q z = 0 + · rwa [W.nonsingular_of_equiv <| W.add_equiv (Setoid.refl P) <| equiv_zero_of_Zeq0 hQ hQz, + W.add_of_Zeq0_right hPz <| by exact rfl] + · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · simpa only [W.add_of_Yeq hPz hQz hx hy] using W.nonsingular_zero + · erw [W.add_of_Yne hPz hQz hx hy, + nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, + addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy] + exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) + (nonsingular_affine_of_Zne0 hQ hQz) fun _ => (negY_divZ hQz).symm ▸ Function.comp + (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp + · erw [W.add_of_Xne hPz hQz hx, + nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, + addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx] + exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) + (nonsingular_affine_of_Zne0 hQ hQz) fun h => False.elim <| hx <| + mul_comm (Q x) _ ▸ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp h + +lemma nonsingular_lift_add_map {P Q : PointClass F} (hP : W.nonsingular_lift P) + (hQ : W.nonsingular_lift Q) : W.nonsingular_lift <| W.add_map P Q := by + rcases P, Q with ⟨⟨_⟩, ⟨_⟩⟩ + exact nonsingular_add hP hQ + +end Representative + +section Point + +/-! ### Group operations on nonsingular rational points -/ + +/-- A nonsingular rational point on `W`. -/ +@[pp_dot] +structure Point where + {point : PointClass R} + (nonsingular : W.nonsingular_lift point) + +attribute [pp_dot] Point.point +attribute [pp_dot] Point.nonsingular + +/-- The point class underlying a nonsingular rational point on `W`. -/ +add_decl_doc Point.point + +/-- The nonsingular condition underlying a nonsingular rational point on `W`. -/ +add_decl_doc Point.nonsingular + +namespace Point + +variable {W} + +instance instZeroPoint [Nontrivial R] : Zero W.Point := + ⟨⟨W.nonsingular_lift_zero⟩⟩ + +lemma zero_def [Nontrivial R] : (⟨W.nonsingular_lift_zero⟩ : W.Point) = 0 := + rfl + +/-- The map from a nonsingular rational point on a Weierstrass curve `W` in affine coordinates +to the corresponding nonsingular rational point on `W` in Jacobian coordinates. -/ +def fromAffine [Nontrivial R] : W.toAffine.Point → W.Point + | 0 => 0 + | Affine.Point.some h => ⟨(W.nonsingular_lift_some ..).mpr h⟩ + +@[simp] +lemma fromAffine_zero [Nontrivial R] : fromAffine 0 = (0 : W.Point) := + rfl + +@[simp] +lemma fromAffine_some [Nontrivial R] {X Y : R} (h : W.toAffine.nonsingular X Y) : + fromAffine (Affine.Point.some h) = ⟨(W.nonsingular_lift_some ..).mpr h⟩ := + rfl + +variable {F : Type u} [Field F] {W : Jacobian F} + +/-- The negation of a nonsingular rational point on `W`. + +Given a nonsingular rational point `P` on `W`, use `-P` instead of `neg P`. -/ +@[pp_dot] +def neg (P : W.Point) : W.Point := + ⟨W.nonsingular_lift_neg_map P.nonsingular⟩ + +instance instNegPoint : Neg W.Point := + ⟨neg⟩ + +lemma neg_def (P : W.Point) : P.neg = -P := + rfl + +@[simp] +lemma neg_zero : (-⟨W.nonsingular_lift_zero⟩ : W.Point) = ⟨W.nonsingular_lift_zero⟩ := by + simp only [← neg_def, neg, neg_map_zero] + +/-- The addition of two nonsingular rational points on `W`. + +Given two nonsingular rational points `P` and `Q` on `W`, use `P + Q` instead of `add P Q`. -/ +@[pp_dot] +noncomputable def add (P Q : W.Point) : W.Point := + ⟨W.nonsingular_lift_add_map P.nonsingular Q.nonsingular⟩ + +noncomputable instance instAddPoint : Add W.Point := + ⟨add⟩ + +lemma add_def (P Q : W.Point) : P.add Q = P + Q := + rfl + +@[simp] +lemma zero_add (P : W.Point) : ⟨W.nonsingular_lift_zero⟩ + P = P := by + simp only [← add_def, add, add_map_zero_left] + +@[simp] +lemma add_zero (P : W.Point) : P + ⟨W.nonsingular_lift_zero⟩ = P := by + simp only [← add_def, add, add_map_zero_right' P.nonsingular] + +noncomputable instance instAddZeroClassPoint : AddZeroClass W.Point := + ⟨zero_add, add_zero⟩ + +open scoped Classical + +/-- The map from a point representative that is nonsingular on a Weierstrass curve `W` in Jacobian +coordinates to the corresponding nonsingular rational point on `W` in affine coordinates. -/ +noncomputable def toAffine {P : Fin 3 → F} (h : W.nonsingular P) : W.toAffine.Point := + if hPz : P z = 0 then 0 else Affine.Point.some <| nonsingular_affine_of_Zne0 h hPz + +lemma toAffine_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z = 0) : toAffine h = 0 := + dif_pos hPz + +lemma toAffine_zero : toAffine W.nonsingular_zero = 0 := + toAffine_of_Zeq0 rfl + +lemma toAffine_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z ≠ 0) : + toAffine h = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := + dif_neg hPz + +lemma toAffine_some {X Y : F} (h : W.nonsingular ![X, Y, 1]) : + toAffine h = Affine.Point.some ((W.nonsingular_some X Y).mp h) := by + rw [toAffine_of_Zne0 <| by exact one_ne_zero] + matrix_simp + simp only [one_pow, div_one] + +lemma toAffine_neg {P : Fin 3 → F} (hP : W.nonsingular P) : + toAffine (nonsingular_neg hP) = -toAffine hP := by + by_cases hPz : P z = 0 + · rw [toAffine_of_Zeq0 <| by exact hPz, toAffine_of_Zeq0 hPz, Affine.Point.neg_zero] + · rw [toAffine_of_Zne0 <| by exact hPz, toAffine_of_Zne0 hPz, Affine.Point.neg_some, + Affine.Point.some.injEq] + exact ⟨rfl, negY_divZ hPz⟩ + +lemma toAffine_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : + toAffine (nonsingular_add hP hQ) = toAffine hP + toAffine hQ := by + by_cases hPz : P z = 0 + · simp only [W.add_of_Zeq0_left (Q := Q) hPz, toAffine_of_Zeq0 hPz, _root_.zero_add] + · by_cases hQz : Q z = 0 + · simp only [W.add_of_Zeq0_right hPz hQz, toAffine_of_Zeq0 hQz, _root_.add_zero] + · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · have hx' : P x / P z ^ 2 = Q x / Q z ^ 2 := + (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mpr <| mul_comm (P z ^ 2) _ ▸ hx + rw [toAffine_of_Zne0 hPz, toAffine_of_Zne0 hQz] + by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · have hy' : P y / P z ^ 3 = W.negY Q / Q z ^ 3 := Iff.mpr + (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)) <| mul_comm (P z ^ 3) _ ▸ hy + simp only [W.add_of_Yeq hPz hQz hx hy] + rw [toAffine_zero, Affine.Point.some_add_some_of_Yeq hx' <| by rwa [← negY_divZ hQz]] + · have hy' : P y / P z ^ 3 ≠ W.negY Q / Q z ^ 3 := Function.comp + (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp + simp only [W.add_of_Yne hPz hQz hx hy] + rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, + Affine.Point.some_add_some_of_Yne hx' <| (negY_divZ hQz).symm ▸ hy', + Affine.Point.some.injEq] + exact ⟨addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy⟩ + · have hx' : P x / P z ^ 2 ≠ Q x / Q z ^ 2 := + (mul_comm (P z ^ 2) _ ▸ hx) ∘ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp + simp only [W.add_of_Xne hPz hQz hx] + rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, toAffine_of_Zne0 hPz, + toAffine_of_Zne0 hQz, Affine.Point.some_add_some_of_Xne hx', Affine.Point.some.injEq] + exact ⟨addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx⟩ + +lemma toAffine_of_equiv (P Q : Fin 3 → F) (h : P ≈ Q) : + HEq (toAffine (W := W) (P := P)) (toAffine (W := W) (P := Q)) := by + rcases h with ⟨u, rfl⟩ + refine Function.hfunext (propext <| W.nonsingular_smul_iff Q u) <| fun _ _ _ => ?_ + by_cases hPz : Q z = 0 + · rw [toAffine_of_Zeq0 <| by exact u.mul_right_eq_zero.mpr hPz, toAffine_of_Zeq0 hPz] + · rw [toAffine_of_Zne0 <| by exact mul_ne_zero u.ne_zero hPz, toAffine_of_Zne0 hPz, heq_eq_eq] + simp only [Affine.Point.some.injEq, smul_fin3_ext, mul_pow, + mul_div_mul_left _ _ <| pow_ne_zero _ u.ne_zero, and_self] + +/-- The map from a nonsingular rational point on a Weierstrass curve `W` in Jacobian coordinates +to the corresponding nonsingular rational point on `W` in affine coordinates. -/ +noncomputable def toAffine_lift (P : W.Point) : W.toAffine.Point := + P.point.hrecOn (fun _ => toAffine) toAffine_of_equiv P.nonsingular + +lemma toAffine_lift_eq {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : + toAffine_lift ⟨h⟩ = toAffine h := + rfl + +lemma toAffine_lift_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z = 0) : + toAffine_lift ⟨h⟩ = 0 := + toAffine_of_Zeq0 hPz (h := h) + +lemma toAffine_lift_zero : toAffine_lift (0 : W.Point) = 0 := + toAffine_zero + +lemma toAffine_lift_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z ≠ 0) : + toAffine_lift ⟨h⟩ = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := + toAffine_of_Zne0 hPz (h := h) + +lemma toAffine_lift_some {X Y : F} (h : W.nonsingular_lift ⟦![X, Y, 1]⟧) : + toAffine_lift ⟨h⟩ = Affine.Point.some ((W.nonsingular_some X Y).mp h) := + toAffine_some h + +lemma toAffine_lift_neg {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : + toAffine_lift (-⟨h⟩) = -toAffine_lift ⟨h⟩ := + toAffine_neg h + +lemma toAffine_lift_add {P Q : Fin 3 → F} (hP : W.nonsingular_lift ⟦P⟧) + (hQ : W.nonsingular_lift ⟦Q⟧) : + toAffine_lift (⟨hP⟩ + ⟨hQ⟩) = toAffine_lift ⟨hP⟩ + toAffine_lift ⟨hQ⟩ := + toAffine_add hP hQ + +/-- The equivalence between the nonsingular rational points on a Weierstrass curve `W` in Jacobian +coordinates with the nonsingular rational points on `W` in affine coordinates. -/ +@[simps] +noncomputable def toAffine_addEquiv : W.Point ≃+ W.toAffine.Point where + toFun := toAffine_lift + invFun := fromAffine + left_inv := by + rintro @⟨⟨P⟩, h⟩ + by_cases hPz : P z = 0 + · erw [toAffine_lift_eq, toAffine_of_Zeq0 hPz, fromAffine_zero, mk.injEq, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Zeq0 h hPz + · erw [toAffine_lift_eq, toAffine_of_Zne0 hPz, fromAffine_some, mk.injEq, Quotient.eq] + exact Setoid.symm <| equiv_some_of_Zne0 hPz + right_inv := by + rintro (_ | _) + · erw [fromAffine_zero, toAffine_lift_zero, Affine.Point.zero_def] + · rw [fromAffine_some, toAffine_lift_some] + map_add' := by + rintro @⟨⟨_⟩, _⟩ @⟨⟨_⟩, _⟩ + simpa only using toAffine_lift_add .. + +end Point + +end Point + +end WeierstrassCurve.Jacobian From 906e8581e77684f8c4425ddf76f9cd0070395cbd Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 4 Jan 2024 09:48:58 +0000 Subject: [PATCH 10/47] Implement equations and nonsingularity for Jacobian coordinates --- Mathlib.lean | 1 + .../EllipticCurve/Jacobian.lean | 314 ++++++++++++++++++ 2 files changed, 315 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean diff --git a/Mathlib.lean b/Mathlib.lean index 72f713857128d4..65cf3ba1d377be 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -492,6 +492,7 @@ import Mathlib.Algebra.Tropical.Lattice import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.Group +import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass import Mathlib.AlgebraicGeometry.FunctionField import Mathlib.AlgebraicGeometry.GammaSpecAdjunction diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean new file mode 100644 index 00000000000000..1d779115ee030b --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -0,0 +1,314 @@ +/- +Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ +import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.Data.MvPolynomial.CommRing + +/-! +# Jacobian coordinates for Weierstrass curves + +This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence +class of triples up to scaling by weights, satisfying a Weierstrass equation with a nonsingular +condition. + +## Mathematical background + +Let `W` be a Weierstrass curve over a field `F`. A point on the weighted projective plane with +weights $(2, 3, 1)$ is an equivalence class of triples $[x:y:z]$ with coordinates in `F` such that +$(x, y, z) \sim (x', y', z')$ precisely if there is some unit $u$ of `F` such that +$(x, y, z) = (u^2x', u^3y', uz')$, with an extra condition that $(x, y, z) \ne (0, 0, 0)$. +A rational point is a point on the $(2, 3, 1)$-projective plane satisfying a $(2, 3, 1)$-homogeneous +Weierstrass equation $Y^2 + a_1XYZ + a_3YZ^3 = X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6$, and being +nonsingular means the partial derivatives $W_X(X, Y, Z)$, $W_Y(X, Y, Z)$, and $W_Z(X, Y, Z)$ do not +vanish simultaneously. Note that the vanishing of the Weierstrass equation and its partial +derivatives are independent of the representative for $[x:y:z]$, and the nonsingularity condition +already implies that $(x, y, z) \ne (0, 0, 0)$, so a nonsingular rational point on `W` can simply be +given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any representative. +In cryptography, as well as in this file, this is often called the Jacobian coordinates of `W`. + +## Main definitions + + * `WeierstrassCurve.Jacobian.PointClass`: the equivalence class of a point representative. + * `WeierstrassCurve.Jacobian.toAffine`: the Weierstrass curve in affine coordinates. + * `WeierstrassCurve.Jacobian.nonsingular`: the nonsingular condition on a point representative. + * `WeierstrassCurve.Jacobian.nonsingular_lift`: the nonsingular condition on a point class. + +## Implementation notes + +A point representative is implemented as a term `P` of type `Fin 3 → R`, which allows for the vector +notation `![x, y, z]`. However, `P` is not definitionally equivalent to the expanded vector +`![P x, P y, P z]`, so the auxiliary lemma `fin3_def` can be used to convert between the two forms. +The equivalence of two point representatives `P` and `Q` is implemented as an equivalence of orbits +of the action of `Rˣ`, or equivalently that there is some unit `u` of `R` such that `P = u • Q`. +However, `u • Q` is again not definitionally equal to `![u² * Q x, u³ * Q y, u * Q z]`, so the +auxiliary lemmas `smul_fin3` and `smul_fin3_ext` can be used to convert between the two forms. + +## References + +[J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, rational point, Jacobian coordinates +-/ + +local notation "x" => 0 + +local notation "y" => 1 + +local notation "z" => 2 + +local macro "matrix_simp" : tactic => + `(tactic| simp only [Matrix.head_cons, Matrix.tail_cons, Matrix.smul_empty, Matrix.smul_cons, + Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.cons_val_two]) + +universe u + +lemma fin3_def {R : Type u} (P : Fin 3 → R) : P = ![P x, P y, P z] := by + ext n; fin_cases n <;> rfl + +private instance {R : Type u} [CommRing R] : SMul Rˣ <| Fin 3 → R := + ⟨fun u P => ![u ^ 2 * P x, u ^ 3 * P y, u * P z]⟩ + +lemma smul_fin3 {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : + u • P = ![u ^ 2 * P x, u ^ 3 * P y, u * P z] := + rfl + +lemma smul_fin3_ext {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : + (u • P) 0 = u ^ 2 * P x ∧ (u • P) 1 = u ^ 3 * P y ∧ (u • P) 2 = u * P z := + ⟨rfl, rfl, rfl⟩ + +private instance {R : Type u} [CommRing R] : MulAction Rˣ <| Fin 3 → R where + one_smul := fun _ => by + simp only [smul_fin3, Units.val_one, one_pow, one_mul, ← fin3_def] + mul_smul := fun u v P => by + simp only [smul_fin3, Units.val_mul, mul_pow, mul_assoc] + matrix_simp + +/-! ## Weierstrass curves -/ + +/-- An abbreviation for a Weierstrass curve in Jacobian coordinates. -/ +abbrev WeierstrassCurve.Jacobian := + WeierstrassCurve + +namespace WeierstrassCurve.Jacobian + +open MvPolynomial + +local macro "eval_simp" : tactic => + `(tactic| simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, eval_pow]) + +variable (R : Type u) [CommRing R] + +/-- The equivalence setoid for a point representative. -/ +def PointSetoid : Setoid <| Fin 3 → R := + MulAction.orbitRel Rˣ <| Fin 3 → R + +attribute [local instance] PointSetoid + +/-- The equivalence class of a point representative. -/ +abbrev PointClass : Type u := + MulAction.orbitRel.Quotient Rˣ <| Fin 3 → R + +variable {R} (W : Jacobian R) + +/-- The coercion to a Weierstrass curve in affine coordinates. -/ +@[pp_dot] +abbrev toAffine : Affine R := + W + +section Equation + +/-! ### Equations and nonsingularity -/ + +/-- The polynomial $W(X, Y, Z) := Y^2 + a_1XYZ + a_3YZ^3 - (X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6)$ +associated to a Weierstrass curve `W` over `R`. This is represented as a term of type +`MvPolynomial (Fin 3) R`, where `X 0`, `X 1`, and `X 2` represent $X$, $Y$, and $Z$ respectively. -/ +@[pp_dot] +noncomputable def polynomial : MvPolynomial (Fin 3) R := + X 1 ^ 2 + C W.a₁ * X 0 * X 1 * X 2 + C W.a₃ * X 1 * X 2 ^ 3 + - (X 0 ^ 3 + C W.a₂ * X 0 ^ 2 * X 2 ^ 2 + C W.a₄ * X 0 * X 2 ^ 4 + C W.a₆ * X 2 ^ 6) + +lemma eval_polynomial (P : Fin 3 → R) : eval P W.polynomial = + P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 + - (P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6) := by + rw [polynomial] + eval_simp + +/-- The proposition that a point representative $(x, y, z)$ lies in `W`. +In other words, $W(x, y, z) = 0$. -/ +@[pp_dot] +def equation (P : Fin 3 → R) : Prop := + eval P W.polynomial = 0 + +lemma equation_iff (P : Fin 3 → R) : W.equation P ↔ + P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 + = P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6 := by + rw [equation, eval_polynomial, sub_eq_zero] + +lemma equation_zero : W.equation ![1, 1, 0] := + (W.equation_iff ![1, 1, 0]).mpr <| by matrix_simp; ring1 + +lemma equation_zero' (Y : R) : W.equation ![Y ^ 2, Y ^ 3, 0] := + (W.equation_iff ![Y ^ 2, Y ^ 3, 0]).mpr <| by matrix_simp; ring1 + +lemma equation_some (X Y : R) : W.equation ![X, Y, 1] ↔ W.toAffine.equation X Y := by + rw [equation_iff, W.toAffine.equation_iff] + congr! 1 <;> matrix_simp <;> ring1 + +lemma equation_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.equation (u • P) ↔ W.equation P := + have (u : Rˣ) {P : Fin 3 → R} (h : W.equation P) : W.equation <| u • P := by + rw [equation_iff] at h ⊢ + linear_combination (norm := (simp only [smul_fin3_ext]; ring1)) (u : R) ^ 6 * h + ⟨fun h => by convert this u⁻¹ h; rw [inv_smul_smul], this u⟩ + +/-- The partial derivative $W_X(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $X$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialX : MvPolynomial (Fin 3) R := + C W.a₁ * X 1 * X 2 - (C 3 * X 0 ^ 2 + C (2 * W.a₂) * X 0 * X 2 ^ 2 + C W.a₄ * X 2 ^ 4) + +lemma eval_polynomialX (P : Fin 3 → R) : eval P W.polynomialX = + W.a₁ * P y * P z - (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4) := by + rw [polynomialX] + eval_simp + +/-- The partial derivative $W_Y(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Y$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialY : MvPolynomial (Fin 3) R := + C 2 * X 1 + C W.a₁ * X 0 * X 2 + C W.a₃ * X 2 ^ 3 + +lemma eval_polynomialY (P : Fin 3 → R) : + eval P W.polynomialY = 2 * P y + W.a₁ * P x * P z + W.a₃ * P z ^ 3 := by + rw [polynomialY] + eval_simp + +/-- The partial derivative $W_Z(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Z$. + +TODO: define this in terms of `MvPolynomial.pderiv`. -/ +@[pp_dot] +noncomputable def polynomialZ : MvPolynomial (Fin 3) R := + C W.a₁ * X 0 * X 1 + C (3 * W.a₃) * X 1 * X 2 ^ 2 + - (C (2 * W.a₂) * X 0 ^ 2 * X 2 + C (4 * W.a₄) * X 0 * X 2 ^ 3 + C (6 * W.a₆) * X 2 ^ 5) + +lemma eval_polynomialZ (P : Fin 3 → R) : eval P W.polynomialZ = + W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 + - (2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by + rw [polynomialZ] + eval_simp + +/-- The proposition that a point representative $(x, y, z)$ in `W` is nonsingular. +In other words, either $W_X(x, y, z) \ne 0$, $W_Y(x, y, z) \ne 0$, or $W_Z(x, y, z) \ne 0$. -/ +@[pp_dot] +def nonsingular (P : Fin 3 → R) : Prop := + W.equation P ∧ (eval P W.polynomialX ≠ 0 ∨ eval P W.polynomialY ≠ 0 ∨ eval P W.polynomialZ ≠ 0) + +lemma nonsingular_iff (P : Fin 3 → R) : W.nonsingular P ↔ W.equation P ∧ + (W.a₁ * P y * P z ≠ 3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 ∨ + P y ≠ -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 ∨ + W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 + ≠ 2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by + rw [nonsingular, eval_polynomialX, eval_polynomialY, eval_polynomialZ, sub_ne_zero, sub_ne_zero, + ← sub_ne_zero (a := P y)] + congr! 4 + ring1 + +lemma nonsingular_zero [Nontrivial R] : W.nonsingular ![1, 1, 0] := + (W.nonsingular_iff ![1, 1, 0]).mpr ⟨W.equation_zero, + by simp; by_contra! h; exact one_ne_zero <| by linear_combination -h.1 - h.2.1⟩ + +lemma nonsingular_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : + W.nonsingular ![Y ^ 2, Y ^ 3, 0] := + (W.nonsingular_iff ![Y ^ 2, Y ^ 3, 0]).mpr ⟨W.equation_zero' Y, + by simp [hy]; by_contra! h; exact pow_ne_zero 3 hy <| by linear_combination Y ^ 3 * h.1 - h.2.1⟩ + +lemma nonsingular_some (X Y : R) : W.nonsingular ![X, Y, 1] ↔ W.toAffine.nonsingular X Y := by + rw [nonsingular_iff] + matrix_simp + simp only [W.toAffine.nonsingular_iff, equation_some, and_congr_right_iff, + W.toAffine.equation_iff, ← not_and_or, not_iff_not, one_pow, mul_one, Iff.comm, iff_self_and] + intro h hX hY + linear_combination (norm := ring1) 6 * h - 2 * X * hX - 3 * Y * hY + +lemma nonsingular_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.nonsingular (u • P) ↔ W.nonsingular P := + have (u : Rˣ) {P : Fin 3 → R} (h : W.nonsingular <| u • P) : W.nonsingular P := by + rcases (W.nonsingular_iff _).mp h with ⟨h, h'⟩ + refine (W.nonsingular_iff P).mpr ⟨(W.equation_smul_iff P u).mp h, ?_⟩ + contrapose! h' + simp only [smul_fin3_ext] + exact ⟨by linear_combination (norm := ring1) (u : R) ^ 4 * h'.left, + by linear_combination (norm := ring1) (u : R) ^ 3 * h'.right.left, + by linear_combination (norm := ring1) (u : R) ^ 5 * h'.right.right⟩ + ⟨this u, fun h => this u⁻¹ <| by rwa [inv_smul_smul]⟩ + +lemma nonsingular_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.nonsingular P ↔ W.nonsingular Q := by + rcases h with ⟨u, rfl⟩ + exact W.nonsingular_smul_iff Q u + +/-- The proposition that a point class on `W` is nonsingular. If `P` is a point representative, +then `W.nonsingular_lift ⟦P⟧` is definitionally equivalent to `W.nonsingular P`. -/ +@[pp_dot] +def nonsingular_lift (P : PointClass R) : Prop := + P.lift W.nonsingular fun _ _ => propext ∘ W.nonsingular_of_equiv + +@[simp] +lemma nonsingular_lift_eq (P : Fin 3 → R) : W.nonsingular_lift ⟦P⟧ = W.nonsingular P := + rfl + +lemma nonsingular_lift_zero [Nontrivial R] : W.nonsingular_lift ⟦![1, 1, 0]⟧ := + W.nonsingular_zero + +lemma nonsingular_lift_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : + W.nonsingular_lift ⟦![Y ^ 2, Y ^ 3, 0]⟧ := + W.nonsingular_zero' hy + +lemma nonsingular_lift_some (X Y : R) : + W.nonsingular_lift ⟦![X, Y, 1]⟧ ↔ W.toAffine.nonsingular X Y := + W.nonsingular_some X Y + +variable {F : Type u} [Field F] {W : Jacobian F} + +lemma equiv_of_Zeq0 {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z = 0) + (hQz : Q z = 0) : P ≈ Q := by + rw [fin3_def P, hPz] at hP ⊢ + rw [fin3_def Q, hQz] at hQ ⊢ + simp [nonsingular_iff, equation_iff] at hP hQ + have hPx : P x ≠ 0 := fun h => by simp [h] at hP; simp [hP] at hP + have hQx : Q x ≠ 0 := fun h => by simp [h] at hQ; simp [hQ] at hQ + have hPy : P y ≠ 0 := fun h => by simp [h] at hP; exact hPx <| pow_eq_zero hP.left.symm + have hQy : Q y ≠ 0 := fun h => by simp [h] at hQ; exact hQx <| pow_eq_zero hQ.left.symm + use Units.mk0 _ <| mul_ne_zero (div_ne_zero hPy hPx) (div_ne_zero hQx hQy) + simp [smul_fin3, mul_pow, div_pow] + congr! 2 + · field_simp [hP.left, hQ.left] + ring1 + · field_simp [← hP.left, ← hQ.left] + ring1 + +lemma equiv_zero_of_Zeq0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z = 0) : P ≈ ![1, 1, 0] := + equiv_of_Zeq0 h W.nonsingular_zero hPz rfl + +lemma equiv_some_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : P ≈ ![P x / P z ^ 2, P y / P z ^ 3, 1] := + ⟨Units.mk0 _ hPz, by simp [smul_fin3, ← fin3_def P, mul_div_cancel' _ <| pow_ne_zero _ hPz]⟩ + +lemma nonsingular_iff_affine_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.nonsingular P ↔ W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := + (W.nonsingular_of_equiv <| equiv_some_of_Zne0 hPz).trans <| W.nonsingular_some .. + +lemma nonsingular_of_affine_of_Zne0 {P : Fin 3 → F} + (h : W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3)) (hPz : P z ≠ 0) : + W.nonsingular P := + (nonsingular_iff_affine_of_Zne0 hPz).mpr h + +lemma nonsingular_affine_of_Zne0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z ≠ 0) : + W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := + (nonsingular_iff_affine_of_Zne0 hPz).mp h + +end Equation + +end WeierstrassCurve.Jacobian From b9e60a1f56cebe170a300216e09bde1043cbd4a9 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 4 Jan 2024 09:56:27 +0000 Subject: [PATCH 11/47] Implement group operation polynomials for Jacobian coordinates --- .../EllipticCurve/Jacobian.lean | 210 ++++++++++++++++++ 1 file changed, 210 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 1d779115ee030b..4afd8aa624c3de 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -28,6 +28,10 @@ already implies that $(x, y, z) \ne (0, 0, 0)$, so a nonsingular rational point given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any representative. In cryptography, as well as in this file, this is often called the Jacobian coordinates of `W`. +As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms +an abelian group under the same secant-and-tangent process, but the polynomials involved are +$(2, 3, 1)$-homogeneous, and any instances of division become multiplication in the $Z$-coordinate. + ## Main definitions * `WeierstrassCurve.Jacobian.PointClass`: the equivalence class of a point representative. @@ -311,4 +315,210 @@ lemma nonsingular_affine_of_Zne0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : end Equation +section Polynomial + +/-! ### Group operation polynomials -/ + +/-- The $Y$-coordinate of the negation of a point representative. -/ +@[pp_dot] +def negY (P : Fin 3 → R) : R := + -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 + +lemma negY_smul (P : Fin 3 → R) (u : Rˣ) : W.negY (u • P) = u ^ 3 * W.negY P := by + simp only [negY, smul_fin3_ext] + ring1 + +/-- The $X$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addX_of_Xne (P Q : Fin 3 → R) : R := + (P y * Q z ^ 3 - P z ^ 3 * Q y) ^ 2 + + W.a₁ * P z * Q z * (P y * Q z ^ 3 - P z ^ 3 * Q y) * (P x * Q z ^ 2 - P z ^ 2 * Q x) + - (W.a₂ * P z ^ 2 * Q z ^ 2 + P x * Q z ^ 2 + P z ^ 2 * Q x) + * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2 + +lemma addX_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addX_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addX_of_Xne P Q := by + simp only [addX_of_Xne, smul_fin3_ext] + ring1 + +/-- The $X$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addX_of_Yne (P : Fin 3 → R) : R := + (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) ^ 2 + + W.a₁ * (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) + * (P y * P z - W.negY P * P z) + - (W.a₂ * P z ^ 2 + 2 * P x) * (P y - W.negY P) ^ 2 + +lemma addX_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addX_of_Yne (u • P) = (u : R) ^ 8 * W.addX_of_Yne P := by + simp only [addX_of_Yne, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the addition of two point representatives, before applying the final +negation that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where their $Z$-coordinates are non-zero and their +$X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addY'_of_Xne (P Q : Fin 3 → R) : R := + (P y * Q z ^ 3 - P z ^ 3 * Q y) + * (W.addX_of_Xne P Q - P x * Q z ^ 2 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) + + P y * Q z ^ 3 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 3 + +lemma addY'_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addY'_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY'_of_Xne P Q := by + simp only [addY'_of_Xne, addX_of_Xne_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the doubling of a point representative, before applying the final negation +that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where its $Z$-coordinate is non-zero and its $Y$-coordinate +is distinct from that of its negation. -/ +@[pp_dot] +def addY'_of_Yne (P : Fin 3 → R) : R := + (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) + * (W.addX_of_Yne P - P x * (P y - W.negY P) ^ 2) + + P y * (P y - W.negY P) ^ 3 + +lemma addY'_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addY'_of_Yne (u • P) = (u : R) ^ 12 * W.addY'_of_Yne P := by + simp only [addY'_of_Yne, addX_of_Yne_smul, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Z$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +def addZ_of_Xne (P Q : Fin 3 → R) : R := + P z * Q z * (P x * Q z ^ 2 - P z ^ 2 * Q x) + +lemma addZ_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + addZ_of_Xne (u • P) (v • Q) = (u : R) ^ 3 * (v : R) ^ 3 * addZ_of_Xne P Q := by + simp only [addZ_of_Xne, smul_fin3_ext] + ring1 + +/-- The $Z$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addZ_of_Yne (P : Fin 3 → R) : R := + P z * (P y - W.negY P) + +lemma addZ_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addZ_of_Yne (u • P) = (u : R) ^ 4 * W.addZ_of_Yne P := by + simp only [addZ_of_Yne, negY_smul, smul_fin3_ext] + ring1 + +/-- The $Y$-coordinate of the addition of two point representatives, where their $Z$-coordinates are +non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ +@[pp_dot] +def addY_of_Xne (P Q : Fin 3 → R) : R := + W.negY ![W.addX_of_Xne P Q, W.addY'_of_Xne P Q, addZ_of_Xne P Q] + +lemma addY_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : + W.addY_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY_of_Xne P Q := by + simp only [addY_of_Xne, negY, addX_of_Xne_smul, addY'_of_Xne_smul, addZ_of_Xne_smul] + matrix_simp + ring1 + +/-- The $Y$-coordinate of the doubling of a point representative, where its $Z$-coordinate is +non-zero and its $Y$-coordinate is distinct from that of its negation. -/ +@[pp_dot] +def addY_of_Yne (P : Fin 3 → R) : R := + W.negY ![W.addX_of_Yne P, W.addY'_of_Yne P, W.addZ_of_Yne P] + +lemma addY_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : + W.addY_of_Yne (u • P) = (u : R) ^ 12 * W.addY_of_Yne P := by + simp only [addY_of_Yne, negY, addX_of_Yne_smul, addY'_of_Yne_smul, addZ_of_Yne_smul] + matrix_simp + ring1 + +variable {F : Type u} [Field F] {W : Jacobian F} + +lemma negY_divZ {P : Fin 3 → F} (hPz : P z ≠ 0) : + W.negY P / P z ^ 3 = W.toAffine.negY (P x / P z ^ 2) (P y / P z ^ 3) := by + field_simp [negY, Affine.negY] + ring1 + +lemma Yne_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z ≠ 0) + (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + P y ≠ W.negY P := by + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + have hx' : P x * (P z / P z ^ 3) = Q x * (Q z / Q z ^ 3) := by + simp_rw [pow_succ _ 2, div_mul_right _ hPz, div_mul_right _ hQz, mul_one_div, hx] + have hy' : P y / P z ^ 3 = Q y / Q z ^ 3 := + Affine.Yeq_of_Yne (nonsingular_affine_of_Zne0 hP hPz).left + (nonsingular_affine_of_Zne0 hQ hQz).left hx <| (negY_divZ hQz).symm ▸ hy + simp_rw [negY, sub_div, neg_div, mul_div_assoc, mul_assoc, ← hy', ← hx', + div_self <| pow_ne_zero 3 hQz, ← div_self <| pow_ne_zero 3 hPz, ← mul_assoc, ← mul_div_assoc, + ← neg_div, ← sub_div, div_left_inj' <| pow_ne_zero 3 hPz] at hy + exact hy + +lemma addX_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addX_of_Xne P Q / addZ_of_Xne P Q ^ 2 = + W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.slope_of_Xne <| + by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] + field_simp [sub_ne_zero_of_ne hx, addX_of_Xne, addZ_of_Xne] + ring1 + +lemma addX_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.addX_of_Yne P / W.addZ_of_Yne P ^ 2 = W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← hx, ← negY_divZ hPz] + field_simp [addX_of_Yne, addX_of_Yne, addZ_of_Yne] + ring1 + +lemma addY'_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY'_of_Xne P Q / addZ_of_Xne P Q ^ 3 = + W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY', ← addX_div_addZ_of_Xne hPz hQz hx, Affine.slope_of_Xne <| + by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] + field_simp [sub_ne_zero_of_ne hx, addY'_of_Xne, addX_of_Xne, addZ_of_Xne] + ring1 + +lemma addY'_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY'_of_Yne P / W.addZ_of_Yne P ^ 3 = + W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + rw [Affine.addY', ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy] + simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] + at hx hy + rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← negY_divZ hPz] + field_simp [addY'_of_Yne, addX_of_Yne, addZ_of_Yne] + ring1 + +lemma addZ_ne_zero_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : addZ_of_Xne P Q ≠ 0 := + mul_ne_zero (mul_ne_zero hPz hQz) <| sub_ne_zero_of_ne hx + +lemma addZ_ne_zero_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addZ_of_Yne P ≠ 0 := + mul_ne_zero hPz <| sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy + +lemma addY_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY_of_Xne P Q / addZ_of_Xne P Q ^ 3 = + W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY, ← addX_div_addZ_of_Xne hPz hQz hx, ← addY'_div_addZ_of_Xne hPz hQz hx] + exact negY_divZ <| addZ_ne_zero_of_Xne hPz hQz hx + +lemma addY_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) + (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY_of_Yne P / W.addZ_of_Yne P ^ 3 = + W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + rw [Affine.addY, ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, + ← addY'_div_addZ_of_Yne hP hQ hPz hQz hx hy] + exact negY_divZ <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy + +end Polynomial + end WeierstrassCurve.Jacobian From d42c841fc75c3f9d614743c815ceb897edc8ade0 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 4 Jan 2024 10:00:52 +0000 Subject: [PATCH 12/47] Implement group operations on point representatives for Jacobian coordinates --- .../EllipticCurve/Jacobian.lean | 239 +++++++++++++++++- 1 file changed, 238 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 4afd8aa624c3de..740b395de09849 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -11,7 +11,8 @@ import Mathlib.Data.MvPolynomial.CommRing This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence class of triples up to scaling by weights, satisfying a Weierstrass equation with a nonsingular -condition. +condition. This file also defines the negation and addition operations of the group law for this +type, and proves that they respect the Weierstrass equation and the nonsingular condition. ## Mathematical background @@ -31,6 +32,7 @@ In cryptography, as well as in this file, this is often called the Jacobian coor As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms an abelian group under the same secant-and-tangent process, but the polynomials involved are $(2, 3, 1)$-homogeneous, and any instances of division become multiplication in the $Z$-coordinate. +Note that most computational proofs follow from their analogous proofs for affine coordinates. ## Main definitions @@ -38,6 +40,15 @@ $(2, 3, 1)$-homogeneous, and any instances of division become multiplication in * `WeierstrassCurve.Jacobian.toAffine`: the Weierstrass curve in affine coordinates. * `WeierstrassCurve.Jacobian.nonsingular`: the nonsingular condition on a point representative. * `WeierstrassCurve.Jacobian.nonsingular_lift`: the nonsingular condition on a point class. + * `WeierstrassCurve.Jacobian.neg`: the negation operation on a point representative. + * `WeierstrassCurve.Jacobian.neg_map`: the negation operation on a point class. + * `WeierstrassCurve.Jacobian.add`: the addition operation on a point representative. + * `WeierstrassCurve.Jacobian.add_map`: the addition operation on a point class. + +## Main statements + + * `WeierstrassCurve.Jacobian.nonsingular_neg`: negation preserves the nonsingular condition. + * `WeierstrassCurve.Jacobian.nonsingular_add`: addition preserves the nonsingular condition. ## Implementation notes @@ -521,4 +532,230 @@ lemma addY_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.no end Polynomial +section Representative + +/-! ### Group operations on point representatives -/ + +/-- The negation of a point representative. -/ +@[pp_dot] +def neg (P : Fin 3 → R) : Fin 3 → R := + ![P x, W.negY P, P z] + +@[simp] +lemma neg_zero : W.neg ![1, 1, 0] = ![1, -1, 0] := by + erw [neg, negY, mul_zero, zero_pow three_pos, mul_zero, sub_zero, sub_zero] + rfl + +@[simp] +lemma neg_some (X Y : R) : W.neg ![X, Y, 1] = ![X, -Y - W.a₁ * X - W.a₃, 1] := by + erw [neg, negY, mul_one, one_pow, mul_one] + rfl + +lemma neg_smul_equiv (P : Fin 3 → R) (u : Rˣ) : W.neg (u • P) ≈ W.neg P := + ⟨u, by simp_rw [neg, negY_smul, smul_fin3]; rfl⟩ + +lemma neg_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.neg P ≈ W.neg Q := by + rcases h with ⟨u, rfl⟩ + exact W.neg_smul_equiv Q u + +/-- The negation of a point class. If `P` is a point representative, +then `W.neg_map ⟦P⟧` is definitionally equivalent to `W.neg P`. -/ +@[pp_dot] +def neg_map (P : PointClass R) : PointClass R := + P.map W.neg fun _ _ => W.neg_equiv + +lemma neg_map_eq {P : Fin 3 → R} : W.neg_map ⟦P⟧ = ⟦W.neg P⟧ := + rfl + +@[simp] +lemma neg_map_zero : W.neg_map ⟦![1, 1, 0]⟧ = ⟦![1, 1, 0]⟧ := by + simpa only [neg_map_eq, neg_zero, Quotient.eq] using ⟨-1, by norm_num [smul_fin3]⟩ + +@[simp] +lemma neg_map_some (X Y : R) : W.neg_map ⟦![X, Y, 1]⟧ = ⟦![X, -Y - W.a₁ * X - W.a₃, 1]⟧ := by + rw [neg_map_eq, neg_some] + +open scoped Classical + +/-- The addition of two point representatives. -/ +@[pp_dot] +noncomputable def add (P Q : Fin 3 → R) : Fin 3 → R := + if P z = 0 then Q else if Q z = 0 then P else if P x * Q z ^ 2 = P z ^ 2 * Q x then + if P y * Q z ^ 3 = P z ^ 3 * W.negY Q then ![1, 1, 0] else + ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] + else ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] + +@[simp] +lemma add_of_Zeq0_left {P Q : Fin 3 → R} (hPz : P z = 0) : W.add P Q = Q := + if_pos hPz + +lemma add_zero_left (P : Fin 3 → R) : W.add ![1, 1, 0] P = P := + W.add_of_Zeq0_left rfl + +@[simp] +lemma add_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : W.add P Q = P := by + rw [add, if_neg hPz, if_pos hQz] + +lemma add_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add P ![1, 1, 0] = P := + W.add_of_Zeq0_right hPz rfl + +@[simp] +lemma add_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : + W.add P Q = ![1, 1, 0] := by + rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_pos hy] + +@[simp] +lemma add_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.add P Q = ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] := by + rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_neg hy] + +@[simp] +lemma add_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.add P Q = ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] := by + rw [add, if_neg hPz, if_neg hQz, if_neg hx] + +variable [IsDomain R] + +lemma add_smul_equiv (P Q : Fin 3 → R) (u v : Rˣ) : W.add (u • P) (v • Q) ≈ W.add P Q := by + have huv (n : ℕ) : (u ^ n * v ^ n : R) ≠ 0 := + mul_ne_zero (pow_ne_zero n u.ne_zero) (pow_ne_zero n v.ne_zero) + by_cases hPz : P z = 0 + · exact ⟨v, by rw [W.add_of_Zeq0_left hPz, + W.add_of_Zeq0_left <| by simp only [smul_fin3_ext, hPz, mul_zero]]⟩ + · have huz : u * P z ≠ 0 := mul_ne_zero u.ne_zero hPz + by_cases hQz : Q z = 0 + · rw [W.add_of_Zeq0_right hPz hQz, + W.add_of_Zeq0_right huz <| by simp only [smul_fin3_ext, hQz, mul_zero]] + exact ⟨u, rfl⟩ + · have hvz : v * Q z ≠ 0 := mul_ne_zero v.ne_zero hQz + by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · rw [W.add_of_Yeq huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by + simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm, hy], + W.add_of_Yeq hPz hQz hx hy] + · rw [W.add_of_Yne huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by + simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm] + exact hy ∘ mul_left_cancel₀ (huv 3), + addX_of_Yne_smul, addY_of_Yne_smul, addZ_of_Yne_smul, W.add_of_Yne hPz hQz hx hy] + exact ⟨u ^ 4, by simp only [smul_fin3, ← Units.val_pow_eq_pow_val, ← pow_mul]; rfl⟩ + · rw [W.add_of_Xne huz hvz <| by + simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm]; exact hx ∘ mul_left_cancel₀ (huv 2), + addX_of_Xne_smul, addY_of_Xne_smul, addZ_of_Xne_smul, W.add_of_Xne hPz hQz hx] + exact ⟨u ^ 3 * v ^ 3, + by simp_rw [smul_fin3, ← Units.val_pow_eq_pow_val, mul_pow, ← pow_mul]; rfl⟩ + +lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : W.add P Q ≈ W.add P' Q' := by + rcases hP, hQ with ⟨⟨u, rfl⟩, ⟨v, rfl⟩⟩ + exact W.add_smul_equiv P' Q' u v + +/-- The addition of two point classes. If `P` is a point representative, +then `W.add_map ⟦P⟧ ⟦Q⟧` is definitionally equivalent to `W.add P Q`. -/ +@[pp_dot] +noncomputable def add_map (P Q : PointClass R) : PointClass R := + Quotient.map₂ W.add (fun _ _ hP _ _ hQ => W.add_equiv hP hQ) P Q + +lemma add_map_eq (P Q : Fin 3 → R) : W.add_map ⟦P⟧ ⟦Q⟧ = ⟦W.add P Q⟧ := + rfl + +@[simp] +lemma add_map_of_Zeq0_left {P : Fin 3 → R} {Q : PointClass R} (hPz : P z = 0) : + W.add_map ⟦P⟧ Q = Q := by + rcases Q with ⟨Q⟩ + erw [add_map_eq, W.add_of_Zeq0_left hPz] + rfl + +lemma add_map_zero_left (P : PointClass R) : W.add_map ⟦![1, 1, 0]⟧ P = P := + W.add_map_of_Zeq0_left rfl + +@[simp] +lemma add_map_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦P⟧ := by + rw [add_map_eq, W.add_of_Zeq0_right hPz hQz] + +lemma add_map_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add_map ⟦P⟧ ⟦![1, 1, 0]⟧ = ⟦P⟧ := by + rw [add_map_eq, W.add_zero_right hPz] + +@[simp] +lemma add_map_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![1, 1, 0]⟧ := by + rw [add_map_eq, W.add_of_Yeq hPz hQz hx hy] + +@[simp] +lemma add_map_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P]⟧ := by + rw [add_map_eq, W.add_of_Yne hPz hQz hx hy] + +@[simp] +lemma add_map_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) + (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q]⟧ := by + rw [add_map_eq, W.add_of_Xne hPz hQz hx] + +variable {F : Type u} [Field F] {W : Jacobian F} + +@[simp] +lemma add_map_of_Zeq0_right' {P : PointClass F} {Q : Fin 3 → F} (hP : W.nonsingular_lift P) + (hQ : W.nonsingular Q) (hQz : Q z = 0) : W.add_map P ⟦Q⟧ = P := by + rcases P with ⟨P⟩ + by_cases hPz : P z = 0 + · erw [W.add_map_of_Zeq0_left hPz, Quotient.eq] + exact equiv_of_Zeq0 hQ hP hQz hPz + · exact W.add_map_of_Zeq0_right hPz hQz + +lemma add_map_zero_right' {P : PointClass F} (hP : W.nonsingular_lift P) : + W.add_map P ⟦![1, 1, 0]⟧ = P := + add_map_of_Zeq0_right' hP W.nonsingular_zero rfl + +variable {F : Type u} [Field F] {W : Jacobian F} + +/-- The negation of a nonsingular point representative in `W` lies in `W`. -/ +lemma nonsingular_neg {P : Fin 3 → F} (h : W.nonsingular P) : W.nonsingular <| W.neg P := by + by_cases hPz : P z = 0 + · rw [W.nonsingular_of_equiv <| W.neg_equiv <| equiv_zero_of_Zeq0 h hPz, neg_zero] + convert W.nonsingular_zero' <| neg_ne_zero.mpr one_ne_zero <;> norm_num1 + · rw [nonsingular_iff_affine_of_Zne0 <| by exact hPz] at h ⊢ + rwa [← Affine.nonsingular_neg_iff, ← negY_divZ hPz] at h + +lemma nonsingular_lift_neg_map {P : PointClass F} (h : W.nonsingular_lift P) : + W.nonsingular_lift <| W.neg_map P := by + rcases P with ⟨_⟩ + exact nonsingular_neg h + +/-- The addition of two nonsingular point representatives in `W` lies in `W`. -/ +lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : + W.nonsingular <| W.add P Q := by + by_cases hPz : P z = 0 + · rwa [W.nonsingular_of_equiv <| W.add_equiv (equiv_zero_of_Zeq0 hP hPz) <| Setoid.refl Q, + W.add_of_Zeq0_left <| by exact rfl] + · by_cases hQz : Q z = 0 + · rwa [W.nonsingular_of_equiv <| W.add_equiv (Setoid.refl P) <| equiv_zero_of_Zeq0 hQ hQz, + W.add_of_Zeq0_right hPz <| by exact rfl] + · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · simpa only [W.add_of_Yeq hPz hQz hx hy] using W.nonsingular_zero + · erw [W.add_of_Yne hPz hQz hx hy, + nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, + addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy] + exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) + (nonsingular_affine_of_Zne0 hQ hQz) fun _ => (negY_divZ hQz).symm ▸ Function.comp + (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp + · erw [W.add_of_Xne hPz hQz hx, + nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, + addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx] + exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) + (nonsingular_affine_of_Zne0 hQ hQz) fun h => False.elim <| hx <| + mul_comm (Q x) _ ▸ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp h + +lemma nonsingular_lift_add_map {P Q : PointClass F} (hP : W.nonsingular_lift P) + (hQ : W.nonsingular_lift Q) : W.nonsingular_lift <| W.add_map P Q := by + rcases P, Q with ⟨⟨_⟩, ⟨_⟩⟩ + exact nonsingular_add hP hQ + +end Representative + end WeierstrassCurve.Jacobian From 59a7cf410b830750d969a7c91986d793bbf2fdfe Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 4 Jan 2024 10:08:06 +0000 Subject: [PATCH 13/47] Implement group operations on nonsingular rational points for Jacobian coordinates --- .../EllipticCurve/Jacobian.lean | 220 ++++++++++++++++++ 1 file changed, 220 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 740b395de09849..3c7bf001914d77 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -44,6 +44,11 @@ Note that most computational proofs follow from their analogous proofs for affin * `WeierstrassCurve.Jacobian.neg_map`: the negation operation on a point class. * `WeierstrassCurve.Jacobian.add`: the addition operation on a point representative. * `WeierstrassCurve.Jacobian.add_map`: the addition operation on a point class. + * `WeierstrassCurve.Jacobian.Point`: a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.neg`: the negation operation on a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.add`: the addition operation on a nonsingular rational point. + * `WeierstrassCurve.Jacobian.Point.toAffine_addEquiv`: the equivalence between the nonsingular + rational points on a Jacobian Weierstrass curve with those on an affine Weierstrass curve. ## Main statements @@ -758,4 +763,219 @@ lemma nonsingular_lift_add_map {P Q : PointClass F} (hP : W.nonsingular_lift P) end Representative +section Point + +/-! ### Group operations on nonsingular rational points -/ + +/-- A nonsingular rational point on `W`. -/ +@[pp_dot] +structure Point where + {point : PointClass R} + (nonsingular : W.nonsingular_lift point) + +attribute [pp_dot] Point.point +attribute [pp_dot] Point.nonsingular + +/-- The point class underlying a nonsingular rational point on `W`. -/ +add_decl_doc Point.point + +/-- The nonsingular condition underlying a nonsingular rational point on `W`. -/ +add_decl_doc Point.nonsingular + +namespace Point + +variable {W} + +instance instZeroPoint [Nontrivial R] : Zero W.Point := + ⟨⟨W.nonsingular_lift_zero⟩⟩ + +lemma zero_def [Nontrivial R] : (⟨W.nonsingular_lift_zero⟩ : W.Point) = 0 := + rfl + +/-- The map from a nonsingular rational point on a Weierstrass curve `W` in affine coordinates +to the corresponding nonsingular rational point on `W` in Jacobian coordinates. -/ +def fromAffine [Nontrivial R] : W.toAffine.Point → W.Point + | 0 => 0 + | Affine.Point.some h => ⟨(W.nonsingular_lift_some ..).mpr h⟩ + +@[simp] +lemma fromAffine_zero [Nontrivial R] : fromAffine 0 = (0 : W.Point) := + rfl + +@[simp] +lemma fromAffine_some [Nontrivial R] {X Y : R} (h : W.toAffine.nonsingular X Y) : + fromAffine (Affine.Point.some h) = ⟨(W.nonsingular_lift_some ..).mpr h⟩ := + rfl + +variable {F : Type u} [Field F] {W : Jacobian F} + +/-- The negation of a nonsingular rational point on `W`. + +Given a nonsingular rational point `P` on `W`, use `-P` instead of `neg P`. -/ +@[pp_dot] +def neg (P : W.Point) : W.Point := + ⟨W.nonsingular_lift_neg_map P.nonsingular⟩ + +instance instNegPoint : Neg W.Point := + ⟨neg⟩ + +lemma neg_def (P : W.Point) : P.neg = -P := + rfl + +@[simp] +lemma neg_zero : (-⟨W.nonsingular_lift_zero⟩ : W.Point) = ⟨W.nonsingular_lift_zero⟩ := by + simp only [← neg_def, neg, neg_map_zero] + +/-- The addition of two nonsingular rational points on `W`. + +Given two nonsingular rational points `P` and `Q` on `W`, use `P + Q` instead of `add P Q`. -/ +@[pp_dot] +noncomputable def add (P Q : W.Point) : W.Point := + ⟨W.nonsingular_lift_add_map P.nonsingular Q.nonsingular⟩ + +noncomputable instance instAddPoint : Add W.Point := + ⟨add⟩ + +lemma add_def (P Q : W.Point) : P.add Q = P + Q := + rfl + +@[simp] +lemma zero_add (P : W.Point) : ⟨W.nonsingular_lift_zero⟩ + P = P := by + simp only [← add_def, add, add_map_zero_left] + +@[simp] +lemma add_zero (P : W.Point) : P + ⟨W.nonsingular_lift_zero⟩ = P := by + simp only [← add_def, add, add_map_zero_right' P.nonsingular] + +noncomputable instance instAddZeroClassPoint : AddZeroClass W.Point := + ⟨zero_add, add_zero⟩ + +open scoped Classical + +/-- The map from a point representative that is nonsingular on a Weierstrass curve `W` in Jacobian +coordinates to the corresponding nonsingular rational point on `W` in affine coordinates. -/ +noncomputable def toAffine {P : Fin 3 → F} (h : W.nonsingular P) : W.toAffine.Point := + if hPz : P z = 0 then 0 else Affine.Point.some <| nonsingular_affine_of_Zne0 h hPz + +lemma toAffine_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z = 0) : toAffine h = 0 := + dif_pos hPz + +lemma toAffine_zero : toAffine W.nonsingular_zero = 0 := + toAffine_of_Zeq0 rfl + +lemma toAffine_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z ≠ 0) : + toAffine h = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := + dif_neg hPz + +lemma toAffine_some {X Y : F} (h : W.nonsingular ![X, Y, 1]) : + toAffine h = Affine.Point.some ((W.nonsingular_some X Y).mp h) := by + rw [toAffine_of_Zne0 <| by exact one_ne_zero] + matrix_simp + simp only [one_pow, div_one] + +lemma toAffine_neg {P : Fin 3 → F} (hP : W.nonsingular P) : + toAffine (nonsingular_neg hP) = -toAffine hP := by + by_cases hPz : P z = 0 + · rw [toAffine_of_Zeq0 <| by exact hPz, toAffine_of_Zeq0 hPz, Affine.Point.neg_zero] + · rw [toAffine_of_Zne0 <| by exact hPz, toAffine_of_Zne0 hPz, Affine.Point.neg_some, + Affine.Point.some.injEq] + exact ⟨rfl, negY_divZ hPz⟩ + +lemma toAffine_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : + toAffine (nonsingular_add hP hQ) = toAffine hP + toAffine hQ := by + by_cases hPz : P z = 0 + · simp only [W.add_of_Zeq0_left (Q := Q) hPz, toAffine_of_Zeq0 hPz, _root_.zero_add] + · by_cases hQz : Q z = 0 + · simp only [W.add_of_Zeq0_right hPz hQz, toAffine_of_Zeq0 hQz, _root_.add_zero] + · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x + · have hx' : P x / P z ^ 2 = Q x / Q z ^ 2 := + (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mpr <| mul_comm (P z ^ 2) _ ▸ hx + rw [toAffine_of_Zne0 hPz, toAffine_of_Zne0 hQz] + by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q + · have hy' : P y / P z ^ 3 = W.negY Q / Q z ^ 3 := Iff.mpr + (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)) <| mul_comm (P z ^ 3) _ ▸ hy + simp only [W.add_of_Yeq hPz hQz hx hy] + rw [toAffine_zero, Affine.Point.some_add_some_of_Yeq hx' <| by rwa [← negY_divZ hQz]] + · have hy' : P y / P z ^ 3 ≠ W.negY Q / Q z ^ 3 := Function.comp + (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp + simp only [W.add_of_Yne hPz hQz hx hy] + rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, + Affine.Point.some_add_some_of_Yne hx' <| (negY_divZ hQz).symm ▸ hy', + Affine.Point.some.injEq] + exact ⟨addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy⟩ + · have hx' : P x / P z ^ 2 ≠ Q x / Q z ^ 2 := + (mul_comm (P z ^ 2) _ ▸ hx) ∘ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp + simp only [W.add_of_Xne hPz hQz hx] + rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, toAffine_of_Zne0 hPz, + toAffine_of_Zne0 hQz, Affine.Point.some_add_some_of_Xne hx', Affine.Point.some.injEq] + exact ⟨addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx⟩ + +lemma toAffine_of_equiv (P Q : Fin 3 → F) (h : P ≈ Q) : + HEq (toAffine (W := W) (P := P)) (toAffine (W := W) (P := Q)) := by + rcases h with ⟨u, rfl⟩ + refine Function.hfunext (propext <| W.nonsingular_smul_iff Q u) <| fun _ _ _ => ?_ + by_cases hPz : Q z = 0 + · rw [toAffine_of_Zeq0 <| by exact u.mul_right_eq_zero.mpr hPz, toAffine_of_Zeq0 hPz] + · rw [toAffine_of_Zne0 <| by exact mul_ne_zero u.ne_zero hPz, toAffine_of_Zne0 hPz, heq_eq_eq] + simp only [Affine.Point.some.injEq, smul_fin3_ext, mul_pow, + mul_div_mul_left _ _ <| pow_ne_zero _ u.ne_zero, and_self] + +/-- The map from a nonsingular rational point on a Weierstrass curve `W` in Jacobian coordinates +to the corresponding nonsingular rational point on `W` in affine coordinates. -/ +noncomputable def toAffine_lift (P : W.Point) : W.toAffine.Point := + P.point.hrecOn (fun _ => toAffine) toAffine_of_equiv P.nonsingular + +lemma toAffine_lift_eq {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : + toAffine_lift ⟨h⟩ = toAffine h := + rfl + +lemma toAffine_lift_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z = 0) : + toAffine_lift ⟨h⟩ = 0 := + toAffine_of_Zeq0 hPz (h := h) + +lemma toAffine_lift_zero : toAffine_lift (0 : W.Point) = 0 := + toAffine_zero + +lemma toAffine_lift_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z ≠ 0) : + toAffine_lift ⟨h⟩ = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := + toAffine_of_Zne0 hPz (h := h) + +lemma toAffine_lift_some {X Y : F} (h : W.nonsingular_lift ⟦![X, Y, 1]⟧) : + toAffine_lift ⟨h⟩ = Affine.Point.some ((W.nonsingular_some X Y).mp h) := + toAffine_some h + +lemma toAffine_lift_neg {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : + toAffine_lift (-⟨h⟩) = -toAffine_lift ⟨h⟩ := + toAffine_neg h + +lemma toAffine_lift_add {P Q : Fin 3 → F} (hP : W.nonsingular_lift ⟦P⟧) + (hQ : W.nonsingular_lift ⟦Q⟧) : + toAffine_lift (⟨hP⟩ + ⟨hQ⟩) = toAffine_lift ⟨hP⟩ + toAffine_lift ⟨hQ⟩ := + toAffine_add hP hQ + +/-- The equivalence between the nonsingular rational points on a Weierstrass curve `W` in Jacobian +coordinates with the nonsingular rational points on `W` in affine coordinates. -/ +@[simps] +noncomputable def toAffine_addEquiv : W.Point ≃+ W.toAffine.Point where + toFun := toAffine_lift + invFun := fromAffine + left_inv := by + rintro @⟨⟨P⟩, h⟩ + by_cases hPz : P z = 0 + · erw [toAffine_lift_eq, toAffine_of_Zeq0 hPz, fromAffine_zero, mk.injEq, Quotient.eq] + exact Setoid.symm <| equiv_zero_of_Zeq0 h hPz + · erw [toAffine_lift_eq, toAffine_of_Zne0 hPz, fromAffine_some, mk.injEq, Quotient.eq] + exact Setoid.symm <| equiv_some_of_Zne0 hPz + right_inv := by + rintro (_ | _) + · erw [fromAffine_zero, toAffine_lift_zero, Affine.Point.zero_def] + · rw [fromAffine_some, toAffine_lift_some] + map_add' := by + rintro @⟨⟨_⟩, _⟩ @⟨⟨_⟩, _⟩ + simpa only using toAffine_lift_add .. + +end Point + +end Point + end WeierstrassCurve.Jacobian From 18e2fd47d03de0cf89b69629239d7dd36bb46f59 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 5 Jan 2024 14:04:07 +0000 Subject: [PATCH 14/47] Remove extraneous factor in addition --- .../EllipticCurve/Jacobian.lean | 85 ++++++++++++------- 1 file changed, 56 insertions(+), 29 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 4afd8aa624c3de..88fbc2c4cd0f75 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -332,13 +332,14 @@ lemma negY_smul (P : Fin 3 → R) (u : Rˣ) : W.negY (u • P) = u ^ 3 * W.negY non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ @[pp_dot] def addX_of_Xne (P Q : Fin 3 → R) : R := - (P y * Q z ^ 3 - P z ^ 3 * Q y) ^ 2 - + W.a₁ * P z * Q z * (P y * Q z ^ 3 - P z ^ 3 * Q y) * (P x * Q z ^ 2 - P z ^ 2 * Q x) - - (W.a₂ * P z ^ 2 * Q z ^ 2 + P x * Q z ^ 2 + P z ^ 2 * Q x) - * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2 + P x * Q x ^ 2 * P z ^ 2 - 2 * P y * Q y * P z * Q z + P x ^ 2 * Q x * Q z ^ 2 + - W.a₁ * P x * Q y * P z ^ 2 * Q z - W.a₁ * P y * Q x * P z * Q z ^ 2 + + 2 * W.a₂ * P x * Q x * P z ^ 2 * Q z ^ 2 - W.a₃ * Q y * P z ^ 4 * Q z + - W.a₃ * P y * P z * Q z ^ 4 + W.a₄ * Q x * P z ^ 4 * Q z ^ 2 + W.a₄ * P x * P z ^ 2 * Q z ^ 4 + + 2 * W.a₆ * P z ^ 4 * Q z ^ 4 lemma addX_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addX_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addX_of_Xne P Q := by + W.addX_of_Xne (u • P) (v • Q) = (u : R) ^ 4 * (v : R) ^ 4 * W.addX_of_Xne P Q := by simp only [addX_of_Xne, smul_fin3_ext] ring1 @@ -361,13 +362,20 @@ negation that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where their $Z$-coordinates are $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ @[pp_dot] def addY'_of_Xne (P Q : Fin 3 → R) : R := - (P y * Q z ^ 3 - P z ^ 3 * Q y) - * (W.addX_of_Xne P Q - P x * Q z ^ 2 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) - + P y * Q z ^ 3 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 3 + -P y * Q x ^ 3 * P z ^ 3 + 2 * P y * Q y ^ 2 * P z ^ 3 - 3 * P x ^ 2 * Q x * Q y * P z ^ 2 * Q z + + 3 * P x * P y * Q x ^ 2 * P z * Q z ^ 2 + P x ^ 3 * Q y * Q z ^ 3 + - 2 * P y ^ 2 * Q y * Q z ^ 3 + W.a₁ * P x * Q y ^ 2 * P z ^ 4 + + W.a₁ * P y * Q x * Q y * P z ^ 3 * Q z - W.a₁ * P x * P y * Q y * P z * Q z ^ 3 + - W.a₁ * P y ^ 2 * Q x * Q z ^ 4 - 2 * W.a₂ * P x * Q x * Q y * P z ^ 4 * Q z + + 2 * W.a₂ * P x * P y * Q x * P z * Q z ^ 4 + W.a₃ * Q y ^ 2 * P z ^ 6 + - W.a₃ * P y ^ 2 * Q z ^ 6 - W.a₄ * Q x * Q y * P z ^ 6 * Q z + - W.a₄ * P x * Q y * P z ^ 4 * Q z ^ 3 + W.a₄ * P y * Q x * P z ^ 3 * Q z ^ 4 + + W.a₄ * P x * P y * P z * Q z ^ 6 - 2 * W.a₆ * Q y * P z ^ 6 * Q z ^ 3 + + 2 * W.a₆ * P y * P z ^ 3 * Q z ^ 6 lemma addY'_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addY'_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY'_of_Xne P Q := by - simp only [addY'_of_Xne, addX_of_Xne_smul, smul_fin3_ext] + W.addY'_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addY'_of_Xne P Q := by + simp only [addY'_of_Xne, smul_fin3_ext] ring1 /-- The $Y$-coordinate of the doubling of a point representative, before applying the final negation @@ -387,10 +395,10 @@ lemma addY'_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : /-- The $Z$-coordinate of the addition of two point representatives, where their $Z$-coordinates are non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ def addZ_of_Xne (P Q : Fin 3 → R) : R := - P z * Q z * (P x * Q z ^ 2 - P z ^ 2 * Q x) + P x * Q z ^ 2 - P z ^ 2 * Q x lemma addZ_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - addZ_of_Xne (u • P) (v • Q) = (u : R) ^ 3 * (v : R) ^ 3 * addZ_of_Xne P Q := by + addZ_of_Xne (u • P) (v • Q) = (u : R) ^ 2 * (v : R) ^ 2 * addZ_of_Xne P Q := by simp only [addZ_of_Xne, smul_fin3_ext] ring1 @@ -412,7 +420,7 @@ def addY_of_Xne (P Q : Fin 3 → R) : R := W.negY ![W.addX_of_Xne P Q, W.addY'_of_Xne P Q, addZ_of_Xne P Q] lemma addY_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addY_of_Xne (u • P) (v • Q) = (u : R) ^ 9 * (v : R) ^ 9 * W.addY_of_Xne P Q := by + W.addY_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addY_of_Xne P Q := by simp only [addY_of_Xne, negY, addX_of_Xne_smul, addY'_of_Xne_smul, addZ_of_Xne_smul] matrix_simp ring1 @@ -451,13 +459,22 @@ lemma Yne_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular ← neg_div, ← sub_div, div_left_inj' <| pow_ne_zero 3 hPz] at hy exact hy -lemma addX_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addX_of_Xne P Q / addZ_of_Xne P Q ^ 2 = - W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by +lemma addX_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.addX_of_Xne P Q / addZ_of_Xne P Q ^ 2 = W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + convert_to + ((P y * Q z ^ 3 - P z ^ 3 * Q y) ^ 2 + + W.a₁ * P z * Q z * (P y * Q z ^ 3 - P z ^ 3 * Q y) * (P x * Q z ^ 2 - P z ^ 2 * Q x) + - (W.a₂ * P z ^ 2 * Q z ^ 2 + P x * Q z ^ 2 + P z ^ 2 * Q x) + * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) + / (P z ^ 2 * Q z ^ 2) / addZ_of_Xne P Q ^ 2 = _ using 2 + · rw [nonsingular_iff, equation_iff] at hP hQ + rw [addX_of_Xne, eq_div_iff_mul_eq <| mul_ne_zero (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)] + linear_combination (norm := ring1) -Q z ^ 6 * hP.left - P z ^ 6 * hQ.left rw [Affine.slope_of_Xne <| by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] - field_simp [sub_ne_zero_of_ne hx, addX_of_Xne, addZ_of_Xne] + field_simp [sub_ne_zero_of_ne hx, addZ_of_Xne] ring1 lemma addX_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) @@ -472,13 +489,22 @@ lemma addX_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.no field_simp [addX_of_Yne, addX_of_Yne, addZ_of_Yne] ring1 -lemma addY'_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY'_of_Xne P Q / addZ_of_Xne P Q ^ 3 = +lemma addY'_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.addY'_of_Xne P Q / addZ_of_Xne P Q ^ 3 = W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - rw [Affine.addY', ← addX_div_addZ_of_Xne hPz hQz hx, Affine.slope_of_Xne <| + convert_to + ((P y * Q z ^ 3 - P z ^ 3 * Q y) * (P z ^ 2 * Q z ^ 2 * W.addX_of_Xne P Q + - P x * Q z ^ 2 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) + + P y * Q z ^ 3 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 3) + / (P z ^ 3 * Q z ^ 3) / addZ_of_Xne P Q ^ 3 = _ using 2 + · rw [addY'_of_Xne, addX_of_Xne, + eq_div_iff_mul_eq <| mul_ne_zero (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)] + ring1 + rw [Affine.addY', ← addX_div_addZ_of_Xne hP hQ hPz hQz hx, Affine.slope_of_Xne <| by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] - field_simp [sub_ne_zero_of_ne hx, addY'_of_Xne, addX_of_Xne, addZ_of_Xne] + field_simp [sub_ne_zero_of_ne hx, addX_of_Xne, addZ_of_Xne] ring1 lemma addY'_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) @@ -494,21 +520,22 @@ lemma addY'_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.n field_simp [addY'_of_Yne, addX_of_Yne, addZ_of_Yne] ring1 -lemma addZ_ne_zero_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : addZ_of_Xne P Q ≠ 0 := - mul_ne_zero (mul_ne_zero hPz hQz) <| sub_ne_zero_of_ne hx +lemma addZ_ne_zero_of_Xne {P Q : Fin 3 → F} (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + addZ_of_Xne P Q ≠ 0 := + sub_ne_zero_of_ne hx lemma addZ_ne_zero_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addZ_of_Yne P ≠ 0 := mul_ne_zero hPz <| sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy -lemma addY_div_addZ_of_Xne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : W.addY_of_Xne P Q / addZ_of_Xne P Q ^ 3 = +lemma addY_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) + (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : + W.addY_of_Xne P Q / addZ_of_Xne P Q ^ 3 = W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - rw [Affine.addY, ← addX_div_addZ_of_Xne hPz hQz hx, ← addY'_div_addZ_of_Xne hPz hQz hx] - exact negY_divZ <| addZ_ne_zero_of_Xne hPz hQz hx + simpa only [Affine.addY, ← addX_div_addZ_of_Xne hP hQ hPz hQz hx, + ← addY'_div_addZ_of_Xne hP hQ hPz hQz hx] using negY_divZ <| addZ_ne_zero_of_Xne hx lemma addY_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) From 44f37d6e55f2c2c0bb9076007f66553a211c5565 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 5 Jan 2024 14:10:44 +0000 Subject: [PATCH 15/47] Fix merge error --- Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index d8fdf423718b77..4bec96013ea68e 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -671,7 +671,7 @@ lemma add_smul_equiv (P Q : Fin 3 → R) (u v : Rˣ) : W.add (u • P) (v • Q) · rw [W.add_of_Xne huz hvz <| by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm]; exact hx ∘ mul_left_cancel₀ (huv 2), addX_of_Xne_smul, addY_of_Xne_smul, addZ_of_Xne_smul, W.add_of_Xne hPz hQz hx] - exact ⟨u ^ 3 * v ^ 3, + exact ⟨u ^ 2 * v ^ 2, by simp_rw [smul_fin3, ← Units.val_pow_eq_pow_val, mul_pow, ← pow_mul]; rfl⟩ lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : W.add P Q ≈ W.add P' Q' := by @@ -772,8 +772,8 @@ lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsing (nonsingular_affine_of_Zne0 hQ hQz) fun _ => (negY_divZ hQz).symm ▸ Function.comp (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp · erw [W.add_of_Xne hPz hQz hx, - nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, - addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx] + nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Xne hx, + addX_div_addZ_of_Xne hP hQ hPz hQz hx, addY_div_addZ_of_Xne hP hQ hPz hQz hx] exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) (nonsingular_affine_of_Zne0 hQ hQz) fun h => False.elim <| hx <| mul_comm (Q x) _ ▸ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp h From 8b3cc8386f33d34fc9b36c7611ac8257f7bbd2e7 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 5 Jan 2024 14:15:47 +0000 Subject: [PATCH 16/47] Fix merge error --- Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 7a0bac4ca7f2b4..949f6aada16571 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -933,9 +933,9 @@ lemma toAffine_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingula · have hx' : P x / P z ^ 2 ≠ Q x / Q z ^ 2 := (mul_comm (P z ^ 2) _ ▸ hx) ∘ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp simp only [W.add_of_Xne hPz hQz hx] - rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Xne hPz hQz hx, toAffine_of_Zne0 hPz, - toAffine_of_Zne0 hQz, Affine.Point.some_add_some_of_Xne hx', Affine.Point.some.injEq] - exact ⟨addX_div_addZ_of_Xne hPz hQz hx, addY_div_addZ_of_Xne hPz hQz hx⟩ + rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Xne hx, toAffine_of_Zne0 hPz, toAffine_of_Zne0 hQz, + Affine.Point.some_add_some_of_Xne hx', Affine.Point.some.injEq] + exact ⟨addX_div_addZ_of_Xne hP hQ hPz hQz hx, addY_div_addZ_of_Xne hP hQ hPz hQz hx⟩ lemma toAffine_of_equiv (P Q : Fin 3 → F) (h : P ≈ Q) : HEq (toAffine (W := W) (P := P)) (toAffine (W := W) (P := Q)) := by From d67fb6adff186ea041ea1697adb63abde5e97b12 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 21 Feb 2024 15:28:34 +0000 Subject: [PATCH 17/47] Expose the auxiliary sequence --- Mathlib/Data/Nat/Parity.lean | 3 + .../EllipticDivisibilitySequence.lean | 172 +++++++++++++----- 2 files changed, 134 insertions(+), 41 deletions(-) diff --git a/Mathlib/Data/Nat/Parity.lean b/Mathlib/Data/Nat/Parity.lean index 93d718727737a3..3bdf30387a396a 100644 --- a/Mathlib/Data/Nat/Parity.lean +++ b/Mathlib/Data/Nat/Parity.lean @@ -134,6 +134,9 @@ set_option linter.deprecated false in theorem not_even_bit1 (n : ℕ) : ¬Even (bit1 n) := by simp [bit1, parity_simps] #align nat.not_even_bit1 Nat.not_even_bit1 +theorem not_even_two_mul_add_one (n : ℕ) : ¬Even (2 * n + 1) := + odd_iff_not_even.mp <| odd_two_mul_add_one n + theorem two_not_dvd_two_mul_add_one (n : ℕ) : ¬2 ∣ 2 * n + 1 := by simp [add_mod] #align nat.two_not_dvd_two_mul_add_one Nat.two_not_dvd_two_mul_add_one diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index a94a4f9d6fa9f8..f111103f83053d 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ - import Mathlib.Data.Nat.Parity import Mathlib.Tactic.Linarith import Mathlib.Tactic.LinearCombination @@ -30,6 +29,8 @@ Some examples of EDSs include * `IsEllSequence`: a sequence indexed by integers is an elliptic sequence. * `IsDivSequence`: a sequence indexed by integers is a divisibility sequence. * `IsEllDivSequence`: a sequence indexed by integers is an EDS. + * `preNormEDS'`: the auxiliary sequence for a normalised EDS indexed by `ℕ`. + * `preNormEDS`: the auxiliary sequence for a normalised EDS indexed by `ℤ`. * `normEDS'`: the canonical example of a normalised EDS indexed by `ℕ`. * `normEDS`: the canonical example of a normalised EDS indexed by `ℤ`. @@ -40,15 +41,17 @@ Some examples of EDSs include ## Implementation notes -`IsEllDivSequence' b c d n` is defined in terms of the private `IsEllDivSequence'' b c d n`, -which are equal when `n` is odd and differ by a factor of `b` when `n` is even. This coincides with -the reference since both agree for `IsEllDivSequence' b c d 2` and for `IsEllDivSequence' b c d 4`, -and the correct factors of `b` are removed in `IsEllDivSequence' b c d (2 * (m + 2) + 1)` and in -`IsEllDivSequence' b c d (2 * (m + 3))`. This is done to avoid the necessity for ring division by -`b` in the inductive definition of `IsEllDivSequence' b c d (2 * (m + 3))`. The idea is that, an -easy lemma shows that `IsEllDivSequence' b c d (2 * (m + 3))` always contains a factor of `b`, so it -is possible to remove a factor of `b` a posteriori, but stating this lemma requires first defining -`IsEllDivSequence' b c d (2 * (m + 3))`, which requires having this factor of `b` a priori. +The normalised EDS `normEDS b c d n` (resp `normEDS' b c d n`) is defined in terms of the auxiliary +sequence `preNormEDS (b ^ 4) c d n` (resp `preNormEDS' (b ^ 4) c d n`), which are equal when `n` is +odd and differ by a factor of `b` when `n` is even. This coincides with the definition in the +reference since both agree for `normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors +of `b` are removed in `normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. +One reason is to avoid the necessity for ring division by `b` in the inductive definition of +`normEDS b c d (2 * (m + 3))`. The idea is that, it can be shown that `normEDS b c d (2 * (m + 3))` +always contains a factor of `b`, so it is possible to remove a factor of `b` *a posteriori*, +but stating this lemma requires first defining `normEDS b c d (2 * (m + 3))`, which requires having +this factor of `b` *a priori*. Another reason is to allow the definition of univariate $n$-division +polynomials of elliptic curves, omitting a factor of the bivariate $2$-division polynomial. ## References @@ -97,7 +100,9 @@ lemma IsEllDivSequence_mul (x : R) {W : ℤ → R} (h : IsEllDivSequence W) : IsEllDivSequence (x • W) := ⟨IsEllSequence_mul x h.left, IsDivSequence_mul x h.right⟩ -private def normEDS'' (b c d : R) : ℕ → R +/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, +with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ +def preNormEDS' (b c d : R) : ℕ → R | 0 => 0 | 1 => 1 | 2 => 1 @@ -109,49 +114,129 @@ private def normEDS'' (b c d : R) : ℕ → R have h2 : m + 2 < n + 5 := (lt_add_one _).trans h3 have h1 : m + 1 < n + 5 := (lt_add_one _).trans h2 if hn : Even n then - normEDS'' b c d (m + 4) * normEDS'' b c d (m + 2) ^ 3 * (if Even m then b ^ 4 else 1) - - normEDS'' b c d (m + 1) * normEDS'' b c d (m + 3) ^ 3 * (if Even m then 1 else b ^ 4) + preNormEDS' b c d (m + 4) * preNormEDS' b c d (m + 2) ^ 3 * (if Even m then b else 1) - + preNormEDS' b c d (m + 1) * preNormEDS' b c d (m + 3) ^ 3 * (if Even m then 1 else b) else have h5 : m + 5 < n + 5 := add_lt_add_right (Nat.div_lt_self (Nat.odd_iff_not_even.mpr hn).pos <| Nat.lt_succ_self 1) 5 - normEDS'' b c d (m + 2) ^ 2 * normEDS'' b c d (m + 3) * normEDS'' b c d (m + 5) - - normEDS'' b c d (m + 1) * normEDS'' b c d (m + 3) * normEDS'' b c d (m + 4) ^ 2 + preNormEDS' b c d (m + 2) ^ 2 * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 5) - + preNormEDS' b c d (m + 1) * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 4) ^ 2 variable (b c d : R) +@[simp] +lemma preNormEDS'_zero : preNormEDS' b c d 0 = 0 := + rfl + +@[simp] +lemma preNormEDS'_one : preNormEDS' b c d 1 = 1 := + rfl + +@[simp] +lemma preNormEDS'_two : preNormEDS' b c d 2 = 1 := + rfl + +@[simp] +lemma preNormEDS'_three : preNormEDS' b c d 3 = c := + rfl + +@[simp] +lemma preNormEDS'_four : preNormEDS' b c d 4 = d := + rfl + +lemma preNormEDS'_odd (m : ℕ) : preNormEDS' b c d (2 * (m + 2) + 1) = + preNormEDS' b c d (m + 4) * preNormEDS' b c d (m + 2) ^ 3 * (if Even m then b else 1) - + preNormEDS' b c d (m + 1) * preNormEDS' b c d (m + 3) ^ 3 * (if Even m then 1 else b) := by + rw [show 2 * (m + 2) + 1 = 2 * m + 5 by rfl, preNormEDS', dif_pos <| even_two_mul _] + simp only [m.mul_div_cancel_left two_pos] + +lemma preNormEDS'_even (m : ℕ) : preNormEDS' b c d (2 * (m + 3)) = + preNormEDS' b c d (m + 2) ^ 2 * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 5) - + preNormEDS' b c d (m + 1) * preNormEDS' b c d (m + 3) * preNormEDS' b c d (m + 4) ^ 2 := by + rw [show 2 * (m + 3) = 2 * m + 1 + 5 by rfl, preNormEDS', dif_neg m.not_even_two_mul_add_one] + simp only [Nat.mul_add_div two_pos] + rfl + +/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, +with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. + +This extends `preNormEDS'` by defining its values at negative integers. -/ +def preNormEDS (n : ℤ) : R := + n.sign * preNormEDS' b c d n.natAbs + +@[simp] +lemma preNormEDS_ofNat (n : ℕ) : preNormEDS b c d n = preNormEDS' b c d n := by + by_cases hn : n = 0 + · rw [hn, preNormEDS, Nat.cast_zero, Int.sign_zero, Int.cast_zero, zero_mul, preNormEDS'_zero] + · rw [preNormEDS, Int.sign_coe_nat_of_nonzero hn, Int.cast_one, one_mul, Int.natAbs_cast] + +@[simp] +lemma preNormEDS_zero : preNormEDS b c d 0 = 0 := by + erw [preNormEDS_ofNat, preNormEDS'_zero] + +@[simp] +lemma preNormEDS_one : preNormEDS b c d 1 = 1 := by + erw [preNormEDS_ofNat, preNormEDS'_one] + +@[simp] +lemma preNormEDS_two : preNormEDS b c d 2 = 1 := by + erw [preNormEDS_ofNat, preNormEDS'_two] + +@[simp] +lemma preNormEDS_three : preNormEDS b c d 3 = c := by + erw [preNormEDS_ofNat, preNormEDS'_three] + +@[simp] +lemma preNormEDS_four : preNormEDS b c d 4 = d := by + erw [preNormEDS_ofNat, preNormEDS'_four] + +lemma preNormEDS_odd (m : ℕ) : preNormEDS b c d (2 * (m + 2) + 1) = + preNormEDS b c d (m + 4) * preNormEDS b c d (m + 2) ^ 3 * (if Even m then b else 1) - + preNormEDS b c d (m + 1) * preNormEDS b c d (m + 3) ^ 3 * (if Even m then 1 else b) := by + repeat erw [preNormEDS_ofNat] + exact preNormEDS'_odd .. + +lemma preNormEDS_even (m : ℕ) : preNormEDS b c d (2 * (m + 3)) = + preNormEDS b c d (m + 2) ^ 2 * preNormEDS b c d (m + 3) * preNormEDS b c d (m + 5) - + preNormEDS b c d (m + 1) * preNormEDS b c d (m + 3) * preNormEDS b c d (m + 4) ^ 2 := by + repeat erw [preNormEDS_ofNat] + exact preNormEDS'_even .. + +@[simp] +lemma preNormEDS_neg (n : ℤ) : preNormEDS b c d (-n) = -preNormEDS b c d n := by + rw [preNormEDS, Int.sign_neg, Int.cast_neg, neg_mul, Int.natAbs_neg, preNormEDS] + /-- The canonical example of a normalised EDS `W : ℕ → R`, -with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = b * d`. +with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = d * b`. -This is defined in terms of a truncated sequence whose even terms differ by a factor of `b`. -/ +This is defined in terms of `preNormEDS'` whose even terms differ by a factor of `b`. -/ def normEDS' (n : ℕ) : R := - normEDS'' b c d n * if Even n then b else 1 + preNormEDS' (b ^ 4) c d n * if Even n then b else 1 @[simp] lemma normEDS'_zero : normEDS' b c d 0 = 0 := by - rw [normEDS', normEDS'', zero_mul] + rw [normEDS', preNormEDS'_zero, zero_mul] @[simp] lemma normEDS'_one : normEDS' b c d 1 = 1 := by - rw [normEDS', normEDS'', one_mul, if_neg Nat.not_even_one] + rw [normEDS', preNormEDS'_one, one_mul, if_neg Nat.not_even_one] @[simp] lemma normEDS'_two : normEDS' b c d 2 = b := by - rw [normEDS', normEDS'', one_mul, if_pos even_two] + rw [normEDS', preNormEDS'_two, one_mul, if_pos even_two] @[simp] lemma normEDS'_three : normEDS' b c d 3 = c := by - rw [normEDS', normEDS'', if_neg <| by decide, mul_one] + rw [normEDS', preNormEDS'_three, if_neg <| by decide, mul_one] @[simp] lemma normEDS'_four : normEDS' b c d 4 = d * b := by - rw [normEDS', normEDS'', if_pos <| by decide] + rw [normEDS', preNormEDS'_four, if_pos <| by decide] lemma normEDS'_odd (m : ℕ) : normEDS' b c d (2 * (m + 2) + 1) = normEDS' b c d (m + 4) * normEDS' b c d (m + 2) ^ 3 - normEDS' b c d (m + 1) * normEDS' b c d (m + 3) ^ 3 := by - rw [normEDS', if_neg <| fun h => Nat.even_add_one.mp h <| even_two_mul _, - show 2 * (m + 2) + 1 = 2 * m + 5 by rfl, normEDS'', dif_pos <| even_two_mul m] - simp only [normEDS', Nat.mul_div_right _ zero_lt_two] + simp only [normEDS', preNormEDS'_odd, if_neg (m + 2).not_even_two_mul_add_one] by_cases hm : Even m · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 @@ -169,9 +254,7 @@ lemma normEDS'_odd (m : ℕ) : normEDS' b c d (2 * (m + 2) + 1) = lemma normEDS'_even (m : ℕ) : normEDS' b c d (2 * (m + 3)) * b = normEDS' b c d (m + 2) ^ 2 * normEDS' b c d (m + 3) * normEDS' b c d (m + 5) - normEDS' b c d (m + 1) * normEDS' b c d (m + 3) * normEDS' b c d (m + 4) ^ 2 := by - rw [normEDS', if_pos <| even_two_mul _, show 2 * (m + 3) = 2 * m + 1 + 5 by rfl, - normEDS'', dif_neg <| fun h => Nat.even_add_one.mp h <| even_two_mul _] - simp only [normEDS', Nat.mul_add_div two_pos, show 1 / 2 = 0 by rfl] + simp only [normEDS', preNormEDS'_even, if_pos <| even_two_mul _] by_cases hm : Even m · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 @@ -189,43 +272,50 @@ lemma normEDS'_even (m : ℕ) : normEDS' b c d (2 * (m + 3)) * b = ring1 /-- The canonical example of a normalised EDS `W : ℤ → R`, -with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = b * d`. +with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = d * b`. This extends `normEDS'` by defining its values at negative integers. -/ -def normEDS (n : ℤ) : R := n.sign * normEDS' b c d n.natAbs +def normEDS (n : ℤ) : R := + n.sign * normEDS' b c d n.natAbs + +@[simp] +lemma normEDS_ofNat (n : ℕ) : normEDS b c d n = normEDS' b c d n := by + by_cases hn : n = 0 + · rw [hn, normEDS, Nat.cast_zero, Int.sign_zero, Int.cast_zero, zero_mul, normEDS'_zero] + · rw [normEDS, Int.sign_coe_nat_of_nonzero hn, Int.cast_one, one_mul, Int.natAbs_cast] @[simp] lemma normEDS_zero : normEDS b c d 0 = 0 := by - erw [normEDS, Int.cast_zero, zero_mul] + erw [normEDS_ofNat, normEDS'_zero] @[simp] lemma normEDS_one : normEDS b c d 1 = 1 := by - erw [normEDS, Int.cast_one, one_mul, normEDS'_one] + erw [normEDS_ofNat, normEDS'_one] @[simp] lemma normEDS_two : normEDS b c d 2 = b := by - erw [normEDS, Int.cast_one, one_mul, normEDS'_two] + erw [normEDS_ofNat, normEDS'_two] @[simp] lemma normEDS_three : normEDS b c d 3 = c := by - erw [normEDS, Int.cast_one, one_mul, normEDS'_three] + erw [normEDS_ofNat, normEDS'_three] @[simp] lemma normEDS_four : normEDS b c d 4 = d * b := by - erw [normEDS, Int.cast_one, one_mul, normEDS'_four] + erw [normEDS_ofNat, normEDS'_four] lemma normEDS_odd (m : ℕ) : normEDS b c d (2 * (m + 2) + 1) = normEDS b c d (m + 4) * normEDS b c d (m + 2) ^ 3 - normEDS b c d (m + 1) * normEDS b c d (m + 3) ^ 3 := by - repeat erw [normEDS, Int.cast_one, one_mul] - exact normEDS'_odd b c d m + repeat erw [normEDS_ofNat] + exact normEDS'_odd .. lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = normEDS b c d (m + 2) ^ 2 * normEDS b c d (m + 3) * normEDS b c d (m + 5) - normEDS b c d (m + 1) * normEDS b c d (m + 3) * normEDS b c d (m + 4) ^ 2 := by - repeat erw [normEDS, Int.cast_one, one_mul] - exact normEDS'_even b c d m + repeat erw [normEDS_ofNat] + exact normEDS'_even .. @[simp] lemma normEDS_neg (n : ℕ) : normEDS b c d (-n) = -normEDS b c d n := by - simp [normEDS] + rw [normEDS, Int.sign_neg, Int.cast_neg, neg_mul, Int.natAbs_neg, normEDS] From 9abf5def655e0d44fd16e9b8b243b67025e7776f Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 22 Feb 2024 00:33:37 +0000 Subject: [PATCH 18/47] Define recursion principle for normalised EDS --- Mathlib/Data/Nat/EvenOddRec.lean | 16 +++++------ .../EllipticDivisibilitySequence.lean | 28 ++++++++++++++++++- 2 files changed, 35 insertions(+), 9 deletions(-) diff --git a/Mathlib/Data/Nat/EvenOddRec.lean b/Mathlib/Data/Nat/EvenOddRec.lean index 085c3ed55f705f..3e46fe9b289aaa 100644 --- a/Mathlib/Data/Nat/EvenOddRec.lean +++ b/Mathlib/Data/Nat/EvenOddRec.lean @@ -27,14 +27,14 @@ def evenOddRec {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ n, P n → P (2 * n) #align nat.even_odd_rec Nat.evenOddRec @[simp] -theorem evenOddRec_zero (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) - (h_odd : ∀ i, P i → P (2 * i + 1)) : @evenOddRec _ h0 h_even h_odd 0 = h0 := +theorem evenOddRec_zero {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) + (h_odd : ∀ i, P i → P (2 * i + 1)) : evenOddRec h0 h_even h_odd 0 = h0 := binaryRec_zero _ _ #align nat.even_odd_rec_zero Nat.evenOddRec_zero @[simp] -theorem evenOddRec_even (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) - (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) : +theorem evenOddRec_even {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) + (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) (n : ℕ) : (2 * n).evenOddRec h0 h_even h_odd = h_even n (evenOddRec h0 h_even h_odd n) := have : ∀ a, bit false n = a → HEq (@evenOddRec _ h0 h_even h_odd a) (h_even n (evenOddRec h0 h_even h_odd n)) @@ -43,8 +43,8 @@ theorem evenOddRec_even (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i #align nat.even_odd_rec_even Nat.evenOddRec_even @[simp] -theorem evenOddRec_odd (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) - (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) : +theorem evenOddRec_odd {P : ℕ → Sort*} (h0 : P 0) (h_even : ∀ i, P i → P (2 * i)) + (h_odd : ∀ i, P i → P (2 * i + 1)) (H : h_even 0 h0 = h0) (n : ℕ) : (2 * n + 1).evenOddRec h0 h_even h_odd = h_odd n (evenOddRec h0 h_even h_odd n) := have : ∀ a, bit true n = a → HEq (@evenOddRec _ h0 h_even h_odd a) (h_odd n (evenOddRec h0 h_even h_odd n)) @@ -56,9 +56,9 @@ theorem evenOddRec_odd (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i, from `P j` for all `j < 2 * i` and we can prove `P (2 * i + 1)` from `P j` for all `j < 2 * i + 1`, then we have `P n` for all `n : ℕ`. -/ @[elab_as_elim] -noncomputable def evenOddStrongRec {P : ℕ → Sort*} (n : ℕ) +noncomputable def evenOddStrongRec {P : ℕ → Sort*} (h_even : ∀ n : ℕ, (∀ k < 2 * n, P k) → P (2 * n)) - (h_odd : ∀ n : ℕ, (∀ k < 2 * n + 1, P k) → P (2 * n + 1)) : P n := + (h_odd : ∀ n : ℕ, (∀ k < 2 * n + 1, P k) → P (2 * n + 1)) (n : ℕ) : P n := n.strongRecOn fun m ih => m.even_or_odd'.choose_spec.by_cases (fun h => h.symm ▸ h_even m.even_or_odd'.choose <| h ▸ ih) (fun h => h.symm ▸ h_odd m.even_or_odd'.choose <| h ▸ ih) diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index f111103f83053d..bcc521a8aea5eb 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.Data.Nat.Parity +import Mathlib.Data.Nat.EvenOddRec import Mathlib.Tactic.Linarith import Mathlib.Tactic.LinearCombination @@ -319,3 +319,29 @@ lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = @[simp] lemma normEDS_neg (n : ℕ) : normEDS b c d (-n) = -normEDS b c d n := by rw [normEDS, Int.sign_neg, Int.cast_neg, neg_mul, Int.natAbs_neg, normEDS] + +/-- Strong recursion principle for a normalised EDS indexed by `ℕ`: if we have + * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, + * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P k` for all `k < 2 * (m + 3)`, and + * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P k` for all `k < 2 * (m + 2) + 1`, +then we have `P n` for all `n : ℕ`. -/ +@[elab_as_elim] +noncomputable def normEDSRec' {P : ℕ → Sort u} (base0 : P 0) (base1 : P 1) (base2 : P 2) + (base3 : P 3) (base4 : P 4) (even : ∀ m : ℕ, (∀ k < 2 * (m + 3), P k) → P (2 * (m + 3))) + (odd : ∀ m : ℕ, (∀ k < 2 * (m + 2) + 1, P k) → P (2 * (m + 2) + 1)) (n : ℕ) : P n := + n.evenOddStrongRec (by rintro (_ | _ | _ | _) h; exacts [base0, base2, base4, even _ h]) + (by rintro (_ | _ | _) h; exacts [base1, base3, odd _ h]) + +/-- Strong recursion principle for a normalised EDS indexed by `ℤ`: if we have + * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, + * for all `m : ℕ` we can prove `P (2 * (m + 3))` from `P k` for all `k < 2 * (m + 3)`, + * for all `m : ℕ` we can prove `P (2 * (m + 2) + 1)` from `P k` for all `k < 2 * (m + 2) + 1`, and + * for all `n : ℕ` we can extend from `P n` to `P (-n)`, +then we have `P n` for all `n : ℤ`. -/ +@[elab_as_elim] +noncomputable def normEDSRec {P : ℤ → Sort u} (base0 : P 0) (base1 : P 1) (base2 : P 2) + (base3 : P 3) (base4 : P 4) (even : ∀ m : ℕ, (∀ k < 2 * (m + 3), P k) → P (2 * (m + 3))) + (odd : ∀ m : ℕ, (∀ k < 2 * (m + 2) + 1, P k) → P (2 * (m + 2) + 1)) + (neg : ∀ n : ℕ, P n → P (-n)) : ∀ n : ℤ, P n + | .ofNat n => normEDSRec' base0 base1 base2 base3 base4 even odd n + | .negSucc n => neg (n + 1) <| normEDSRec' base0 base1 base2 base3 base4 even odd <| n + 1 From 59b7909f867ab074628916e90e7c2d20f6e82fea Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 22 Feb 2024 00:43:10 +0000 Subject: [PATCH 19/47] Redefine division polynomials --- .../EllipticCurve/DivisionPolynomial.lean | 308 ++++++++++++++++++ .../EllipticCurve/Torsion.lean | 70 ---- 2 files changed, 308 insertions(+), 70 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean delete mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean new file mode 100644 index 00000000000000..9d195c6e534669 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -0,0 +1,308 @@ +/- +Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ +import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.Data.Int.Parity +import Mathlib.NumberTheory.EllipticDivisibilitySequence + +/-! +# Division polynomials of Weierstrass curves + +This file defines certain polynomials associated to division polynomials of Weierstrass curves and +computes their leading terms. These are defined in terms of the auxiliary sequences of normalised +elliptic divisibility sequences defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`. + +## Mathematical background + +Let $W$ be a Weierstrass curve over a commutative ring $R$. The sequence of $n$-division polynomials +of $W$ is a normalised elliptic divisibility sequence $\psi_n \in R[X][Y]$ with initial values + * $\psi_0 = 0$, + * $\psi_1 = 1$, + * $\psi_2 = 2Y + a_1X + a_3$, + * $\psi_3 = 3X^4 + b_2X^3 + 3b_4X ^ 2 + 3b_6X + b_8$, and + * $\psi_4 = \psi_2 \cdot + (2X^6 + b_2X^5 + 5b_4X^4 + 10b_6X^3 + 10b_8X^2 + (b_2b_8 - b_4b_6)X + (b_4b_8 - b_6^2))$. +Furthermore, define associated sequences $\phi_n, \omega_n, \widetilde{\psi_n} \in R[X][Y]$ by + * $\phi_n := x\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, + * $\omega_n := \tfrac{1}{2\psi_n} \cdot (\psi_{2n} - \psi_n^2(a_1\phi_n + a_3\psi_n^2)$, and + * $\tilde{\psi}_n := \psi_n / \psi_2$ if $n$ is even and $\tilde{\psi}_n := \psi_n$ otherwise. +Under the coordinate ring $R[W]$ of $W$, the $2$-division polynomial $\psi_2$ satisfies the identity +$\psi_2^2 \equiv 4X^3 + b_2X^2 + 2b_4X + b_6$, so that $\phi_n, \psi_n^2, \tilde{\psi}_n \in R[X]$. +Furthermore, their leading terms can be computed via induction to be $\phi_n = X^{n^2} + \dots$ +and $\psi_n^2 = n^2X^{n^2 - 1} + \dots$, and $\tilde{\psi}_n = nX^{\tfrac{n^2 - 4}{2}} + \dots$ +if $n$ is even or $\tilde{\psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ otherwise. + +## Main definitions + + * `WeierstrassCurve.divisionPolynomial2Sq`: $\psi_2^2$ viewed as a univariate polynomial. + * `WeierstrassCurve.divisionPolynomial`: $\tilde{\psi}_n$ viewed as a univariate polynomial. + * `WeierstrassCurve.divisionPolynomialZSq`: $\psi_n^2$ viewed as a univariate polynomial. + * `WeierstrassCurve.divisionPolynomialX`: $\phi_n$ viewed as a univariate polynomial. + * `WeierstrassCurve.divisionPolynomialZ`: $\psi_n$ viewed as a bivariate polynomial. + * TODO: $\omega$ viewed as a bivariate polynomial. + +## Implementation notes + +Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate +polynomials $\psi_n$ are defined in terms of the univariate polynomials $\tilde{\psi}_n$. This is +done partially to avoid ring division, but more crucially to allow the definition of $\phi_n$ and +$\psi_n^2$ as univariate polynomials without needing to work under the coordinate ring, and to allow +the computation of their leading terms without ambiguity. Furthermore, evaluating these polynomials +at a rational point $(x, y)$ on $W$ recovers their original definition up to linear combinations of +the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring. + +## References + +[J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] + +## Tags + +elliptic curve, division polynomial, torsion point +-/ + +open Polynomial PolynomialPolynomial + +local macro "C_simp" : tactic => + `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) + +universe u + +namespace WeierstrassCurve + +variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) + +/-- The polynomial $4X^3 + b_2X^2 + 2b_4X + b_6$ congruent to the square $\psi_2^2$ of the +$2$-division polynomial $\psi_2 = 2Y + a_1X + a_3$ under $R[W]$. -/ +@[pp_dot] +noncomputable def divisionPolynomial2Sq : R[X] := + C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ + +-- TODO: remove `twoTorsionPolynomial` in favour of `divisionPolynomial2Sq` +lemma divisionPolynomial2Sq_eq : W.divisionPolynomial2Sq = W.twoTorsionPolynomial.toPoly := + rfl + +/-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials +$\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{N}$. -/ +@[pp_dot] +noncomputable def divisionPolynomial' (n : ℕ) : R[X] := + preNormEDS' (W.divisionPolynomial2Sq ^ 2) + (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) + (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) n + +@[simp] +lemma divisionPolynomial'_zero : W.divisionPolynomial' 0 = 0 := + preNormEDS'_zero .. + +@[simp] +lemma divisionPolynomial'_one : W.divisionPolynomial' 1 = 1 := + preNormEDS'_one .. + +@[simp] +lemma divisionPolynomial'_two : W.divisionPolynomial' 2 = 1 := + preNormEDS'_two .. + +@[simp] +lemma divisionPolynomial'_three : W.divisionPolynomial' 3 = + 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ := + preNormEDS'_three .. + +@[simp] +lemma divisionPolynomial'_four : W.divisionPolynomial' 4 = + 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) := + preNormEDS'_four .. + +lemma divisionPolynomial'_odd (m : ℕ) : W.divisionPolynomial' (2 * (m + 2) + 1) = + W.divisionPolynomial' (m + 4) * W.divisionPolynomial' (m + 2) ^ 3 * + (if Even m then W.divisionPolynomial2Sq ^ 2 else 1) - + W.divisionPolynomial' (m + 1) * W.divisionPolynomial' (m + 3) ^ 3 * + (if Even m then 1 else W.divisionPolynomial2Sq ^ 2) := + preNormEDS'_odd .. + +lemma divisionPolynomial'_even (m : ℕ) : W.divisionPolynomial' (2 * (m + 3)) = + W.divisionPolynomial' (m + 2) ^ 2 * W.divisionPolynomial' (m + 3) * + W.divisionPolynomial' (m + 5) - + W.divisionPolynomial' (m + 1) * W.divisionPolynomial' (m + 3) * + W.divisionPolynomial' (m + 4) ^ 2 := + preNormEDS'_even .. + +/-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials +$\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ +@[pp_dot] +noncomputable def divisionPolynomial (n : ℤ) : R[X] := + n.sign * W.divisionPolynomial' n.natAbs + +@[simp] +lemma divisionPolynomial_ofNat (n : ℕ) : W.divisionPolynomial n = W.divisionPolynomial' n := + preNormEDS_ofNat .. + +@[simp] +lemma divisionPolynomial_zero : W.divisionPolynomial 0 = 0 := + preNormEDS_zero .. + +@[simp] +lemma divisionPolynomial_one : W.divisionPolynomial 1 = 1 := + preNormEDS_one .. + +@[simp] +lemma divisionPolynomial_two : W.divisionPolynomial 2 = 1 := + preNormEDS_two .. + +@[simp] +lemma divisionPolynomial_three : W.divisionPolynomial 3 = W.divisionPolynomial' 3 := + preNormEDS_three .. + +@[simp] +lemma divisionPolynomial_four : W.divisionPolynomial 4 = W.divisionPolynomial' 4 := + preNormEDS_four .. + +lemma divisionPolynomial_odd (m : ℕ) : W.divisionPolynomial (2 * (m + 2) + 1) = + W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 * + (if Even m then W.divisionPolynomial2Sq ^ 2 else 1) - + W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 * + (if Even m then 1 else W.divisionPolynomial2Sq ^ 2) := + preNormEDS_odd .. + +lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) = + W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * + W.divisionPolynomial (m + 5) - + W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) * + W.divisionPolynomial (m + 4) ^ 2 := + preNormEDS_even .. + +@[simp] +lemma divisionPolynomial_neg (n : ℤ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := + preNormEDS_neg .. + +/-- The univariate polynomials congruent under $R[W]$ to the bivariate squares $\psi_n^2$ of the +$n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ +@[pp_dot] +noncomputable def divisionPolynomialZSq (n : ℤ) : R[X] := + W.divisionPolynomial n ^ 2 * if Even n.natAbs then W.divisionPolynomial2Sq else 1 + +@[simp] +lemma divisionPolynomialZSq_ofNat (n : ℕ) : W.divisionPolynomialZSq n = + W.divisionPolynomial' n ^ 2 * if Even n then W.divisionPolynomial2Sq else 1 := by + rw [divisionPolynomialZSq, divisionPolynomial_ofNat, Int.natAbs_cast] + +@[simp] +lemma divisionPolynomialZSq_zero : W.divisionPolynomialZSq 0 = 0 := by + erw [divisionPolynomialZSq_ofNat, zero_pow two_ne_zero, zero_mul] + +@[simp] +lemma divisionPolynomialZSq_one : W.divisionPolynomialZSq 1 = 1 := by + erw [divisionPolynomialZSq_ofNat, one_pow, mul_one] + +@[simp] +lemma divisionPolynomialZSq_two : W.divisionPolynomialZSq 2 = W.divisionPolynomial2Sq := by + erw [divisionPolynomialZSq_ofNat, one_pow, one_mul, if_pos even_two] + +@[simp] +lemma divisionPolynomialZSq_three : W.divisionPolynomialZSq 3 = (W.divisionPolynomial' 3) ^ 2 := by + erw [divisionPolynomialZSq_ofNat, mul_one] + +@[simp] +lemma divisionPolynomialZSq_four : W.divisionPolynomialZSq 4 = + (W.divisionPolynomial' 4) ^ 2 * W.divisionPolynomial2Sq := by + erw [divisionPolynomialZSq_ofNat, if_pos <| by decide] + +@[simp] +lemma divisionPolynomialZSq_neg (n : ℤ) : + W.divisionPolynomialZSq (-n) = W.divisionPolynomialZSq n := by + rw [divisionPolynomialZSq, divisionPolynomial_neg, neg_sq, Int.natAbs_neg, divisionPolynomialZSq] + +/-- The univariate polynomials congruent under $R[W]$ to the bivariate polynomials $\phi_n$ defined +in terms of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ +@[pp_dot] +noncomputable def divisionPolynomialX (n : ℤ) : R[X] := + X * W.divisionPolynomialZSq n - W.divisionPolynomial (n + 1) * W.divisionPolynomial (n - 1) * + if Even n.natAbs then 1 else W.divisionPolynomial2Sq + +@[simp] +lemma divisionPolynomialX_ofNat (n : ℕ) : W.divisionPolynomialX (n + 1) = + X * W.divisionPolynomial' (n + 1) ^ 2 * + (if Even n then 1 else W.divisionPolynomial2Sq) - + W.divisionPolynomial' (n + 2) * W.divisionPolynomial' n * + (if Even n then W.divisionPolynomial2Sq else 1) := by + erw [divisionPolynomialX, divisionPolynomialZSq_ofNat, ← mul_assoc, divisionPolynomial_ofNat, + add_sub_cancel, divisionPolynomial_ofNat, Int.natAbs_cast] + simp only [Nat.even_add_one, ite_not] + +@[simp] +lemma divisionPolynomialX_zero : W.divisionPolynomialX 0 = 1 := by + rw [divisionPolynomialX, divisionPolynomialZSq_zero, mul_zero, zero_sub, zero_add, + divisionPolynomial_one, one_mul, zero_sub, divisionPolynomial_neg, divisionPolynomial_one, + neg_one_mul, neg_neg, if_pos even_zero.natAbs] + +@[simp] +lemma divisionPolynomialX_one : W.divisionPolynomialX 1 = X := by + erw [divisionPolynomialX_ofNat, one_pow, mul_one, mul_one, mul_zero, zero_mul, sub_zero] + +@[simp] +lemma divisionPolynomialX_two : W.divisionPolynomialX 2 = + X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by + erw [divisionPolynomialX_ofNat, divisionPolynomial'_two, if_neg Nat.not_even_one, + divisionPolynomial2Sq, divisionPolynomial'_three, if_neg Nat.not_even_one] + C_simp + ring1 + +@[simp] +lemma divisionPolynomialX_three : W.divisionPolynomialX 3 = + X * W.divisionPolynomial' 3 ^ 2 - W.divisionPolynomial' 4 * W.divisionPolynomial2Sq := by + erw [divisionPolynomialX_ofNat, divisionPolynomial'_three, mul_one, mul_one, if_pos even_two] + +@[simp] +lemma divisionPolynomialX_four : W.divisionPolynomialX 4 = + X * W.divisionPolynomial' 4 ^ 2 * W.divisionPolynomial2Sq - W.divisionPolynomial' 3 * + (W.divisionPolynomial' 4 * W.divisionPolynomial2Sq ^ 2 - W.divisionPolynomial' 3 ^ 3) := by + erw [divisionPolynomialX_ofNat, if_neg <| by decide, if_neg <| by decide, + show 3 + 2 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial'_two, + if_pos even_zero, if_pos even_zero] + ring1 + +@[simp] +lemma divisionPolynomialX_neg (n : ℤ) : W.divisionPolynomialX (-n) = W.divisionPolynomialX n := by + rw [divisionPolynomialX, divisionPolynomialZSq_neg, neg_add_eq_sub, ← neg_sub n, + divisionPolynomial_neg, ← neg_add', divisionPolynomial_neg, neg_mul_neg, + mul_comm <| W.divisionPolynomial _, Int.natAbs_neg, divisionPolynomialX] + +/-- The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ +@[pp_dot] +noncomputable def divisionPolynomialZ (n : ℤ) : R[X][Y] := + C (W.divisionPolynomial n) * if Even n.natAbs then Y - W.toAffine.negPolynomial else 1 + +@[simp] +lemma divisionPolynomialZ_ofNat (n : ℕ) : W.divisionPolynomialZ n = + C (W.divisionPolynomial' n) * if Even n then Y - W.toAffine.negPolynomial else 1 := by + rw [divisionPolynomialZ, divisionPolynomial_ofNat, Int.natAbs_cast] + +@[simp] +lemma divisionPolynomialZ_zero : W.divisionPolynomialZ 0 = 0 := by + erw [divisionPolynomialZ_ofNat, C_0, zero_mul] + +@[simp] +lemma divisionPolynomialZ_one : W.divisionPolynomialZ 1 = 1 := by + erw [divisionPolynomialZ_ofNat, C_1, mul_one] + +@[simp] +lemma divisionPolynomialZ_two : W.divisionPolynomialZ 2 = Y - W.toAffine.negPolynomial := by + erw [divisionPolynomialZ_ofNat, one_mul, if_pos even_two] + +@[simp] +lemma divisionPolynomialZ_three : W.divisionPolynomialZ 3 = C (W.divisionPolynomial' 3) := by + erw [divisionPolynomialZ_ofNat, mul_one] + +@[simp] +lemma divisionPolynomialZ_four : W.divisionPolynomialZ 4 = + C (W.divisionPolynomial' 4) * (Y - W.toAffine.negPolynomial) := by + erw [divisionPolynomialZ_ofNat, if_pos <| by decide] + +@[simp] +lemma divisionPolynomialZ_neg (n : ℤ) : W.divisionPolynomialZ (-n) = -W.divisionPolynomialZ n := by + rw [divisionPolynomialZ, divisionPolynomial_neg, C_neg, neg_mul (α := R[X][Y]), Int.natAbs_neg, + divisionPolynomialZ] + +end WeierstrassCurve diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean deleted file mode 100644 index ff6a5eae2bffd8..00000000000000 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Torsion.lean +++ /dev/null @@ -1,70 +0,0 @@ -/- -Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Kurniadi Angdinata --/ - -import Mathlib.AlgebraicGeometry.EllipticCurve.Affine -import Mathlib.NumberTheory.EllipticDivisibilitySequence - -/-! -# Torsion points on Weierstrass curves --/ - -open Polynomial PolynomialPolynomial - -universe u - -namespace WeierstrassCurve.Affine - -variable {R : Type u} [CommRing R] (W : Affine R) - -noncomputable def divisionPolynomial (n : ℤ) : R[X][Y] := - EllDivSequence (Y - W.negPolynomial) - (C <| 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) - (C <| 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + - 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) - n - -@[simp] -lemma divisionPolynomial_zero : W.divisionPolynomial 0 = 0 := - EllDivSequence_zero _ _ _ - -@[simp] -lemma divisionPolynomial_one : W.divisionPolynomial 1 = 1 := - EllDivSequence_one _ _ _ - -@[simp] -lemma divisionPolynomial_two : W.divisionPolynomial 2 = Y - W.negPolynomial := - EllDivSequence_two _ _ _ - -@[simp] -lemma divisionPolynomial_three : W.divisionPolynomial 3 = - C (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) := - EllDivSequence_three _ _ _ - -@[simp] -lemma divisionPolynomial_four : W.divisionPolynomial 4 = - C (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + - C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) * - (Y - W.negPolynomial) := - EllDivSequence_four _ _ _ - -@[simp] -lemma divisionPolynomial_odd (m : ℕ) : W.divisionPolynomial (2 * (m + 2) + 1) = - W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 := - EllDivSequence_odd _ _ _ m - -@[simp] -lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) * (Y - W.negPolynomial) = - W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * W.divisionPolynomial (m + 5) - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) * - W.divisionPolynomial (m + 4) ^ 2 := - EllDivSequence_even _ _ _ m - -@[simp] -lemma divisionPolynomial_neg (n : ℕ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := - EllDivSequence_neg _ _ _ n - -end WeierstrassCurve.Affine From 9f4fb43302f3717021122ea043a549fe53978567 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 22 Feb 2024 00:48:34 +0000 Subject: [PATCH 20/47] Fix dependencies --- Mathlib.lean | 3 +- .../EllipticCurve/Group.lean | 45 +- .../EllipticCurve/Jacobian.lean | 1009 ----------------- 3 files changed, 7 insertions(+), 1050 deletions(-) delete mode 100644 Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean diff --git a/Mathlib.lean b/Mathlib.lean index db72a01e600067..4871100644c28c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -517,9 +517,8 @@ import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial import Mathlib.AlgebraicGeometry.EllipticCurve.Group -import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian -import Mathlib.AlgebraicGeometry.EllipticCurve.Torsion import Mathlib.AlgebraicGeometry.EllipticCurve.Weierstrass import Mathlib.AlgebraicGeometry.FunctionField import Mathlib.AlgebraicGeometry.GammaSpecAdjunction diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index a8c3d2199b8ccd..72a86d1fb0c0ac 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian +import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.LinearAlgebra.FreeModule.Norm import Mathlib.RingTheory.ClassGroup @@ -11,8 +11,7 @@ import Mathlib.RingTheory.ClassGroup # Group law on Weierstrass curves This file proves that the nonsingular rational points on a Weierstrass curve forms an abelian group -under the geometric group law defined in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine` and in -`Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian`. +under the geometric group law defined in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`. ## Mathematical background @@ -32,9 +31,6 @@ Injectivity can then be shown by computing the degree of such a norm $N(p + qY)$ ways, which is done in `WeierstrassCurve.Affine.CoordinateRing.degree_norm_smul_basis` and in the auxiliary lemmas in the proof of `WeierstrassCurve.Affine.Point.instAddCommGroupPoint`. -When `W` is given in Jacobian coordinates, `WeierstrassCurve.Jacobian.Point.toAffine_addEquiv` -pulls back the group law on `WeierstrassCurve.Affine.Point` to `WeierstrassCurve.Jacobian.Point`. - ## Main definitions * `WeierstrassCurve.Affine.CoordinateRing`: the coordinate ring `F[W]` of a Weierstrass curve `W`. @@ -42,14 +38,12 @@ pulls back the group law on `WeierstrassCurve.Affine.Point` to `WeierstrassCurve ## Main statements - * `WeierstrassCurve.Affine.CoordinateRing.instIsDomainCoordinateRing`: the coordinate ring of an - affine Weierstrass curve is an integral domain. + * `WeierstrassCurve.Affine.CoordinateRing.instIsDomainCoordinateRing`: the coordinate ring of a + Weierstrass curve is an integral domain. * `WeierstrassCurve.Affine.CoordinateRing.degree_norm_smul_basis`: the degree of the norm of an - element in the coordinate ring of an affine Weierstrass curve in terms of the power basis. + element in the coordinate ring in terms of the power basis. * `WeierstrassCurve.Affine.Point.instAddCommGroupPoint`: the type of nonsingular rational points on - an affine Weierstrass curve forms an abelian group under addition. - * `WeierstrassCurve.Jacobian.Point.instAddCommGroupPoint`: the type of nonsingular rational - points on a Jacobian Weierstrass curve forms an abelian group under addition. + a Weierstrass curve forms an abelian group under addition. ## References @@ -647,33 +641,6 @@ end Point end WeierstrassCurve.Affine -namespace WeierstrassCurve.Jacobian.Point - -/-! ## Weierstrass curves in Jacobian coordinates -/ - -variable {F : Type u} [Field F] {W : Jacobian F} - -lemma add_left_neg (P : W.Point) : -P + P = 0 := - toAffine_addEquiv.injective <| by - rcases P with @⟨⟨_⟩, _⟩ - simp only [map_add, toAffine_addEquiv_apply, toAffine_lift_neg, Affine.Point.add_left_neg, - toAffine_lift_zero] - -lemma add_comm (P Q : W.Point) : P + Q = Q + P := - toAffine_addEquiv.injective <| by simp_rw [map_add, Affine.Point.add_comm] - -lemma add_assoc (P Q R : W.Point) : P + Q + R = P + (Q + R) := - toAffine_addEquiv.injective <| by simp only [map_add, Affine.Point.add_assoc] - -noncomputable instance instAddCommGroupPoint : AddCommGroup W.Point where - zero_add := zero_add - add_zero := add_zero - add_left_neg := add_left_neg - add_comm := add_comm - add_assoc := add_assoc - -end WeierstrassCurve.Jacobian.Point - namespace EllipticCurve.Affine /-! ## Elliptic curves in affine coordinates -/ diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean deleted file mode 100644 index 5badff32ab8054..00000000000000 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ /dev/null @@ -1,1009 +0,0 @@ -/- -Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Kurniadi Angdinata --/ -import Mathlib.AlgebraicGeometry.EllipticCurve.Affine -import Mathlib.Data.MvPolynomial.CommRing - -/-! -# Jacobian coordinates for Weierstrass curves - -This file defines the type of points on a Weierstrass curve as a tuple, consisting of an equivalence -class of triples up to scaling by weights, satisfying a Weierstrass equation with a nonsingular -condition. This file also defines the negation and addition operations of the group law for this -type, and proves that they respect the Weierstrass equation and the nonsingular condition. The fact -that they form an abelian group is proven in `Mathlib.AlgebraicGeometry.EllipticCurve.Group`. - -## Mathematical background - -Let `W` be a Weierstrass curve over a field `F`. A point on the weighted projective plane with -weights $(2, 3, 1)$ is an equivalence class of triples $[x:y:z]$ with coordinates in `F` such that -$(x, y, z) \sim (x', y', z')$ precisely if there is some unit $u$ of `F` such that -$(x, y, z) = (u^2x', u^3y', uz')$, with an extra condition that $(x, y, z) \ne (0, 0, 0)$. -A rational point is a point on the $(2, 3, 1)$-projective plane satisfying a $(2, 3, 1)$-homogeneous -Weierstrass equation $Y^2 + a_1XYZ + a_3YZ^3 = X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6$, and being -nonsingular means the partial derivatives $W_X(X, Y, Z)$, $W_Y(X, Y, Z)$, and $W_Z(X, Y, Z)$ do not -vanish simultaneously. Note that the vanishing of the Weierstrass equation and its partial -derivatives are independent of the representative for $[x:y:z]$, and the nonsingularity condition -already implies that $(x, y, z) \ne (0, 0, 0)$, so a nonsingular rational point on `W` can simply be -given by a tuple consisting of $[x:y:z]$ and the nonsingular condition on any representative. -In cryptography, as well as in this file, this is often called the Jacobian coordinates of `W`. - -As in `Mathlib.AlgebraicGeometry.EllipticCurve.Affine`, the set of nonsingular rational points forms -an abelian group under the same secant-and-tangent process, but the polynomials involved are -$(2, 3, 1)$-homogeneous, and any instances of division become multiplication in the $Z$-coordinate. -Note that most computational proofs follow from their analogous proofs for affine coordinates. - -## Main definitions - - * `WeierstrassCurve.Jacobian.PointClass`: the equivalence class of a point representative. - * `WeierstrassCurve.Jacobian.toAffine`: the Weierstrass curve in affine coordinates. - * `WeierstrassCurve.Jacobian.nonsingular`: the nonsingular condition on a point representative. - * `WeierstrassCurve.Jacobian.nonsingular_lift`: the nonsingular condition on a point class. - * `WeierstrassCurve.Jacobian.neg`: the negation operation on a point representative. - * `WeierstrassCurve.Jacobian.neg_map`: the negation operation on a point class. - * `WeierstrassCurve.Jacobian.add`: the addition operation on a point representative. - * `WeierstrassCurve.Jacobian.add_map`: the addition operation on a point class. - * `WeierstrassCurve.Jacobian.Point`: a nonsingular rational point. - * `WeierstrassCurve.Jacobian.Point.neg`: the negation operation on a nonsingular rational point. - * `WeierstrassCurve.Jacobian.Point.add`: the addition operation on a nonsingular rational point. - * `WeierstrassCurve.Jacobian.Point.toAffine_addEquiv`: the equivalence between the nonsingular - rational points on a Jacobian Weierstrass curve with those on an affine Weierstrass curve. - -## Main statements - - * `WeierstrassCurve.Jacobian.nonsingular_neg`: negation preserves the nonsingular condition. - * `WeierstrassCurve.Jacobian.nonsingular_add`: addition preserves the nonsingular condition. - -## Implementation notes - -A point representative is implemented as a term `P` of type `Fin 3 → R`, which allows for the vector -notation `![x, y, z]`. However, `P` is not definitionally equivalent to the expanded vector -`![P x, P y, P z]`, so the auxiliary lemma `fin3_def` can be used to convert between the two forms. -The equivalence of two point representatives `P` and `Q` is implemented as an equivalence of orbits -of the action of `Rˣ`, or equivalently that there is some unit `u` of `R` such that `P = u • Q`. -However, `u • Q` is again not definitionally equal to `![u² * Q x, u³ * Q y, u * Q z]`, so the -auxiliary lemmas `smul_fin3` and `smul_fin3_ext` can be used to convert between the two forms. - -## References - -[J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] - -## Tags - -elliptic curve, rational point, Jacobian coordinates --/ - -local notation "x" => 0 - -local notation "y" => 1 - -local notation "z" => 2 - -local macro "matrix_simp" : tactic => - `(tactic| simp only [Matrix.head_cons, Matrix.tail_cons, Matrix.smul_empty, Matrix.smul_cons, - Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.cons_val_two]) - -universe u - -lemma fin3_def {R : Type u} (P : Fin 3 → R) : P = ![P x, P y, P z] := by - ext n; fin_cases n <;> rfl - -private instance {R : Type u} [CommRing R] : SMul Rˣ <| Fin 3 → R := - ⟨fun u P => ![u ^ 2 * P x, u ^ 3 * P y, u * P z]⟩ - -lemma smul_fin3 {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : - u • P = ![u ^ 2 * P x, u ^ 3 * P y, u * P z] := - rfl - -lemma smul_fin3_ext {R : Type u} [CommRing R] (P : Fin 3 → R) (u : Rˣ) : - (u • P) 0 = u ^ 2 * P x ∧ (u • P) 1 = u ^ 3 * P y ∧ (u • P) 2 = u * P z := - ⟨rfl, rfl, rfl⟩ - -private instance {R : Type u} [CommRing R] : MulAction Rˣ <| Fin 3 → R where - one_smul := fun _ => by - simp only [smul_fin3, Units.val_one, one_pow, one_mul, ← fin3_def] - mul_smul := fun u v P => by - simp only [smul_fin3, Units.val_mul, mul_pow, mul_assoc] - matrix_simp - -/-! ## Weierstrass curves -/ - -/-- An abbreviation for a Weierstrass curve in Jacobian coordinates. -/ -abbrev WeierstrassCurve.Jacobian := - WeierstrassCurve - -namespace WeierstrassCurve.Jacobian - -open MvPolynomial - -local macro "eval_simp" : tactic => - `(tactic| simp only [eval_C, eval_X, eval_add, eval_sub, eval_mul, eval_pow]) - -variable (R : Type u) [CommRing R] - -/-- The equivalence setoid for a point representative. -/ -def PointSetoid : Setoid <| Fin 3 → R := - MulAction.orbitRel Rˣ <| Fin 3 → R - -attribute [local instance] PointSetoid - -/-- The equivalence class of a point representative. -/ -abbrev PointClass : Type u := - MulAction.orbitRel.Quotient Rˣ <| Fin 3 → R - -variable {R} (W : Jacobian R) - -/-- The coercion to a Weierstrass curve in affine coordinates. -/ -@[pp_dot] -abbrev toAffine : Affine R := - W - -section Equation - -/-! ### Equations and nonsingularity -/ - -/-- The polynomial $W(X, Y, Z) := Y^2 + a_1XYZ + a_3YZ^3 - (X^3 + a_2X^2Z^2 + a_4XZ^4 + a_6Z^6)$ -associated to a Weierstrass curve `W` over `R`. This is represented as a term of type -`MvPolynomial (Fin 3) R`, where `X 0`, `X 1`, and `X 2` represent $X$, $Y$, and $Z$ respectively. -/ -@[pp_dot] -noncomputable def polynomial : MvPolynomial (Fin 3) R := - X 1 ^ 2 + C W.a₁ * X 0 * X 1 * X 2 + C W.a₃ * X 1 * X 2 ^ 3 - - (X 0 ^ 3 + C W.a₂ * X 0 ^ 2 * X 2 ^ 2 + C W.a₄ * X 0 * X 2 ^ 4 + C W.a₆ * X 2 ^ 6) - -lemma eval_polynomial (P : Fin 3 → R) : eval P W.polynomial = - P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 - - (P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6) := by - rw [polynomial] - eval_simp - -/-- The proposition that a point representative $(x, y, z)$ lies in `W`. -In other words, $W(x, y, z) = 0$. -/ -@[pp_dot] -def equation (P : Fin 3 → R) : Prop := - eval P W.polynomial = 0 - -lemma equation_iff (P : Fin 3 → R) : W.equation P ↔ - P y ^ 2 + W.a₁ * P x * P y * P z + W.a₃ * P y * P z ^ 3 - = P x ^ 3 + W.a₂ * P x ^ 2 * P z ^ 2 + W.a₄ * P x * P z ^ 4 + W.a₆ * P z ^ 6 := by - rw [equation, eval_polynomial, sub_eq_zero] - -lemma equation_zero : W.equation ![1, 1, 0] := - (W.equation_iff ![1, 1, 0]).mpr <| by matrix_simp; ring1 - -lemma equation_zero' (Y : R) : W.equation ![Y ^ 2, Y ^ 3, 0] := - (W.equation_iff ![Y ^ 2, Y ^ 3, 0]).mpr <| by matrix_simp; ring1 - -lemma equation_some (X Y : R) : W.equation ![X, Y, 1] ↔ W.toAffine.equation X Y := by - rw [equation_iff, W.toAffine.equation_iff] - congr! 1 <;> matrix_simp <;> ring1 - -lemma equation_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.equation (u • P) ↔ W.equation P := - have (u : Rˣ) {P : Fin 3 → R} (h : W.equation P) : W.equation <| u • P := by - rw [equation_iff] at h ⊢ - linear_combination (norm := (simp only [smul_fin3_ext]; ring1)) (u : R) ^ 6 * h - ⟨fun h => by convert this u⁻¹ h; rw [inv_smul_smul], this u⟩ - -/-- The partial derivative $W_X(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $X$. - -TODO: define this in terms of `MvPolynomial.pderiv`. -/ -@[pp_dot] -noncomputable def polynomialX : MvPolynomial (Fin 3) R := - C W.a₁ * X 1 * X 2 - (C 3 * X 0 ^ 2 + C (2 * W.a₂) * X 0 * X 2 ^ 2 + C W.a₄ * X 2 ^ 4) - -lemma eval_polynomialX (P : Fin 3 → R) : eval P W.polynomialX = - W.a₁ * P y * P z - (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4) := by - rw [polynomialX] - eval_simp - -/-- The partial derivative $W_Y(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Y$. - -TODO: define this in terms of `MvPolynomial.pderiv`. -/ -@[pp_dot] -noncomputable def polynomialY : MvPolynomial (Fin 3) R := - C 2 * X 1 + C W.a₁ * X 0 * X 2 + C W.a₃ * X 2 ^ 3 - -lemma eval_polynomialY (P : Fin 3 → R) : - eval P W.polynomialY = 2 * P y + W.a₁ * P x * P z + W.a₃ * P z ^ 3 := by - rw [polynomialY] - eval_simp - -/-- The partial derivative $W_Z(X, Y, Z)$ of $W(X, Y, Z)$ with respect to $Z$. - -TODO: define this in terms of `MvPolynomial.pderiv`. -/ -@[pp_dot] -noncomputable def polynomialZ : MvPolynomial (Fin 3) R := - C W.a₁ * X 0 * X 1 + C (3 * W.a₃) * X 1 * X 2 ^ 2 - - (C (2 * W.a₂) * X 0 ^ 2 * X 2 + C (4 * W.a₄) * X 0 * X 2 ^ 3 + C (6 * W.a₆) * X 2 ^ 5) - -lemma eval_polynomialZ (P : Fin 3 → R) : eval P W.polynomialZ = - W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 - - (2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by - rw [polynomialZ] - eval_simp - -/-- The proposition that a point representative $(x, y, z)$ in `W` is nonsingular. -In other words, either $W_X(x, y, z) \ne 0$, $W_Y(x, y, z) \ne 0$, or $W_Z(x, y, z) \ne 0$. -/ -@[pp_dot] -def nonsingular (P : Fin 3 → R) : Prop := - W.equation P ∧ (eval P W.polynomialX ≠ 0 ∨ eval P W.polynomialY ≠ 0 ∨ eval P W.polynomialZ ≠ 0) - -lemma nonsingular_iff (P : Fin 3 → R) : W.nonsingular P ↔ W.equation P ∧ - (W.a₁ * P y * P z ≠ 3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 ∨ - P y ≠ -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 ∨ - W.a₁ * P x * P y + 3 * W.a₃ * P y * P z ^ 2 - ≠ 2 * W.a₂ * P x ^ 2 * P z + 4 * W.a₄ * P x * P z ^ 3 + 6 * W.a₆ * P z ^ 5) := by - rw [nonsingular, eval_polynomialX, eval_polynomialY, eval_polynomialZ, sub_ne_zero, sub_ne_zero, - ← sub_ne_zero (a := P y)] - congr! 4 - ring1 - -lemma nonsingular_zero [Nontrivial R] : W.nonsingular ![1, 1, 0] := - (W.nonsingular_iff ![1, 1, 0]).mpr ⟨W.equation_zero, - by simp; by_contra! h; exact one_ne_zero <| by linear_combination -h.1 - h.2.1⟩ - -lemma nonsingular_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : - W.nonsingular ![Y ^ 2, Y ^ 3, 0] := - (W.nonsingular_iff ![Y ^ 2, Y ^ 3, 0]).mpr ⟨W.equation_zero' Y, - by simp [hy]; by_contra! h; exact pow_ne_zero 3 hy <| by linear_combination Y ^ 3 * h.1 - h.2.1⟩ - -lemma nonsingular_some (X Y : R) : W.nonsingular ![X, Y, 1] ↔ W.toAffine.nonsingular X Y := by - rw [nonsingular_iff] - matrix_simp - simp only [W.toAffine.nonsingular_iff, equation_some, and_congr_right_iff, - W.toAffine.equation_iff, ← not_and_or, not_iff_not, one_pow, mul_one, Iff.comm, iff_self_and] - intro h hX hY - linear_combination (norm := ring1) 6 * h - 2 * X * hX - 3 * Y * hY - -lemma nonsingular_smul_iff (P : Fin 3 → R) (u : Rˣ) : W.nonsingular (u • P) ↔ W.nonsingular P := - have (u : Rˣ) {P : Fin 3 → R} (h : W.nonsingular <| u • P) : W.nonsingular P := by - rcases (W.nonsingular_iff _).mp h with ⟨h, h'⟩ - refine (W.nonsingular_iff P).mpr ⟨(W.equation_smul_iff P u).mp h, ?_⟩ - contrapose! h' - simp only [smul_fin3_ext] - exact ⟨by linear_combination (norm := ring1) (u : R) ^ 4 * h'.left, - by linear_combination (norm := ring1) (u : R) ^ 3 * h'.right.left, - by linear_combination (norm := ring1) (u : R) ^ 5 * h'.right.right⟩ - ⟨this u, fun h => this u⁻¹ <| by rwa [inv_smul_smul]⟩ - -lemma nonsingular_of_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.nonsingular P ↔ W.nonsingular Q := by - rcases h with ⟨u, rfl⟩ - exact W.nonsingular_smul_iff Q u - -/-- The proposition that a point class on `W` is nonsingular. If `P` is a point representative, -then `W.nonsingular_lift ⟦P⟧` is definitionally equivalent to `W.nonsingular P`. -/ -@[pp_dot] -def nonsingular_lift (P : PointClass R) : Prop := - P.lift W.nonsingular fun _ _ => propext ∘ W.nonsingular_of_equiv - -@[simp] -lemma nonsingular_lift_eq (P : Fin 3 → R) : W.nonsingular_lift ⟦P⟧ = W.nonsingular P := - rfl - -lemma nonsingular_lift_zero [Nontrivial R] : W.nonsingular_lift ⟦![1, 1, 0]⟧ := - W.nonsingular_zero - -lemma nonsingular_lift_zero' [NoZeroDivisors R] {Y : R} (hy : Y ≠ 0) : - W.nonsingular_lift ⟦![Y ^ 2, Y ^ 3, 0]⟧ := - W.nonsingular_zero' hy - -lemma nonsingular_lift_some (X Y : R) : - W.nonsingular_lift ⟦![X, Y, 1]⟧ ↔ W.toAffine.nonsingular X Y := - W.nonsingular_some X Y - -variable {F : Type u} [Field F] {W : Jacobian F} - -lemma equiv_of_Zeq0 {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z = 0) - (hQz : Q z = 0) : P ≈ Q := by - rw [fin3_def P, hPz] at hP ⊢ - rw [fin3_def Q, hQz] at hQ ⊢ - simp [nonsingular_iff, equation_iff] at hP hQ - have hPx : P x ≠ 0 := fun h => by simp [h] at hP; simp [hP] at hP - have hQx : Q x ≠ 0 := fun h => by simp [h] at hQ; simp [hQ] at hQ - have hPy : P y ≠ 0 := fun h => by simp [h] at hP; exact hPx <| pow_eq_zero hP.left.symm - have hQy : Q y ≠ 0 := fun h => by simp [h] at hQ; exact hQx <| pow_eq_zero hQ.left.symm - use Units.mk0 _ <| mul_ne_zero (div_ne_zero hPy hPx) (div_ne_zero hQx hQy) - simp [smul_fin3, mul_pow, div_pow] - congr! 2 - · field_simp [hP.left, hQ.left] - ring1 - · field_simp [← hP.left, ← hQ.left] - ring1 - -lemma equiv_zero_of_Zeq0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z = 0) : P ≈ ![1, 1, 0] := - equiv_of_Zeq0 h W.nonsingular_zero hPz rfl - -lemma equiv_some_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : P ≈ ![P x / P z ^ 2, P y / P z ^ 3, 1] := - ⟨Units.mk0 _ hPz, by simp [smul_fin3, ← fin3_def P, mul_div_cancel' _ <| pow_ne_zero _ hPz]⟩ - -lemma nonsingular_iff_affine_of_Zne0 {P : Fin 3 → F} (hPz : P z ≠ 0) : - W.nonsingular P ↔ W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := - (W.nonsingular_of_equiv <| equiv_some_of_Zne0 hPz).trans <| W.nonsingular_some .. - -lemma nonsingular_of_affine_of_Zne0 {P : Fin 3 → F} - (h : W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3)) (hPz : P z ≠ 0) : - W.nonsingular P := - (nonsingular_iff_affine_of_Zne0 hPz).mpr h - -lemma nonsingular_affine_of_Zne0 {P : Fin 3 → F} (h : W.nonsingular P) (hPz : P z ≠ 0) : - W.toAffine.nonsingular (P x / P z ^ 2) (P y / P z ^ 3) := - (nonsingular_iff_affine_of_Zne0 hPz).mp h - -end Equation - -section Polynomial - -/-! ### Group operation polynomials -/ - -/-- The $Y$-coordinate of the negation of a point representative. -/ -@[pp_dot] -def negY (P : Fin 3 → R) : R := - -P y - W.a₁ * P x * P z - W.a₃ * P z ^ 3 - -lemma negY_smul (P : Fin 3 → R) (u : Rˣ) : W.negY (u • P) = u ^ 3 * W.negY P := by - simp only [negY, smul_fin3_ext] - ring1 - -/-- The $X$-coordinate of the addition of two point representatives, where their $Z$-coordinates are -non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ -@[pp_dot] -def addX_of_Xne (P Q : Fin 3 → R) : R := - P x * Q x ^ 2 * P z ^ 2 - 2 * P y * Q y * P z * Q z + P x ^ 2 * Q x * Q z ^ 2 - - W.a₁ * P x * Q y * P z ^ 2 * Q z - W.a₁ * P y * Q x * P z * Q z ^ 2 - + 2 * W.a₂ * P x * Q x * P z ^ 2 * Q z ^ 2 - W.a₃ * Q y * P z ^ 4 * Q z - - W.a₃ * P y * P z * Q z ^ 4 + W.a₄ * Q x * P z ^ 4 * Q z ^ 2 + W.a₄ * P x * P z ^ 2 * Q z ^ 4 - + 2 * W.a₆ * P z ^ 4 * Q z ^ 4 - -lemma addX_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addX_of_Xne (u • P) (v • Q) = (u : R) ^ 4 * (v : R) ^ 4 * W.addX_of_Xne P Q := by - simp only [addX_of_Xne, smul_fin3_ext] - ring1 - -/-- The $X$-coordinate of the doubling of a point representative, where its $Z$-coordinate is -non-zero and its $Y$-coordinate is distinct from that of its negation. -/ -@[pp_dot] -def addX_of_Yne (P : Fin 3 → R) : R := - (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) ^ 2 - + W.a₁ * (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) - * (P y * P z - W.negY P * P z) - - (W.a₂ * P z ^ 2 + 2 * P x) * (P y - W.negY P) ^ 2 - -lemma addX_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : - W.addX_of_Yne (u • P) = (u : R) ^ 8 * W.addX_of_Yne P := by - simp only [addX_of_Yne, negY_smul, smul_fin3_ext] - ring1 - -/-- The $Y$-coordinate of the addition of two point representatives, before applying the final -negation that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where their $Z$-coordinates are non-zero and their -$X$-coordinates divided by $Z$-coordinates squared are distinct. -/ -@[pp_dot] -def addY'_of_Xne (P Q : Fin 3 → R) : R := - -P y * Q x ^ 3 * P z ^ 3 + 2 * P y * Q y ^ 2 * P z ^ 3 - 3 * P x ^ 2 * Q x * Q y * P z ^ 2 * Q z - + 3 * P x * P y * Q x ^ 2 * P z * Q z ^ 2 + P x ^ 3 * Q y * Q z ^ 3 - - 2 * P y ^ 2 * Q y * Q z ^ 3 + W.a₁ * P x * Q y ^ 2 * P z ^ 4 - + W.a₁ * P y * Q x * Q y * P z ^ 3 * Q z - W.a₁ * P x * P y * Q y * P z * Q z ^ 3 - - W.a₁ * P y ^ 2 * Q x * Q z ^ 4 - 2 * W.a₂ * P x * Q x * Q y * P z ^ 4 * Q z - + 2 * W.a₂ * P x * P y * Q x * P z * Q z ^ 4 + W.a₃ * Q y ^ 2 * P z ^ 6 - - W.a₃ * P y ^ 2 * Q z ^ 6 - W.a₄ * Q x * Q y * P z ^ 6 * Q z - - W.a₄ * P x * Q y * P z ^ 4 * Q z ^ 3 + W.a₄ * P y * Q x * P z ^ 3 * Q z ^ 4 - + W.a₄ * P x * P y * P z * Q z ^ 6 - 2 * W.a₆ * Q y * P z ^ 6 * Q z ^ 3 - + 2 * W.a₆ * P y * P z ^ 3 * Q z ^ 6 - -lemma addY'_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addY'_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addY'_of_Xne P Q := by - simp only [addY'_of_Xne, smul_fin3_ext] - ring1 - -/-- The $Y$-coordinate of the doubling of a point representative, before applying the final negation -that maps $Y$ to $-Y - a_1XZ - a_3Z^3$, where its $Z$-coordinate is non-zero and its $Y$-coordinate -is distinct from that of its negation. -/ -@[pp_dot] -def addY'_of_Yne (P : Fin 3 → R) : R := - (3 * P x ^ 2 + 2 * W.a₂ * P x * P z ^ 2 + W.a₄ * P z ^ 4 - W.a₁ * P y * P z) - * (W.addX_of_Yne P - P x * (P y - W.negY P) ^ 2) - + P y * (P y - W.negY P) ^ 3 - -lemma addY'_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : - W.addY'_of_Yne (u • P) = (u : R) ^ 12 * W.addY'_of_Yne P := by - simp only [addY'_of_Yne, addX_of_Yne_smul, negY_smul, smul_fin3_ext] - ring1 - -/-- The $Z$-coordinate of the addition of two point representatives, where their $Z$-coordinates are -non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ -def addZ_of_Xne (P Q : Fin 3 → R) : R := - P x * Q z ^ 2 - P z ^ 2 * Q x - -lemma addZ_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - addZ_of_Xne (u • P) (v • Q) = (u : R) ^ 2 * (v : R) ^ 2 * addZ_of_Xne P Q := by - simp only [addZ_of_Xne, smul_fin3_ext] - ring1 - -/-- The $Z$-coordinate of the doubling of a point representative, where its $Z$-coordinate is -non-zero and its $Y$-coordinate is distinct from that of its negation. -/ -@[pp_dot] -def addZ_of_Yne (P : Fin 3 → R) : R := - P z * (P y - W.negY P) - -lemma addZ_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : - W.addZ_of_Yne (u • P) = (u : R) ^ 4 * W.addZ_of_Yne P := by - simp only [addZ_of_Yne, negY_smul, smul_fin3_ext] - ring1 - -/-- The $Y$-coordinate of the addition of two point representatives, where their $Z$-coordinates are -non-zero and their $X$-coordinates divided by $Z$-coordinates squared are distinct. -/ -@[pp_dot] -def addY_of_Xne (P Q : Fin 3 → R) : R := - W.negY ![W.addX_of_Xne P Q, W.addY'_of_Xne P Q, addZ_of_Xne P Q] - -lemma addY_of_Xne_smul (P Q : Fin 3 → R) (u v : Rˣ) : - W.addY_of_Xne (u • P) (v • Q) = (u : R) ^ 6 * (v : R) ^ 6 * W.addY_of_Xne P Q := by - simp only [addY_of_Xne, negY, addX_of_Xne_smul, addY'_of_Xne_smul, addZ_of_Xne_smul] - matrix_simp - ring1 - -/-- The $Y$-coordinate of the doubling of a point representative, where its $Z$-coordinate is -non-zero and its $Y$-coordinate is distinct from that of its negation. -/ -@[pp_dot] -def addY_of_Yne (P : Fin 3 → R) : R := - W.negY ![W.addX_of_Yne P, W.addY'_of_Yne P, W.addZ_of_Yne P] - -lemma addY_of_Yne_smul (P : Fin 3 → R) (u : Rˣ) : - W.addY_of_Yne (u • P) = (u : R) ^ 12 * W.addY_of_Yne P := by - simp only [addY_of_Yne, negY, addX_of_Yne_smul, addY'_of_Yne_smul, addZ_of_Yne_smul] - matrix_simp - ring1 - -variable {F : Type u} [Field F] {W : Jacobian F} - -lemma negY_divZ {P : Fin 3 → F} (hPz : P z ≠ 0) : - W.negY P / P z ^ 3 = W.toAffine.negY (P x / P z ^ 2) (P y / P z ^ 3) := by - field_simp [negY, Affine.negY] - ring1 - -lemma Yne_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) (hPz : P z ≠ 0) - (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : - P y ≠ W.negY P := by - simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] - at hx hy - have hx' : P x * (P z / P z ^ 3) = Q x * (Q z / Q z ^ 3) := by - simp_rw [pow_succ _ 2, div_mul_right _ hPz, div_mul_right _ hQz, mul_one_div, hx] - have hy' : P y / P z ^ 3 = Q y / Q z ^ 3 := - Affine.Yeq_of_Yne (nonsingular_affine_of_Zne0 hP hPz).left - (nonsingular_affine_of_Zne0 hQ hQz).left hx <| (negY_divZ hQz).symm ▸ hy - simp_rw [negY, sub_div, neg_div, mul_div_assoc, mul_assoc, ← hy', ← hx', - div_self <| pow_ne_zero 3 hQz, ← div_self <| pow_ne_zero 3 hPz, ← mul_assoc, ← mul_div_assoc, - ← neg_div, ← sub_div, div_left_inj' <| pow_ne_zero 3 hPz] at hy - exact hy - -lemma addX_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - W.addX_of_Xne P Q / addZ_of_Xne P Q ^ 2 = W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - convert_to - ((P y * Q z ^ 3 - P z ^ 3 * Q y) ^ 2 - + W.a₁ * P z * Q z * (P y * Q z ^ 3 - P z ^ 3 * Q y) * (P x * Q z ^ 2 - P z ^ 2 * Q x) - - (W.a₂ * P z ^ 2 * Q z ^ 2 + P x * Q z ^ 2 + P z ^ 2 * Q x) - * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) - / (P z ^ 2 * Q z ^ 2) / addZ_of_Xne P Q ^ 2 = _ using 2 - · rw [nonsingular_iff, equation_iff] at hP hQ - rw [addX_of_Xne, eq_div_iff_mul_eq <| mul_ne_zero (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)] - linear_combination (norm := ring1) -Q z ^ 6 * hP.left - P z ^ 6 * hQ.left - rw [Affine.slope_of_Xne <| - by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] - field_simp [sub_ne_zero_of_ne hx, addZ_of_Xne] - ring1 - -lemma addX_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) - (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : - W.addX_of_Yne P / W.addZ_of_Yne P ^ 2 = W.toAffine.addX (P x / P z ^ 2) (Q x / Q z ^ 2) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy - simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] - at hx hy - rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← hx, ← negY_divZ hPz] - field_simp [addX_of_Yne, addX_of_Yne, addZ_of_Yne] - ring1 - -lemma addY'_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - W.addY'_of_Xne P Q / addZ_of_Xne P Q ^ 3 = - W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - convert_to - ((P y * Q z ^ 3 - P z ^ 3 * Q y) * (P z ^ 2 * Q z ^ 2 * W.addX_of_Xne P Q - - P x * Q z ^ 2 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 2) - + P y * Q z ^ 3 * (P x * Q z ^ 2 - P z ^ 2 * Q x) ^ 3) - / (P z ^ 3 * Q z ^ 3) / addZ_of_Xne P Q ^ 3 = _ using 2 - · rw [addY'_of_Xne, addX_of_Xne, - eq_div_iff_mul_eq <| mul_ne_zero (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)] - ring1 - rw [Affine.addY', ← addX_div_addZ_of_Xne hP hQ hPz hQz hx, Affine.slope_of_Xne <| - by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| Q x]] - field_simp [sub_ne_zero_of_ne hx, addX_of_Xne, addZ_of_Xne] - ring1 - -lemma addY'_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) - (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY'_of_Yne P / W.addZ_of_Yne P ^ 3 = - W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - have := sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy - rw [Affine.addY', ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy] - simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] - at hx hy - rw [Affine.slope_of_Yne hx <| (negY_divZ hQz).symm ▸ hy, ← negY_divZ hPz] - field_simp [addY'_of_Yne, addX_of_Yne, addZ_of_Yne] - ring1 - -lemma addZ_ne_zero_of_Xne {P Q : Fin 3 → F} (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - addZ_of_Xne P Q ≠ 0 := - sub_ne_zero_of_ne hx - -lemma addZ_ne_zero_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) - (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addZ_of_Yne P ≠ 0 := - mul_ne_zero hPz <| sub_ne_zero_of_ne <| Yne_of_Yne hP hQ hPz hQz hx hy - -lemma addY_div_addZ_of_Xne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - W.addY_of_Xne P Q / addZ_of_Xne P Q ^ 3 = - W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - simpa only [Affine.addY, ← addX_div_addZ_of_Xne hP hQ hPz hQz hx, - ← addY'_div_addZ_of_Xne hP hQ hPz hQz hx] using negY_divZ <| addZ_ne_zero_of_Xne hx - -lemma addY_div_addZ_of_Yne {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) - (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) - (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : W.addY_of_Yne P / W.addZ_of_Yne P ^ 3 = - W.toAffine.addY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by - rw [Affine.addY, ← addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, - ← addY'_div_addZ_of_Yne hP hQ hPz hQz hx hy] - exact negY_divZ <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy - -end Polynomial - -section Representative - -/-! ### Group operations on point representatives -/ - -/-- The negation of a point representative. -/ -@[pp_dot] -def neg (P : Fin 3 → R) : Fin 3 → R := - ![P x, W.negY P, P z] - -@[simp] -lemma neg_zero : W.neg ![1, 1, 0] = ![1, -1, 0] := by - erw [neg, negY, mul_zero, zero_pow three_pos, mul_zero, sub_zero, sub_zero] - rfl - -@[simp] -lemma neg_some (X Y : R) : W.neg ![X, Y, 1] = ![X, -Y - W.a₁ * X - W.a₃, 1] := by - erw [neg, negY, mul_one, one_pow, mul_one] - rfl - -lemma neg_smul_equiv (P : Fin 3 → R) (u : Rˣ) : W.neg (u • P) ≈ W.neg P := - ⟨u, by simp_rw [neg, negY_smul, smul_fin3]; rfl⟩ - -lemma neg_equiv {P Q : Fin 3 → R} (h : P ≈ Q) : W.neg P ≈ W.neg Q := by - rcases h with ⟨u, rfl⟩ - exact W.neg_smul_equiv Q u - -/-- The negation of a point class. If `P` is a point representative, -then `W.neg_map ⟦P⟧` is definitionally equivalent to `W.neg P`. -/ -@[pp_dot] -def neg_map (P : PointClass R) : PointClass R := - P.map W.neg fun _ _ => W.neg_equiv - -lemma neg_map_eq {P : Fin 3 → R} : W.neg_map ⟦P⟧ = ⟦W.neg P⟧ := - rfl - -@[simp] -lemma neg_map_zero : W.neg_map ⟦![1, 1, 0]⟧ = ⟦![1, 1, 0]⟧ := by - simpa only [neg_map_eq, neg_zero, Quotient.eq] using ⟨-1, by norm_num [smul_fin3]⟩ - -@[simp] -lemma neg_map_some (X Y : R) : W.neg_map ⟦![X, Y, 1]⟧ = ⟦![X, -Y - W.a₁ * X - W.a₃, 1]⟧ := by - rw [neg_map_eq, neg_some] - -open scoped Classical - -/-- The addition of two point representatives. -/ -@[pp_dot] -noncomputable def add (P Q : Fin 3 → R) : Fin 3 → R := - if P z = 0 then Q else if Q z = 0 then P else if P x * Q z ^ 2 = P z ^ 2 * Q x then - if P y * Q z ^ 3 = P z ^ 3 * W.negY Q then ![1, 1, 0] else - ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] - else ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] - -@[simp] -lemma add_of_Zeq0_left {P Q : Fin 3 → R} (hPz : P z = 0) : W.add P Q = Q := - if_pos hPz - -lemma add_zero_left (P : Fin 3 → R) : W.add ![1, 1, 0] P = P := - W.add_of_Zeq0_left rfl - -@[simp] -lemma add_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : W.add P Q = P := by - rw [add, if_neg hPz, if_pos hQz] - -lemma add_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add P ![1, 1, 0] = P := - W.add_of_Zeq0_right hPz rfl - -@[simp] -lemma add_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : - W.add P Q = ![1, 1, 0] := by - rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_pos hy] - -@[simp] -lemma add_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : - W.add P Q = ![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P] := by - rw [add, if_neg hPz, if_neg hQz, if_pos hx, if_neg hy] - -@[simp] -lemma add_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - W.add P Q = ![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q] := by - rw [add, if_neg hPz, if_neg hQz, if_neg hx] - -variable [IsDomain R] - -lemma add_smul_equiv (P Q : Fin 3 → R) (u v : Rˣ) : W.add (u • P) (v • Q) ≈ W.add P Q := by - have huv (n : ℕ) : (u ^ n * v ^ n : R) ≠ 0 := - mul_ne_zero (pow_ne_zero n u.ne_zero) (pow_ne_zero n v.ne_zero) - by_cases hPz : P z = 0 - · exact ⟨v, by rw [W.add_of_Zeq0_left hPz, - W.add_of_Zeq0_left <| by simp only [smul_fin3_ext, hPz, mul_zero]]⟩ - · have huz : u * P z ≠ 0 := mul_ne_zero u.ne_zero hPz - by_cases hQz : Q z = 0 - · rw [W.add_of_Zeq0_right hPz hQz, - W.add_of_Zeq0_right huz <| by simp only [smul_fin3_ext, hQz, mul_zero]] - exact ⟨u, rfl⟩ - · have hvz : v * Q z ≠ 0 := mul_ne_zero v.ne_zero hQz - by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x - · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q - · rw [W.add_of_Yeq huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by - simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm, hy], - W.add_of_Yeq hPz hQz hx hy] - · rw [W.add_of_Yne huz hvz (by simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm, hx]) <| by - simp_rw [smul_fin3_ext, mul_pow, negY_smul, mul_mul_mul_comm] - exact hy ∘ mul_left_cancel₀ (huv 3), - addX_of_Yne_smul, addY_of_Yne_smul, addZ_of_Yne_smul, W.add_of_Yne hPz hQz hx hy] - exact ⟨u ^ 4, by simp only [smul_fin3, ← Units.val_pow_eq_pow_val, ← pow_mul]; rfl⟩ - · rw [W.add_of_Xne huz hvz <| by - simp_rw [smul_fin3_ext, mul_pow, mul_mul_mul_comm]; exact hx ∘ mul_left_cancel₀ (huv 2), - addX_of_Xne_smul, addY_of_Xne_smul, addZ_of_Xne_smul, W.add_of_Xne hPz hQz hx] - exact ⟨u ^ 2 * v ^ 2, - by simp_rw [smul_fin3, ← Units.val_pow_eq_pow_val, mul_pow, ← pow_mul]; rfl⟩ - -lemma add_equiv {P P' Q Q' : Fin 3 → R} (hP : P ≈ P') (hQ : Q ≈ Q') : W.add P Q ≈ W.add P' Q' := by - rcases hP, hQ with ⟨⟨u, rfl⟩, ⟨v, rfl⟩⟩ - exact W.add_smul_equiv P' Q' u v - -/-- The addition of two point classes. If `P` is a point representative, -then `W.add_map ⟦P⟧ ⟦Q⟧` is definitionally equivalent to `W.add P Q`. -/ -@[pp_dot] -noncomputable def add_map (P Q : PointClass R) : PointClass R := - Quotient.map₂ W.add (fun _ _ hP _ _ hQ => W.add_equiv hP hQ) P Q - -lemma add_map_eq (P Q : Fin 3 → R) : W.add_map ⟦P⟧ ⟦Q⟧ = ⟦W.add P Q⟧ := - rfl - -@[simp] -lemma add_map_of_Zeq0_left {P : Fin 3 → R} {Q : PointClass R} (hPz : P z = 0) : - W.add_map ⟦P⟧ Q = Q := by - rcases Q with ⟨Q⟩ - erw [add_map_eq, W.add_of_Zeq0_left hPz] - rfl - -lemma add_map_zero_left (P : PointClass R) : W.add_map ⟦![1, 1, 0]⟧ P = P := - W.add_map_of_Zeq0_left rfl - -@[simp] -lemma add_map_of_Zeq0_right {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z = 0) : - W.add_map ⟦P⟧ ⟦Q⟧ = ⟦P⟧ := by - rw [add_map_eq, W.add_of_Zeq0_right hPz hQz] - -lemma add_map_zero_right {P : Fin 3 → R} (hPz : P z ≠ 0) : W.add_map ⟦P⟧ ⟦![1, 1, 0]⟧ = ⟦P⟧ := by - rw [add_map_eq, W.add_zero_right hPz] - -@[simp] -lemma add_map_of_Yeq {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q) : - W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![1, 1, 0]⟧ := by - rw [add_map_eq, W.add_of_Yeq hPz hQz hx hy] - -@[simp] -lemma add_map_of_Yne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 = P z ^ 2 * Q x) (hy : P y * Q z ^ 3 ≠ P z ^ 3 * W.negY Q) : - W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Yne P, W.addY_of_Yne P, W.addZ_of_Yne P]⟧ := by - rw [add_map_eq, W.add_of_Yne hPz hQz hx hy] - -@[simp] -lemma add_map_of_Xne {P Q : Fin 3 → R} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) - (hx : P x * Q z ^ 2 ≠ P z ^ 2 * Q x) : - W.add_map ⟦P⟧ ⟦Q⟧ = ⟦![W.addX_of_Xne P Q, W.addY_of_Xne P Q, addZ_of_Xne P Q]⟧ := by - rw [add_map_eq, W.add_of_Xne hPz hQz hx] - -variable {F : Type u} [Field F] {W : Jacobian F} - -@[simp] -lemma add_map_of_Zeq0_right' {P : PointClass F} {Q : Fin 3 → F} (hP : W.nonsingular_lift P) - (hQ : W.nonsingular Q) (hQz : Q z = 0) : W.add_map P ⟦Q⟧ = P := by - rcases P with ⟨P⟩ - by_cases hPz : P z = 0 - · erw [W.add_map_of_Zeq0_left hPz, Quotient.eq] - exact equiv_of_Zeq0 hQ hP hQz hPz - · exact W.add_map_of_Zeq0_right hPz hQz - -lemma add_map_zero_right' {P : PointClass F} (hP : W.nonsingular_lift P) : - W.add_map P ⟦![1, 1, 0]⟧ = P := - add_map_of_Zeq0_right' hP W.nonsingular_zero rfl - -variable {F : Type u} [Field F] {W : Jacobian F} - -/-- The negation of a nonsingular point representative in `W` lies in `W`. -/ -lemma nonsingular_neg {P : Fin 3 → F} (h : W.nonsingular P) : W.nonsingular <| W.neg P := by - by_cases hPz : P z = 0 - · rw [W.nonsingular_of_equiv <| W.neg_equiv <| equiv_zero_of_Zeq0 h hPz, neg_zero] - convert W.nonsingular_zero' <| neg_ne_zero.mpr one_ne_zero <;> norm_num1 - · rw [nonsingular_iff_affine_of_Zne0 <| by exact hPz] at h ⊢ - rwa [← Affine.nonsingular_neg_iff, ← negY_divZ hPz] at h - -lemma nonsingular_lift_neg_map {P : PointClass F} (h : W.nonsingular_lift P) : - W.nonsingular_lift <| W.neg_map P := by - rcases P with ⟨_⟩ - exact nonsingular_neg h - -/-- The addition of two nonsingular point representatives in `W` lies in `W`. -/ -lemma nonsingular_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : - W.nonsingular <| W.add P Q := by - by_cases hPz : P z = 0 - · rwa [W.nonsingular_of_equiv <| W.add_equiv (equiv_zero_of_Zeq0 hP hPz) <| Setoid.refl Q, - W.add_of_Zeq0_left <| by exact rfl] - · by_cases hQz : Q z = 0 - · rwa [W.nonsingular_of_equiv <| W.add_equiv (Setoid.refl P) <| equiv_zero_of_Zeq0 hQ hQz, - W.add_of_Zeq0_right hPz <| by exact rfl] - · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x - · by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q - · simpa only [W.add_of_Yeq hPz hQz hx hy] using W.nonsingular_zero - · erw [W.add_of_Yne hPz hQz hx hy, - nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, - addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy] - exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) - (nonsingular_affine_of_Zne0 hQ hQz) fun _ => (negY_divZ hQz).symm ▸ Function.comp - (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp - · erw [W.add_of_Xne hPz hQz hx, - nonsingular_iff_affine_of_Zne0 <| addZ_ne_zero_of_Xne hx, - addX_div_addZ_of_Xne hP hQ hPz hQz hx, addY_div_addZ_of_Xne hP hQ hPz hQz hx] - exact W.toAffine.nonsingular_add (nonsingular_affine_of_Zne0 hP hPz) - (nonsingular_affine_of_Zne0 hQ hQz) fun h => False.elim <| hx <| - mul_comm (Q x) _ ▸ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp h - -lemma nonsingular_lift_add_map {P Q : PointClass F} (hP : W.nonsingular_lift P) - (hQ : W.nonsingular_lift Q) : W.nonsingular_lift <| W.add_map P Q := by - rcases P, Q with ⟨⟨_⟩, ⟨_⟩⟩ - exact nonsingular_add hP hQ - -end Representative - -section Point - -/-! ### Group operations on nonsingular rational points -/ - -/-- A nonsingular rational point on `W`. -/ -@[pp_dot] -structure Point where - {point : PointClass R} - (nonsingular : W.nonsingular_lift point) - -attribute [pp_dot] Point.point -attribute [pp_dot] Point.nonsingular - -/-- The point class underlying a nonsingular rational point on `W`. -/ -add_decl_doc Point.point - -/-- The nonsingular condition underlying a nonsingular rational point on `W`. -/ -add_decl_doc Point.nonsingular - -namespace Point - -variable {W} - -instance instZeroPoint [Nontrivial R] : Zero W.Point := - ⟨⟨W.nonsingular_lift_zero⟩⟩ - -lemma zero_def [Nontrivial R] : (⟨W.nonsingular_lift_zero⟩ : W.Point) = 0 := - rfl - -/-- The map from a nonsingular rational point on a Weierstrass curve `W` in affine coordinates -to the corresponding nonsingular rational point on `W` in Jacobian coordinates. -/ -def fromAffine [Nontrivial R] : W.toAffine.Point → W.Point - | 0 => 0 - | Affine.Point.some h => ⟨(W.nonsingular_lift_some ..).mpr h⟩ - -@[simp] -lemma fromAffine_zero [Nontrivial R] : fromAffine 0 = (0 : W.Point) := - rfl - -@[simp] -lemma fromAffine_some [Nontrivial R] {X Y : R} (h : W.toAffine.nonsingular X Y) : - fromAffine (Affine.Point.some h) = ⟨(W.nonsingular_lift_some ..).mpr h⟩ := - rfl - -variable {F : Type u} [Field F] {W : Jacobian F} - -/-- The negation of a nonsingular rational point on `W`. - -Given a nonsingular rational point `P` on `W`, use `-P` instead of `neg P`. -/ -@[pp_dot] -def neg (P : W.Point) : W.Point := - ⟨W.nonsingular_lift_neg_map P.nonsingular⟩ - -instance instNegPoint : Neg W.Point := - ⟨neg⟩ - -lemma neg_def (P : W.Point) : P.neg = -P := - rfl - -@[simp] -lemma neg_zero : (-⟨W.nonsingular_lift_zero⟩ : W.Point) = ⟨W.nonsingular_lift_zero⟩ := by - simp only [← neg_def, neg, neg_map_zero] - -/-- The addition of two nonsingular rational points on `W`. - -Given two nonsingular rational points `P` and `Q` on `W`, use `P + Q` instead of `add P Q`. -/ -@[pp_dot] -noncomputable def add (P Q : W.Point) : W.Point := - ⟨W.nonsingular_lift_add_map P.nonsingular Q.nonsingular⟩ - -noncomputable instance instAddPoint : Add W.Point := - ⟨add⟩ - -lemma add_def (P Q : W.Point) : P.add Q = P + Q := - rfl - -@[simp] -lemma zero_add (P : W.Point) : ⟨W.nonsingular_lift_zero⟩ + P = P := by - simp only [← add_def, add, add_map_zero_left] - -@[simp] -lemma add_zero (P : W.Point) : P + ⟨W.nonsingular_lift_zero⟩ = P := by - simp only [← add_def, add, add_map_zero_right' P.nonsingular] - -noncomputable instance instAddZeroClassPoint : AddZeroClass W.Point := - ⟨zero_add, add_zero⟩ - -open scoped Classical - -/-- The map from a point representative that is nonsingular on a Weierstrass curve `W` in Jacobian -coordinates to the corresponding nonsingular rational point on `W` in affine coordinates. -/ -noncomputable def toAffine {P : Fin 3 → F} (h : W.nonsingular P) : W.toAffine.Point := - if hPz : P z = 0 then 0 else Affine.Point.some <| nonsingular_affine_of_Zne0 h hPz - -lemma toAffine_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z = 0) : toAffine h = 0 := - dif_pos hPz - -lemma toAffine_zero : toAffine W.nonsingular_zero = 0 := - toAffine_of_Zeq0 rfl - -lemma toAffine_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular P} (hPz : P z ≠ 0) : - toAffine h = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := - dif_neg hPz - -lemma toAffine_some {X Y : F} (h : W.nonsingular ![X, Y, 1]) : - toAffine h = Affine.Point.some ((W.nonsingular_some X Y).mp h) := by - rw [toAffine_of_Zne0 <| by exact one_ne_zero] - matrix_simp - simp only [one_pow, div_one] - -lemma toAffine_neg {P : Fin 3 → F} (hP : W.nonsingular P) : - toAffine (nonsingular_neg hP) = -toAffine hP := by - by_cases hPz : P z = 0 - · rw [toAffine_of_Zeq0 <| by exact hPz, toAffine_of_Zeq0 hPz, Affine.Point.neg_zero] - · rw [toAffine_of_Zne0 <| by exact hPz, toAffine_of_Zne0 hPz, Affine.Point.neg_some, - Affine.Point.some.injEq] - exact ⟨rfl, negY_divZ hPz⟩ - -lemma toAffine_add {P Q : Fin 3 → F} (hP : W.nonsingular P) (hQ : W.nonsingular Q) : - toAffine (nonsingular_add hP hQ) = toAffine hP + toAffine hQ := by - by_cases hPz : P z = 0 - · simp only [W.add_of_Zeq0_left (Q := Q) hPz, toAffine_of_Zeq0 hPz, _root_.zero_add] - · by_cases hQz : Q z = 0 - · simp only [W.add_of_Zeq0_right hPz hQz, toAffine_of_Zeq0 hQz, _root_.add_zero] - · by_cases hx : P x * Q z ^ 2 = P z ^ 2 * Q x - · have hx' : P x / P z ^ 2 = Q x / Q z ^ 2 := - (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mpr <| mul_comm (P z ^ 2) _ ▸ hx - rw [toAffine_of_Zne0 hPz, toAffine_of_Zne0 hQz] - by_cases hy : P y * Q z ^ 3 = P z ^ 3 * W.negY Q - · have hy' : P y / P z ^ 3 = W.negY Q / Q z ^ 3 := Iff.mpr - (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)) <| mul_comm (P z ^ 3) _ ▸ hy - simp only [W.add_of_Yeq hPz hQz hx hy] - rw [toAffine_zero, Affine.Point.some_add_some_of_Yeq hx' <| by rwa [← negY_divZ hQz]] - · have hy' : P y / P z ^ 3 ≠ W.negY Q / Q z ^ 3 := Function.comp - (mul_comm (P z ^ 3) _ ▸ hy) (div_eq_div_iff (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)).mp - simp only [W.add_of_Yne hPz hQz hx hy] - rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Yne hP hQ hPz hQz hx hy, - Affine.Point.some_add_some_of_Yne hx' <| (negY_divZ hQz).symm ▸ hy', - Affine.Point.some.injEq] - exact ⟨addX_div_addZ_of_Yne hP hQ hPz hQz hx hy, addY_div_addZ_of_Yne hP hQ hPz hQz hx hy⟩ - · have hx' : P x / P z ^ 2 ≠ Q x / Q z ^ 2 := - (mul_comm (P z ^ 2) _ ▸ hx) ∘ (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mp - simp only [W.add_of_Xne hPz hQz hx] - rw [toAffine_of_Zne0 <| addZ_ne_zero_of_Xne hx, toAffine_of_Zne0 hPz, toAffine_of_Zne0 hQz, - Affine.Point.some_add_some_of_Xne hx', Affine.Point.some.injEq] - exact ⟨addX_div_addZ_of_Xne hP hQ hPz hQz hx, addY_div_addZ_of_Xne hP hQ hPz hQz hx⟩ - -lemma toAffine_of_equiv (P Q : Fin 3 → F) (h : P ≈ Q) : - HEq (toAffine (W := W) (P := P)) (toAffine (W := W) (P := Q)) := by - rcases h with ⟨u, rfl⟩ - refine Function.hfunext (propext <| W.nonsingular_smul_iff Q u) <| fun _ _ _ => ?_ - by_cases hPz : Q z = 0 - · rw [toAffine_of_Zeq0 <| by exact u.mul_right_eq_zero.mpr hPz, toAffine_of_Zeq0 hPz] - · rw [toAffine_of_Zne0 <| by exact mul_ne_zero u.ne_zero hPz, toAffine_of_Zne0 hPz, heq_eq_eq] - simp only [Affine.Point.some.injEq, smul_fin3_ext, mul_pow, - mul_div_mul_left _ _ <| pow_ne_zero _ u.ne_zero, and_self] - -/-- The map from a nonsingular rational point on a Weierstrass curve `W` in Jacobian coordinates -to the corresponding nonsingular rational point on `W` in affine coordinates. -/ -noncomputable def toAffine_lift (P : W.Point) : W.toAffine.Point := - P.point.hrecOn (fun _ => toAffine) toAffine_of_equiv P.nonsingular - -lemma toAffine_lift_eq {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : - toAffine_lift ⟨h⟩ = toAffine h := - rfl - -lemma toAffine_lift_of_Zeq0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z = 0) : - toAffine_lift ⟨h⟩ = 0 := - toAffine_of_Zeq0 hPz (h := h) - -lemma toAffine_lift_zero : toAffine_lift (0 : W.Point) = 0 := - toAffine_zero - -lemma toAffine_lift_of_Zne0 {P : Fin 3 → F} {h : W.nonsingular_lift ⟦P⟧} (hPz : P z ≠ 0) : - toAffine_lift ⟨h⟩ = Affine.Point.some (nonsingular_affine_of_Zne0 h hPz) := - toAffine_of_Zne0 hPz (h := h) - -lemma toAffine_lift_some {X Y : F} (h : W.nonsingular_lift ⟦![X, Y, 1]⟧) : - toAffine_lift ⟨h⟩ = Affine.Point.some ((W.nonsingular_some X Y).mp h) := - toAffine_some h - -lemma toAffine_lift_neg {P : Fin 3 → F} (h : W.nonsingular_lift ⟦P⟧) : - toAffine_lift (-⟨h⟩) = -toAffine_lift ⟨h⟩ := - toAffine_neg h - -lemma toAffine_lift_add {P Q : Fin 3 → F} (hP : W.nonsingular_lift ⟦P⟧) - (hQ : W.nonsingular_lift ⟦Q⟧) : - toAffine_lift (⟨hP⟩ + ⟨hQ⟩) = toAffine_lift ⟨hP⟩ + toAffine_lift ⟨hQ⟩ := - toAffine_add hP hQ - -/-- The equivalence between the nonsingular rational points on a Weierstrass curve `W` in Jacobian -coordinates with the nonsingular rational points on `W` in affine coordinates. -/ -@[simps] -noncomputable def toAffine_addEquiv : W.Point ≃+ W.toAffine.Point where - toFun := toAffine_lift - invFun := fromAffine - left_inv := by - rintro @⟨⟨P⟩, h⟩ - by_cases hPz : P z = 0 - · erw [toAffine_lift_eq, toAffine_of_Zeq0 hPz, fromAffine_zero, mk.injEq, Quotient.eq] - exact Setoid.symm <| equiv_zero_of_Zeq0 h hPz - · erw [toAffine_lift_eq, toAffine_of_Zne0 hPz, fromAffine_some, mk.injEq, Quotient.eq] - exact Setoid.symm <| equiv_some_of_Zne0 hPz - right_inv := by - rintro (_ | _) - · erw [fromAffine_zero, toAffine_lift_zero, Affine.Point.zero_def] - · rw [fromAffine_some, toAffine_lift_some] - map_add' := by - rintro @⟨⟨_⟩, _⟩ @⟨⟨_⟩, _⟩ - simpa only using toAffine_lift_add .. - -end Point - -end Point - -end WeierstrassCurve.Jacobian From 2e23f4199277c130d1786ebf13e5d1cc90cf4ed9 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 23 Feb 2024 10:48:17 +0000 Subject: [PATCH 21/47] Compute degrees of division polynomials --- .../EllipticCurve/DivisionPolynomial.lean | 531 +++++++++++++++++- 1 file changed, 529 insertions(+), 2 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index 9d195c6e534669..29db4273f88f7b 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -6,6 +6,7 @@ Authors: David Kurniadi Angdinata import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.Data.Int.Parity import Mathlib.NumberTheory.EllipticDivisibilitySequence +import Mathlib.Tactic.ComputeDegree /-! # Division polynomials of Weierstrass curves @@ -43,6 +44,17 @@ if $n$ is even or $\tilde{\psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ otherwise. * `WeierstrassCurve.divisionPolynomialZ`: $\psi_n$ viewed as a bivariate polynomial. * TODO: $\omega$ viewed as a bivariate polynomial. +## Main statements + + * `WeierstrassCurve.natDegree_divisionPolynomial2Sq`: the degree of $\psi_2^2$. + * `WeierstrassCurve.coeff_divisionPolynomial2Sq`: the leading coefficient of $\psi_2^2$. + * `WeierstrassCurve.natDegree_divisionPolynomial`: the degree of $\tilde{\psi}_n$. + * `WeierstrassCurve.coeff_divisionPolynomial`: the leading coefficient of $\tilde{\psi}_n$. + * `WeierstrassCurve.natDegree_divisionPolynomialX`: the degree of $\phi_n$. + * `WeierstrassCurve.coeff_divisionPolynomialX`: the leading coefficient of $\phi_n$. + * `WeierstrassCurve.natDegree_divisionPolynomialZSq`: the degree of $\psi_n^2$. + * `WeierstrassCurve.coeff_divisionPolynomialZSq`: the leading coefficient of $\psi_n^2$. + ## Implementation notes Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate @@ -73,6 +85,8 @@ namespace WeierstrassCurve variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) +/-! ### The univariate polynomial congruent to $\psi_2^2$ -/ + /-- The polynomial $4X^3 + b_2X^2 + 2b_4X + b_6$ congruent to the square $\psi_2^2$ of the $2$-division polynomial $\psi_2 = 2Y + a_1X + a_3$ under $R[W]$. -/ @[pp_dot] @@ -83,6 +97,17 @@ noncomputable def divisionPolynomial2Sq : R[X] := lemma divisionPolynomial2Sq_eq : W.divisionPolynomial2Sq = W.twoTorsionPolynomial.toPoly := rfl +lemma natDegree_divisionPolynomial2Sq : W.divisionPolynomial2Sq.natDegree ≤ 3 := by + rw [divisionPolynomial2Sq] + compute_degree + +@[simp] +lemma coeff_divisionPolynomial2Sq : W.divisionPolynomial2Sq.coeff 3 = 4 := by + rw [divisionPolynomial2Sq] + compute_degree! + +/-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{N}$ -/ + /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials $\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{N}$. -/ @[pp_dot] @@ -129,6 +154,355 @@ lemma divisionPolynomial'_even (m : ℕ) : W.divisionPolynomial' (2 * (m + 3)) = W.divisionPolynomial' (m + 4) ^ 2 := preNormEDS'_even .. +private lemma natDegree_divisionPolynomial'_zero : (W.divisionPolynomial' 0).natDegree = 0 := by + rw [divisionPolynomial'_zero, natDegree_zero] + +private lemma coeff_divisionPolynomial'_zero : (W.divisionPolynomial' 0).coeff 0 = 0 := by + rw [divisionPolynomial'_zero, coeff_zero] + +private lemma natDegree_divisionPolynomial'_one : (W.divisionPolynomial' 1).natDegree = 0 := by + rw [divisionPolynomial'_one, natDegree_one] + +private lemma coeff_divisionPolynomial'_one : (W.divisionPolynomial' 1).coeff 0 = 1 := by + rw [divisionPolynomial'_one, coeff_one_zero] + +private lemma natDegree_divisionPolynomial'_two : (W.divisionPolynomial' 2).natDegree = 0 := by + rw [divisionPolynomial'_two, natDegree_one] + +private lemma coeff_divisionPolynomial'_two : (W.divisionPolynomial' 2).coeff 0 = 1 := by + rw [divisionPolynomial'_two, coeff_one_zero] + +private lemma natDegree_divisionPolynomial'_three : (W.divisionPolynomial' 3).natDegree ≤ 4 := by + rw [divisionPolynomial'_three] + compute_degree + +private lemma coeff_divisionPolynomial'_three : (W.divisionPolynomial' 3).coeff 4 = 3 := by + rw [divisionPolynomial'_three] + compute_degree! + +private lemma natDegree_divisionPolynomial'_four : (W.divisionPolynomial' 4).natDegree ≤ 6 := by + rw [divisionPolynomial'_four] + compute_degree + +private lemma coeff_divisionPolynomial'_four : (W.divisionPolynomial' 4).coeff 6 = 2 := by + rw [divisionPolynomial'_four] + compute_degree! + +private lemma natDegree_divisionPolynomial'_five : (W.divisionPolynomial' 5).natDegree ≤ 12 := by + rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial2Sq] + simp + compute_degree + +private lemma coeff_divisionPolynomial'_five : (W.divisionPolynomial' 5).coeff 12 = 5 := by + rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial2Sq] + simp [-coeff_sub] + compute_degree! + +private lemma natDegree_divisionPolynomial'_six : (W.divisionPolynomial' 6).natDegree ≤ 16 := by + rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, + divisionPolynomial'_odd, divisionPolynomial2Sq] + simp + compute_degree + +private lemma coeff_divisionPolynomial'_six : (W.divisionPolynomial' 6).coeff 16 = 3 := by + rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, + divisionPolynomial'_odd, divisionPolynomial2Sq] + simp [-coeff_sub] + compute_degree! + +private lemma natDegree_divisionPolynomial'_eight : (W.divisionPolynomial' 8).natDegree ≤ 30 := by + rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, show 1 + 5 = 2 * 3 by rfl, + divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, + divisionPolynomial2Sq] + simp + compute_degree + +private lemma coeff_divisionPolynomial'_eight : (W.divisionPolynomial' 8).coeff 30 = 4 := by + rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, show 1 + 5 = 2 * 3 by rfl, + divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, + divisionPolynomial2Sq] + simp [-coeff_sub] + compute_degree! + +section Inductive + +variable {W} (m : ℕ) (ih : ∀ k < 2 * m + 2, + ((W.divisionPolynomial' <| 2 * k).natDegree ≤ 2 * k ^ 2 - 2 ∧ + (W.divisionPolynomial' <| 2 * k).coeff (2 * k ^ 2 - 2) = k) ∧ + ((W.divisionPolynomial' <| 2 * k + 1).natDegree ≤ 2 * k ^ 2 + 2 * k ∧ + (W.divisionPolynomial' <| 2 * k + 1).coeff (2 * k ^ 2 + 2 * k) = (2 * k + 1 : ℕ))) + +private lemma natDegree_divisionPolynomial'_add_one : + (W.divisionPolynomial' <| 2 * m + 1).natDegree ≤ 2 * m * (m + 1) := by + convert (ih m <| by linarith only).right.left using 1 + ring1 + +private lemma coeff_divisionPolynomial'_add_one : + (W.divisionPolynomial' <| 2 * m + 1).coeff (2 * m * (m + 1)) = (2 * m + 1 : ℕ) := by + convert (ih m <| by linarith only).right.right using 2 + ring1 + +private lemma natDegree_divisionPolynomial'_add_two : + (W.divisionPolynomial' <| 2 * m + 2).natDegree ≤ 2 * m * (m + 2) := by + convert (ih (m + 1) <| by linarith only).left.left using 1 + rw [show 2 * (m + 1) ^ 2 = 2 * m * (m + 2) + 2 by ring1, Nat.add_sub_cancel] + +private lemma coeff_divisionPolynomial'_add_two : + (W.divisionPolynomial' <| 2 * m + 2).coeff (2 * m * (m + 2)) = (m + 1 : ℕ) := by + convert (ih (m + 1) <| by linarith only).left.right using 2 + rw [show 2 * (m + 1) ^ 2 = 2 * m * (m + 2) + 2 by ring1, Nat.add_sub_cancel] + +private lemma natDegree_divisionPolynomial'_add_three : + (W.divisionPolynomial' <| 2 * m + 3).natDegree ≤ 2 * (m + 1) * (m + 2) := by + convert (ih (m + 1) <| by linarith only).right.left using 1 + ring1 + +private lemma coeff_divisionPolynomial'_add_three : + (W.divisionPolynomial' <| 2 * m + 3).coeff (2 * (m + 1) * (m + 2)) = (2 * m + 3 : ℕ) := by + convert (ih (m + 1) <| by linarith only).right.right using 2 + ring1 + +private lemma natDegree_divisionPolynomial'_add_four : + (W.divisionPolynomial' <| 2 * m + 4).natDegree ≤ 2 * (m + 1) * (m + 3) := by + rcases m with _ | m + · exact W.natDegree_divisionPolynomial'_four + · convert (ih (m + 3) <| by linarith only).left.left using 1 + rw [show 2 * (m + 3) ^ 2 = 2 * (m + 2) * (m + 4) + 2 by ring1, Nat.add_sub_cancel] + +private lemma coeff_divisionPolynomial'_add_four : + (W.divisionPolynomial' <| 2 * m + 4).coeff (2 * (m + 1) * (m + 3)) = (m + 2 : ℕ) := by + rcases m with _ | m + · exact W.coeff_divisionPolynomial'_four + · convert (ih (m + 3) <| by linarith only).left.right using 2 + rw [show 2 * (m + 3) ^ 2 = 2 * (m + 2) * (m + 4) + 2 by ring1, Nat.add_sub_cancel] + +private lemma natDegree_divisionPolynomial'_add_five : + (W.divisionPolynomial' <| 2 * m + 5).natDegree ≤ 2 * (m + 2) * (m + 3) := by + rcases m with _ | m + · exact W.natDegree_divisionPolynomial'_five + · convert (ih (m + 3) <| by linarith only).right.left using 1 + rw [← Nat.add_one] + ring1 + +private lemma coeff_divisionPolynomial'_add_five : + (W.divisionPolynomial' <| 2 * m + 5).coeff (2 * (m + 2) * (m + 3)) = (2 * m + 5 : ℕ) := by + rcases m with _ | m + · exact W.coeff_divisionPolynomial'_five + · convert (ih (m + 3) <| by linarith only).right.right using 2 + rw [← Nat.add_one] + ring1 + +private lemma natDegree_divisionPolynomial'_add_six : + (W.divisionPolynomial' <| 2 * m + 6).natDegree ≤ 2 * (m + 2) * (m + 4) := by + rcases m with _ | _ | m + · exact W.natDegree_divisionPolynomial'_six + · exact W.natDegree_divisionPolynomial'_eight + · convert (ih (m + 5) <| by linarith only).left.left using 1 + rw [show 2 * (m + 5) ^ 2 = 2 * (m + 4) * (m + 6) + 2 by ring1, Nat.add_sub_cancel] + +private lemma coeff_divisionPolynomial'_add_six : + (W.divisionPolynomial' <| 2 * m + 6).coeff (2 * (m + 2) * (m + 4)) = (m + 3 : ℕ) := by + rcases m with _ | _ | m + · exact W.coeff_divisionPolynomial'_six + · exact W.coeff_divisionPolynomial'_eight + · convert (ih (m + 5) <| by linarith only).left.right using 2 + rw [show 2 * (m + 5) ^ 2 = 2 * (m + 4) * (m + 6) + 2 by ring1, Nat.add_sub_cancel] + +private lemma natDegree_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : + (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).natDegree ≤ n := by + nth_rw 1 [← max_self n, habc, hdef] + convert natDegree_sub_le_of_le .. using 1 + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| + natDegree_pow_le_of_le 2 hc + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| + natDegree_pow_le_of_le 2 hf + +private lemma natDegree_odd_divisionPolynomial'_even : + (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 3) := by + rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, + ← @one_pow R[X] _ 2] + exact natDegree_odd (natDegree_divisionPolynomial'_add_four m ih) + (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq + (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) + natDegree_one.le (by ring1) (by ring1) + +private lemma natDegree_odd_divisionPolynomial'_odd : + (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).natDegree ≤ 2 * (2 * m + 3) * (2 * m + 4) := by + rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X] _ 2, + if_neg m.not_even_two_mul_add_one] + exact natDegree_odd (natDegree_divisionPolynomial'_add_five m ih) + (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le + (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) + W.natDegree_divisionPolynomial2Sq (by ring1) (by ring1) + +private lemma coeff_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : + (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).coeff n = a.coeff da * b.coeff db ^ 3 * c.coeff dc ^ 2 - + d.coeff dd * e.coeff de ^ 3 * f.coeff df ^ 2 := by + rw [coeff_sub, habc, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| natDegree_pow_le_of_le 2 hc, + coeff_pow_of_natDegree_le hc, coeff_mul_of_natDegree_le ha <| natDegree_pow_le_of_le 3 hb, + coeff_pow_of_natDegree_le hb, ← habc, hdef, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| natDegree_pow_le_of_le 2 hf, + coeff_mul_of_natDegree_le hd <| natDegree_pow_le_of_le 3 he, coeff_pow_of_natDegree_le he, + coeff_pow_of_natDegree_le hf] + +private lemma coeff_odd_divisionPolynomial'_even : + (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).coeff (2 * (2 * m + 2) * (2 * m + 3)) = + (2 * (2 * m + 2) + 1 : ℕ) := by + rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, + ← @one_pow R[X] _ 2, coeff_odd (natDegree_divisionPolynomial'_add_four m ih) + (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq + (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) + natDegree_one.le (by ring1) (by ring1), coeff_divisionPolynomial'_add_four m ih, + coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial2Sq, + coeff_divisionPolynomial'_add_one m ih, coeff_divisionPolynomial'_add_three m ih, + coeff_one_zero] + push_cast + ring1 + +private lemma coeff_odd_divisionPolynomial'_odd : + (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).coeff (2 * (2 * m + 3) * (2 * m + 4)) = + (2 * (2 * m + 3) + 1 : ℕ) := by + rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X] _ 2, + if_neg m.not_even_two_mul_add_one, coeff_odd (natDegree_divisionPolynomial'_add_five m ih) + (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le + (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) + W.natDegree_divisionPolynomial2Sq (by ring1) (by ring1), + coeff_divisionPolynomial'_add_five m ih, coeff_divisionPolynomial'_add_three m ih, + coeff_one_zero, coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial'_add_four m ih, + coeff_divisionPolynomial2Sq] + push_cast + ring1 + +private lemma natDegree_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : + (a ^ 2 * b * c - d * e * f ^ 2).natDegree ≤ n := by + nth_rw 1 [← max_self n, habc, hdef] + convert natDegree_sub_le_of_le .. using 1 + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd he) <| natDegree_pow_le_of_le 2 hf + +private lemma natDegree_even_divisionPolynomial'_even : + (W.divisionPolynomial' <| 2 * (2 * m + 3)).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 4) := by + rw [divisionPolynomial'_even] + exact natDegree_even (natDegree_divisionPolynomial'_add_two m ih) + (natDegree_divisionPolynomial'_add_three m ih) (natDegree_divisionPolynomial'_add_five m ih) + (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) + (natDegree_divisionPolynomial'_add_four m ih) (by ring1) (by ring1) + +private lemma natDegree_even_divisionPolynomial'_odd : + (W.divisionPolynomial' <| 2 * (2 * m + 4)).natDegree ≤ 2 * (2 * m + 3) * (2 * m + 5) := by + rw [divisionPolynomial'_even] + exact natDegree_even (natDegree_divisionPolynomial'_add_three m ih) + (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_six m ih) + (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) + (natDegree_divisionPolynomial'_add_five m ih) (by ring1) (by ring1) + +private lemma coeff_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : + (a ^ 2 * b * c - d * e * f ^ 2).coeff n = + a.coeff da ^ 2 * b.coeff db * c.coeff dc - d.coeff dd * e.coeff de * f.coeff df ^ 2 := by + rw [coeff_sub, habc, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc, + coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 ha) hb, coeff_pow_of_natDegree_le ha, + ← habc, hdef, coeff_mul_of_natDegree_le (natDegree_mul_le_of_le hd he) <| + natDegree_pow_le_of_le 2 hf, coeff_mul_of_natDegree_le hd he, coeff_pow_of_natDegree_le hf] + +private lemma coeff_even_divisionPolynomial'_even : + (W.divisionPolynomial' <| 2 * (2 * m + 3)).coeff (2 * (2 * m + 2) * (2 * m + 4)) = + (2 * m + 3 : ℕ) := by + rw [divisionPolynomial'_even, coeff_even (natDegree_divisionPolynomial'_add_two m ih) + (natDegree_divisionPolynomial'_add_three m ih) (natDegree_divisionPolynomial'_add_five m ih) + (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) + (natDegree_divisionPolynomial'_add_four m ih) (by ring1) (by ring1), + coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial'_add_three m ih, + coeff_divisionPolynomial'_add_five m ih, coeff_divisionPolynomial'_add_one m ih, + coeff_divisionPolynomial'_add_four m ih] + push_cast + ring1 + +private lemma coeff_even_divisionPolynomial'_odd : + (W.divisionPolynomial' <| 2 * (2 * m + 4)).coeff (2 * (2 * m + 3) * (2 * m + 5)) = + (2 * m + 4 : ℕ) := by + rw [divisionPolynomial'_even, coeff_even (natDegree_divisionPolynomial'_add_three m ih) + (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_six m ih) + (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) + (natDegree_divisionPolynomial'_add_five m ih) (by ring1) (by ring1), + coeff_divisionPolynomial'_add_three m ih, coeff_divisionPolynomial'_add_four m ih, + coeff_divisionPolynomial'_add_six m ih, coeff_divisionPolynomial'_add_two m ih, + coeff_divisionPolynomial'_add_five m ih] + push_cast + ring1 + +end Inductive + +private lemma natDegree_coeff_divisionPolynomial' (n : ℕ) : + if Even n then (W.divisionPolynomial' n).natDegree ≤ (n ^ 2 - 4) / 2 ∧ + (W.divisionPolynomial' n).coeff ((n ^ 2 - 4) / 2) = (n / 2 : ℕ) + else (W.divisionPolynomial' n).natDegree ≤ (n ^ 2 - 1) / 2 ∧ + (W.divisionPolynomial' n).coeff ((n ^ 2 - 1) / 2) = n := by + induction n using normEDSRec' with + | base0 => exact ⟨W.natDegree_divisionPolynomial'_zero.le, + by erw [coeff_divisionPolynomial'_zero, Nat.cast_zero]⟩ + | base1 => exact ⟨W.natDegree_divisionPolynomial'_one.le, + by erw [coeff_divisionPolynomial'_one, Nat.cast_one]⟩ + | base2 => exact ⟨W.natDegree_divisionPolynomial'_two.le, + by erw [coeff_divisionPolynomial'_two, Nat.cast_one]⟩ + | base3 => exact ⟨W.natDegree_divisionPolynomial'_three, W.coeff_divisionPolynomial'_three⟩ + | base4 => exact ⟨W.natDegree_divisionPolynomial'_four, W.coeff_divisionPolynomial'_four⟩ + | _ m ih => + replace ih (k : ℕ) (hk : k < m + 2) : + ((W.divisionPolynomial' <| 2 * k).natDegree ≤ 2 * k ^ 2 - 2 ∧ + (W.divisionPolynomial' <| 2 * k).coeff (2 * k ^ 2 - 2) = k) ∧ + ((W.divisionPolynomial' <| 2 * k + 1).natDegree ≤ 2 * k ^ 2 + 2 * k ∧ + (W.divisionPolynomial' <| 2 * k + 1).coeff (2 * k ^ 2 + 2 * k) = (2 * k + 1 : ℕ)) := by + rw [← show ((2 * k) ^ 2 - 4) / 2 = 2 * k ^ 2 - 2 by + exact Nat.div_eq_of_eq_mul_right two_pos <| by rw [Nat.mul_sub_left_distrib]; ring_nf, + ← show ((2 * k + 1) ^ 2 - 1) / 2 = 2 * k ^ 2 + 2 * k by + exact Nat.div_eq_of_eq_mul_right two_pos <| by erw [add_sq, Nat.add_sub_cancel]; ring1] + nth_rw 5 [← k.mul_div_cancel_left two_pos] + exact ⟨imp_of_if_pos (ih (2 * k) <| by linarith only [hk]) <| even_two_mul k, + imp_of_if_neg (ih (2 * k + 1) <| by linarith only [hk]) k.not_even_two_mul_add_one⟩ + simp only [if_pos <| even_two_mul _, if_neg (m + 2).not_even_two_mul_add_one, + show (2 * (m + 3)) ^ 2 = 2 * (2 * (m + 2) * (m + 4)) + 4 by ring1, + show (2 * (m + 2) + 1) ^ 2 = 2 * (2 * (m + 2) * (m + 3)) + 1 by ring1, Nat.add_sub_cancel, + Nat.mul_div_cancel_left _ two_pos] + by_cases hm : Even m + · rcases (even_iff_exists_two_mul m).mp hm with ⟨m, rfl⟩ + simp only [natDegree_odd_divisionPolynomial'_even m ih, + coeff_odd_divisionPolynomial'_even m ih, natDegree_even_divisionPolynomial'_even m ih, + coeff_even_divisionPolynomial'_even m ih, and_self] + · replace ih (k : ℕ) (hk : k < m + 1) := ih k <| Nat.lt.step hk + rcases Nat.odd_iff_not_even.mpr hm with ⟨m, rfl⟩ + simp only [natDegree_odd_divisionPolynomial'_odd m ih, coeff_odd_divisionPolynomial'_odd m ih, + natDegree_even_divisionPolynomial'_odd m ih, coeff_even_divisionPolynomial'_odd m ih, + and_self] + +lemma natDegree_divisionPolynomial' (n : ℕ) : (W.divisionPolynomial' n).natDegree ≤ + if Even n then (n ^ 2 - 4) / 2 else (n ^ 2 - 1) / 2 := by + by_cases hn : Even n + · simpa only [if_pos hn] using (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial' n) hn).left + · simpa only [if_neg hn] using (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial' n) hn).left + +lemma coeff_divisionPolynomial' (n : ℕ) : + if Even n then (W.divisionPolynomial' n).coeff ((n ^ 2 - 4) / 2) = (n / 2 : ℕ) + else (W.divisionPolynomial' n).coeff ((n ^ 2 - 1) / 2) = n := by + by_cases hn : Even n + · rw [if_pos hn, (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial' n) hn).right] + · rw [if_neg hn, (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial' n) hn).right] + +/-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{Z}$ -/ + /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials $\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ @[pp_dot] @@ -177,6 +551,37 @@ lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) = lemma divisionPolynomial_neg (n : ℤ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := preNormEDS_neg .. +private lemma natDegree_coeff_divisionPolynomial (n : ℤ) : + if Even n then (W.divisionPolynomial n).natDegree ≤ (n.natAbs ^ 2 - 4) / 2 ∧ + (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 4) / 2) = (n / 2 : ℤ) + else (W.divisionPolynomial n).natDegree ≤ (n.natAbs ^ 2 - 1) / 2 ∧ + (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 1) / 2) = n := by + rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ + · norm_cast + exact W.divisionPolynomial_ofNat n ▸ W.natDegree_coeff_divisionPolynomial' n + · simp only [even_neg, Int.even_coe_nat, divisionPolynomial_neg, natDegree_neg, + divisionPolynomial_ofNat, Int.natAbs_neg, Int.natAbs_cast, coeff_neg, neg_eq_iff_eq_neg, + Int.cast_neg, Int.cast_ofNat, neg_neg] + convert W.natDegree_coeff_divisionPolynomial' n using 3 with hn + rcases (even_iff_exists_two_mul _).mp hn with ⟨_, rfl⟩ + rw [Nat.cast_mul, neg_mul_eq_mul_neg, Nat.cast_ofNat, Int.mul_ediv_cancel_left _ two_ne_zero, + Int.cast_neg, neg_neg, Int.cast_ofNat, Nat.mul_div_cancel_left _ two_pos] + +lemma natDegree_divisionPolynomial (n : ℤ) : (W.divisionPolynomial n).natDegree ≤ + if Even n then (n.natAbs ^ 2 - 4) / 2 else (n.natAbs ^ 2 - 1) / 2 := by + by_cases hn : Even n + · simpa only [if_pos hn] using (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial n) hn).left + · simpa only [if_neg hn] using (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial n) hn).left + +lemma coeff_divisionPolynomial (n : ℤ) : + if Even n then (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 4) / 2) = (n / 2 : ℤ) + else (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 1) / 2) = n := by + by_cases hn : Even n + · rw [if_pos hn, (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial n) hn).right] + · rw [if_neg hn, (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial n) hn).right] + +/-! ### The univariate polynomials congruent to $\psi_n^2$ for $n \in \mathbb{Z}$ -/ + /-- The univariate polynomials congruent under $R[W]$ to the bivariate squares $\psi_n^2$ of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ @[pp_dot] @@ -201,12 +606,12 @@ lemma divisionPolynomialZSq_two : W.divisionPolynomialZSq 2 = W.divisionPolynomi erw [divisionPolynomialZSq_ofNat, one_pow, one_mul, if_pos even_two] @[simp] -lemma divisionPolynomialZSq_three : W.divisionPolynomialZSq 3 = (W.divisionPolynomial' 3) ^ 2 := by +lemma divisionPolynomialZSq_three : W.divisionPolynomialZSq 3 = W.divisionPolynomial' 3 ^ 2 := by erw [divisionPolynomialZSq_ofNat, mul_one] @[simp] lemma divisionPolynomialZSq_four : W.divisionPolynomialZSq 4 = - (W.divisionPolynomial' 4) ^ 2 * W.divisionPolynomial2Sq := by + W.divisionPolynomial' 4 ^ 2 * W.divisionPolynomial2Sq := by erw [divisionPolynomialZSq_ofNat, if_pos <| by decide] @[simp] @@ -214,6 +619,58 @@ lemma divisionPolynomialZSq_neg (n : ℤ) : W.divisionPolynomialZSq (-n) = W.divisionPolynomialZSq n := by rw [divisionPolynomialZSq, divisionPolynomial_neg, neg_sq, Int.natAbs_neg, divisionPolynomialZSq] +private lemma natDegree_coeff_divisionPolynomialZSq_ofNat (n : ℕ) : + (W.divisionPolynomialZSq n).natDegree ≤ n ^ 2 - 1 ∧ + (W.divisionPolynomialZSq n).coeff (n ^ 2 - 1) = (n ^ 2 : ℕ) := by + rcases n with _ | n + · erw [divisionPolynomialZSq_zero] + exact ⟨natDegree_zero.le, Nat.cast_zero.symm⟩ + · have h := W.natDegree_coeff_divisionPolynomial' <| n + 1 + simp only [divisionPolynomialZSq, divisionPolynomial_ofNat, Int.natAbs_even, Int.even_coe_nat, + Nat.even_add_one, ite_not] at h ⊢ + by_cases hn : Even n + · rcases (even_iff_exists_two_mul n).mp hn with ⟨m, rfl⟩ + rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, + Nat.mul_div_cancel_left _ two_pos] at h + rw [if_pos hn, mul_one, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, + Nat.add_sub_cancel] + constructor + · exact natDegree_pow_le_of_le 2 h.left + · rw [coeff_pow_of_natDegree_le h.left, h.right] + push_cast + ring1 + · rcases Nat.odd_iff_not_even.mpr hn with ⟨m, rfl⟩ + rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.add_sub_cancel, + Nat.mul_div_cancel_left _ two_pos, show (2 * m + 2) / 2 = 2 * (m + 1) / 2 by rfl, + Nat.mul_div_cancel_left _ two_pos] at h + rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.succ_sub_one] + constructor + · exact natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 h.left) + W.natDegree_divisionPolynomial2Sq + · rw [coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 h.left) + W.natDegree_divisionPolynomial2Sq, coeff_pow_of_natDegree_le h.left, h.right, + coeff_divisionPolynomial2Sq] + push_cast + ring1 + +private lemma natDegree_coeff_divisionPolynomialZSq (n : ℤ) : + (W.divisionPolynomialZSq n).natDegree ≤ n.natAbs ^ 2 - 1 ∧ + (W.divisionPolynomialZSq n).coeff (n.natAbs ^ 2 - 1) = (n.natAbs ^ 2 : ℕ) := by + rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ + · exact W.natDegree_coeff_divisionPolynomialZSq_ofNat n + · simp only [divisionPolynomialZSq_neg, Int.natAbs_neg] + exact W.natDegree_coeff_divisionPolynomialZSq_ofNat n + +lemma natDegree_divisionPolynomialZSq (n : ℤ) : + (W.divisionPolynomialZSq n).natDegree ≤ n.natAbs ^ 2 - 1 := + (W.natDegree_coeff_divisionPolynomialZSq n).left + +lemma coeff_divisionPolynomialZSq (n : ℤ) : + (W.divisionPolynomialZSq n).coeff (n.natAbs ^ 2 - 1) = (n.natAbs ^ 2 : ℕ) := + (W.natDegree_coeff_divisionPolynomialZSq n).right + +/-! ### The univariate polynomials congruent to $\phi_n$ for $n \in \mathbb{Z}$ -/ + /-- The univariate polynomials congruent under $R[W]$ to the bivariate polynomials $\phi_n$ defined in terms of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ @[pp_dot] @@ -269,6 +726,76 @@ lemma divisionPolynomialX_neg (n : ℤ) : W.divisionPolynomialX (-n) = W.divisio divisionPolynomial_neg, ← neg_add', divisionPolynomial_neg, neg_mul_neg, mul_comm <| W.divisionPolynomial _, Int.natAbs_neg, divisionPolynomialX] +private lemma natDegree_coeff_divisionPolynomialX_ofNat (n : ℕ) : + (W.divisionPolynomialX n).natDegree ≤ n ^ 2 ∧ (W.divisionPolynomialX n).coeff (n ^ 2) = 1 := by + rcases n with _ | _ | n + · erw [divisionPolynomialX_zero] + exact ⟨natDegree_one.le, coeff_one_zero⟩ + · erw [divisionPolynomialX_one] + exact ⟨natDegree_X_le, coeff_X⟩ + · have h1 := W.natDegree_coeff_divisionPolynomial' <| n + 1 + have h2 := W.natDegree_coeff_divisionPolynomialZSq (n + 2 : ℕ) + have h3 := W.natDegree_coeff_divisionPolynomial' <| n + 3 + simp only [divisionPolynomialX, Int.natAbs_cast, Nat.even_add_one, ite_not] at h1 h3 ⊢ + erw [← Nat.cast_add, divisionPolynomial_ofNat, ← Nat.cast_sub <| by linarith only, + divisionPolynomial_ofNat, ← Nat.add_one, Nat.add_sub_cancel] + by_cases hn : Even n + · rcases (even_iff_exists_two_mul n).mp hn with ⟨m, rfl⟩ + rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, + Nat.mul_div_cancel_left _ two_pos] at h1 + rw [Int.natAbs_cast, show (2 * m + 2) ^ 2 = 4 * m * (m + 2) + 4 by ring1, Nat.succ_sub_one] + at h2 + rw [if_pos hn, show (2 * m + 3) ^ 2 = 2 * (2 * (m + 1) * (m + 2)) + 1 by ring1, + Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos] at h3 + rw [if_pos hn, mul_one] + constructor + · convert natDegree_sub_le_of_le (natDegree_mul_le_of_le natDegree_X_le h2.left) <| + natDegree_mul_le_of_le h3.left h1.left using 1 + convert (max_self _).symm using 2 <;> ring1 + · rw [coeff_sub, show (2 * m + 2) ^ 2 = 4 * m * (m + 2) + 4 by ring1, coeff_X_mul, h2.right, + show _ + _ = (2 * (m + 1) * (m + 2)) + (2 * m * (m + 1)) by ring1, + coeff_mul_of_natDegree_le h3.left h1.left, h3.right, h1.right] + push_cast + ring1 + · rcases Nat.odd_iff_not_even.mpr hn with ⟨m, rfl⟩ + rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.add_sub_cancel, + Nat.mul_div_cancel_left _ two_pos, show (2 * m + 2) / 2 = 2 * (m + 1) / 2 by rfl, + Nat.mul_div_cancel_left _ two_pos] at h1 + rw [Int.natAbs_cast, show (2 * m + 3) ^ 2 = 4 * (m + 1) * (m + 2) + 1 by ring1, + Nat.add_sub_cancel] at h2 + rw [if_neg hn, show (2 * m + 4) ^ 2 = 2 * (2 * (m + 1) * (m + 3)) + 4 by ring1, + Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos, + show (2 * m + 4) / 2 = 2 * (m + 2) / 2 by rfl, Nat.mul_div_cancel_left _ two_pos] at h3 + rw [if_neg hn] + constructor + · convert natDegree_sub_le_of_le (natDegree_mul_le_of_le natDegree_X_le h2.left) <| + natDegree_mul_le_of_le (natDegree_mul_le_of_le h3.left h1.left) + W.natDegree_divisionPolynomial2Sq using 1 + convert (max_self _).symm using 2 <;> ring1 + · rw [coeff_sub, show (2 * m + 3) ^ 2 = 4 * (m + 1) * (m + 2) + 1 by ring1, coeff_X_mul, + h2.right, show _ + _ = (2 * (m + 1) * (m + 3)) + (2 * m * (m + 2)) + 3 by ring1, + coeff_mul_of_natDegree_le (natDegree_mul_le_of_le h3.left h1.left) + W.natDegree_divisionPolynomial2Sq, coeff_mul_of_natDegree_le h3.left h1.left, h3.right, + h1.right, coeff_divisionPolynomial2Sq] + push_cast + ring1 + +private lemma natDegree_coeff_divisionPolynomialX (n : ℤ) : + (W.divisionPolynomialX n).natDegree ≤ n.natAbs ^ 2 ∧ + (W.divisionPolynomialX n).coeff (n.natAbs ^ 2) = 1 := by + rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ + · exact W.natDegree_coeff_divisionPolynomialX_ofNat n + · rw [divisionPolynomialX_neg, Int.natAbs_neg] + exact W.natDegree_coeff_divisionPolynomialX_ofNat n + +lemma natDegree_divisionPolynomialX (n : ℤ) : (W.divisionPolynomialX n).natDegree ≤ n.natAbs ^ 2 := + (W.natDegree_coeff_divisionPolynomialX n).left + +lemma coeff_divisionPolynomialX (n : ℤ) : (W.divisionPolynomialX n).coeff (n.natAbs ^ 2) = 1 := + (W.natDegree_coeff_divisionPolynomialX n).right + +/-! ### The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$ -/ + /-- The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ @[pp_dot] noncomputable def divisionPolynomialZ (n : ℤ) : R[X][Y] := From 863e6977a8631e82a30e6d92ff5a801624e42329 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 22 May 2024 21:18:11 +0100 Subject: [PATCH 22/47] Fix merge --- .../EllipticCurve/DivisionPolynomial.lean | 192 +++++++++--------- 1 file changed, 97 insertions(+), 95 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index 29db4273f88f7b..e79339bf2586fc 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -89,7 +89,6 @@ variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) /-- The polynomial $4X^3 + b_2X^2 + 2b_4X + b_6$ congruent to the square $\psi_2^2$ of the $2$-division polynomial $\psi_2 = 2Y + a_1X + a_3$ under $R[W]$. -/ -@[pp_dot] noncomputable def divisionPolynomial2Sq : R[X] := C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ @@ -110,7 +109,6 @@ lemma coeff_divisionPolynomial2Sq : W.divisionPolynomial2Sq.coeff 3 = 4 := by /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials $\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{N}$. -/ -@[pp_dot] noncomputable def divisionPolynomial' (n : ℕ) : R[X] := preNormEDS' (W.divisionPolynomial2Sq ^ 2) (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) @@ -154,6 +152,54 @@ lemma divisionPolynomial'_even (m : ℕ) : W.divisionPolynomial' (2 * (m + 3)) = W.divisionPolynomial' (m + 4) ^ 2 := preNormEDS'_even .. +private lemma natDegree_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : + (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).natDegree ≤ n := by + nth_rw 1 [← max_self n, habc, hdef] + convert natDegree_sub_le_of_le .. using 1 + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| + natDegree_pow_le_of_le 2 hc + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| + natDegree_pow_le_of_le 2 hf + +private lemma coeff_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : + (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).coeff n = a.coeff da * b.coeff db ^ 3 * c.coeff dc ^ 2 - + d.coeff dd * e.coeff de ^ 3 * f.coeff df ^ 2 := by + rw [coeff_sub, habc, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| natDegree_pow_le_of_le 2 hc, + coeff_pow_of_natDegree_le hc, coeff_mul_of_natDegree_le ha <| natDegree_pow_le_of_le 3 hb, + coeff_pow_of_natDegree_le hb, ← habc, hdef, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| natDegree_pow_le_of_le 2 hf, + coeff_mul_of_natDegree_le hd <| natDegree_pow_le_of_le 3 he, coeff_pow_of_natDegree_le he, + coeff_pow_of_natDegree_le hf] + +private lemma natDegree_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : + (a ^ 2 * b * c - d * e * f ^ 2).natDegree ≤ n := by + nth_rw 1 [← max_self n, habc, hdef] + convert natDegree_sub_le_of_le .. using 1 + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc + · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd he) <| natDegree_pow_le_of_le 2 hf + +private lemma coeff_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} + (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) + (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) + (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : + (a ^ 2 * b * c - d * e * f ^ 2).coeff n = + a.coeff da ^ 2 * b.coeff db * c.coeff dc - d.coeff dd * e.coeff de * f.coeff df ^ 2 := by + rw [coeff_sub, habc, coeff_mul_of_natDegree_le + (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc, + coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 ha) hb, coeff_pow_of_natDegree_le ha, + ← habc, hdef, coeff_mul_of_natDegree_le (natDegree_mul_le_of_le hd he) <| + natDegree_pow_le_of_le 2 hf, coeff_mul_of_natDegree_le hd he, coeff_pow_of_natDegree_le hf] + private lemma natDegree_divisionPolynomial'_zero : (W.divisionPolynomial' 0).natDegree = 0 := by rw [divisionPolynomial'_zero, natDegree_zero] @@ -189,40 +235,51 @@ private lemma coeff_divisionPolynomial'_four : (W.divisionPolynomial' 4).coeff 6 compute_degree! private lemma natDegree_divisionPolynomial'_five : (W.divisionPolynomial' 5).natDegree ≤ 12 := by - rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial2Sq] - simp - compute_degree + rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, if_pos even_zero, if_pos even_zero, + ← @one_pow R[X]] + exact natDegree_odd W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_two.le + W.natDegree_divisionPolynomial2Sq W.natDegree_divisionPolynomial'_one.le + W.natDegree_divisionPolynomial'_three natDegree_one.le rfl rfl private lemma coeff_divisionPolynomial'_five : (W.divisionPolynomial' 5).coeff 12 = 5 := by - rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial2Sq] - simp [-coeff_sub] - compute_degree! + rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, if_pos even_zero, if_pos even_zero, + ← @one_pow R[X], coeff_odd W.natDegree_divisionPolynomial'_four + W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial2Sq + W.natDegree_divisionPolynomial'_one.le W.natDegree_divisionPolynomial'_three natDegree_one.le + rfl rfl, coeff_divisionPolynomial'_four, coeff_divisionPolynomial'_two, + coeff_divisionPolynomial2Sq, coeff_divisionPolynomial'_one, coeff_divisionPolynomial'_three, + coeff_one_zero] + norm_num1 private lemma natDegree_divisionPolynomial'_six : (W.divisionPolynomial' 6).natDegree ≤ 16 := by - rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, - divisionPolynomial'_odd, divisionPolynomial2Sq] - simp - compute_degree + rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even] + exact natDegree_even W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial'_three + W.natDegree_divisionPolynomial'_five W.natDegree_divisionPolynomial'_one.le + W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four rfl rfl private lemma coeff_divisionPolynomial'_six : (W.divisionPolynomial' 6).coeff 16 = 3 := by - rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, - divisionPolynomial'_odd, divisionPolynomial2Sq] - simp [-coeff_sub] - compute_degree! + rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, coeff_even + W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial'_three + W.natDegree_divisionPolynomial'_five W.natDegree_divisionPolynomial'_one.le + W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four rfl rfl, + coeff_divisionPolynomial'_two, coeff_divisionPolynomial'_three, coeff_divisionPolynomial'_five, + coeff_divisionPolynomial'_one, coeff_divisionPolynomial'_four] + norm_num1 private lemma natDegree_divisionPolynomial'_eight : (W.divisionPolynomial' 8).natDegree ≤ 30 := by - rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, show 1 + 5 = 2 * 3 by rfl, - divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, - divisionPolynomial2Sq] - simp - compute_degree + rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even] + exact natDegree_even W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four + W.natDegree_divisionPolynomial'_six W.natDegree_divisionPolynomial'_two.le + W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_five rfl rfl private lemma coeff_divisionPolynomial'_eight : (W.divisionPolynomial' 8).coeff 30 = 4 := by - rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, show 1 + 5 = 2 * 3 by rfl, - divisionPolynomial'_even, show 0 + 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, - divisionPolynomial2Sq] - simp [-coeff_sub] - compute_degree! + rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, coeff_even + W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four + W.natDegree_divisionPolynomial'_six W.natDegree_divisionPolynomial'_two.le + W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_five rfl rfl, + coeff_divisionPolynomial'_three, coeff_divisionPolynomial'_four, coeff_divisionPolynomial'_six, + coeff_divisionPolynomial'_two, coeff_divisionPolynomial'_five] + norm_num1 section Inductive @@ -281,7 +338,6 @@ private lemma natDegree_divisionPolynomial'_add_five : rcases m with _ | m · exact W.natDegree_divisionPolynomial'_five · convert (ih (m + 3) <| by linarith only).right.left using 1 - rw [← Nat.add_one] ring1 private lemma coeff_divisionPolynomial'_add_five : @@ -289,7 +345,6 @@ private lemma coeff_divisionPolynomial'_add_five : rcases m with _ | m · exact W.coeff_divisionPolynomial'_five · convert (ih (m + 3) <| by linarith only).right.right using 2 - rw [← Nat.add_one] ring1 private lemma natDegree_divisionPolynomial'_add_six : @@ -308,22 +363,9 @@ private lemma coeff_divisionPolynomial'_add_six : · convert (ih (m + 5) <| by linarith only).left.right using 2 rw [show 2 * (m + 5) ^ 2 = 2 * (m + 4) * (m + 6) + 2 by ring1, Nat.add_sub_cancel] -private lemma natDegree_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : - (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).natDegree ≤ n := by - nth_rw 1 [← max_self n, habc, hdef] - convert natDegree_sub_le_of_le .. using 1 - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| - natDegree_pow_le_of_le 2 hc - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| - natDegree_pow_le_of_le 2 hf - private lemma natDegree_odd_divisionPolynomial'_even : (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 3) := by - rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, - ← @one_pow R[X] _ 2] + rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, ← @one_pow R[X]] exact natDegree_odd (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) @@ -331,32 +373,18 @@ private lemma natDegree_odd_divisionPolynomial'_even : private lemma natDegree_odd_divisionPolynomial'_odd : (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).natDegree ≤ 2 * (2 * m + 3) * (2 * m + 4) := by - rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X] _ 2, + rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X], if_neg m.not_even_two_mul_add_one] exact natDegree_odd (natDegree_divisionPolynomial'_add_five m ih) (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) W.natDegree_divisionPolynomial2Sq (by ring1) (by ring1) -private lemma coeff_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : - (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).coeff n = a.coeff da * b.coeff db ^ 3 * c.coeff dc ^ 2 - - d.coeff dd * e.coeff de ^ 3 * f.coeff df ^ 2 := by - rw [coeff_sub, habc, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| natDegree_pow_le_of_le 2 hc, - coeff_pow_of_natDegree_le hc, coeff_mul_of_natDegree_le ha <| natDegree_pow_le_of_le 3 hb, - coeff_pow_of_natDegree_le hb, ← habc, hdef, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| natDegree_pow_le_of_le 2 hf, - coeff_mul_of_natDegree_le hd <| natDegree_pow_le_of_le 3 he, coeff_pow_of_natDegree_le he, - coeff_pow_of_natDegree_le hf] - private lemma coeff_odd_divisionPolynomial'_even : (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).coeff (2 * (2 * m + 2) * (2 * m + 3)) = (2 * (2 * m + 2) + 1 : ℕ) := by - rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, - ← @one_pow R[X] _ 2, coeff_odd (natDegree_divisionPolynomial'_add_four m ih) + rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, ← @one_pow R[X], + coeff_odd (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le (by ring1) (by ring1), coeff_divisionPolynomial'_add_four m ih, @@ -369,7 +397,7 @@ private lemma coeff_odd_divisionPolynomial'_even : private lemma coeff_odd_divisionPolynomial'_odd : (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).coeff (2 * (2 * m + 3) * (2 * m + 4)) = (2 * (2 * m + 3) + 1 : ℕ) := by - rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X] _ 2, + rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X], if_neg m.not_even_two_mul_add_one, coeff_odd (natDegree_divisionPolynomial'_add_five m ih) (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) @@ -380,16 +408,6 @@ private lemma coeff_odd_divisionPolynomial'_odd : push_cast ring1 -private lemma natDegree_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : - (a ^ 2 * b * c - d * e * f ^ 2).natDegree ≤ n := by - nth_rw 1 [← max_self n, habc, hdef] - convert natDegree_sub_le_of_le .. using 1 - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd he) <| natDegree_pow_le_of_le 2 hf - private lemma natDegree_even_divisionPolynomial'_even : (W.divisionPolynomial' <| 2 * (2 * m + 3)).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 4) := by rw [divisionPolynomial'_even] @@ -406,18 +424,6 @@ private lemma natDegree_even_divisionPolynomial'_odd : (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_five m ih) (by ring1) (by ring1) -private lemma coeff_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : - (a ^ 2 * b * c - d * e * f ^ 2).coeff n = - a.coeff da ^ 2 * b.coeff db * c.coeff dc - d.coeff dd * e.coeff de * f.coeff df ^ 2 := by - rw [coeff_sub, habc, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc, - coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 ha) hb, coeff_pow_of_natDegree_le ha, - ← habc, hdef, coeff_mul_of_natDegree_le (natDegree_mul_le_of_le hd he) <| - natDegree_pow_le_of_le 2 hf, coeff_mul_of_natDegree_le hd he, coeff_pow_of_natDegree_le hf] - private lemma coeff_even_divisionPolynomial'_even : (W.divisionPolynomial' <| 2 * (2 * m + 3)).coeff (2 * (2 * m + 2) * (2 * m + 4)) = (2 * m + 3 : ℕ) := by @@ -478,7 +484,7 @@ private lemma natDegree_coeff_divisionPolynomial' (n : ℕ) : show (2 * (m + 2) + 1) ^ 2 = 2 * (2 * (m + 2) * (m + 3)) + 1 by ring1, Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos] by_cases hm : Even m - · rcases (even_iff_exists_two_mul m).mp hm with ⟨m, rfl⟩ + · rcases even_iff_exists_two_mul.mp hm with ⟨m, rfl⟩ simp only [natDegree_odd_divisionPolynomial'_even m ih, coeff_odd_divisionPolynomial'_even m ih, natDegree_even_divisionPolynomial'_even m ih, coeff_even_divisionPolynomial'_even m ih, and_self] @@ -505,7 +511,6 @@ lemma coeff_divisionPolynomial' (n : ℕ) : /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials $\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -@[pp_dot] noncomputable def divisionPolynomial (n : ℤ) : R[X] := n.sign * W.divisionPolynomial' n.natAbs @@ -561,11 +566,11 @@ private lemma natDegree_coeff_divisionPolynomial (n : ℤ) : exact W.divisionPolynomial_ofNat n ▸ W.natDegree_coeff_divisionPolynomial' n · simp only [even_neg, Int.even_coe_nat, divisionPolynomial_neg, natDegree_neg, divisionPolynomial_ofNat, Int.natAbs_neg, Int.natAbs_cast, coeff_neg, neg_eq_iff_eq_neg, - Int.cast_neg, Int.cast_ofNat, neg_neg] + Int.cast_neg, Int.cast_natCast, neg_neg] convert W.natDegree_coeff_divisionPolynomial' n using 3 with hn - rcases (even_iff_exists_two_mul _).mp hn with ⟨_, rfl⟩ + rcases even_iff_exists_two_mul.mp hn with ⟨_, rfl⟩ rw [Nat.cast_mul, neg_mul_eq_mul_neg, Nat.cast_ofNat, Int.mul_ediv_cancel_left _ two_ne_zero, - Int.cast_neg, neg_neg, Int.cast_ofNat, Nat.mul_div_cancel_left _ two_pos] + Int.cast_neg, neg_neg, Int.cast_natCast, Nat.mul_div_cancel_left _ two_pos] lemma natDegree_divisionPolynomial (n : ℤ) : (W.divisionPolynomial n).natDegree ≤ if Even n then (n.natAbs ^ 2 - 4) / 2 else (n.natAbs ^ 2 - 1) / 2 := by @@ -584,7 +589,6 @@ lemma coeff_divisionPolynomial (n : ℤ) : /-- The univariate polynomials congruent under $R[W]$ to the bivariate squares $\psi_n^2$ of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -@[pp_dot] noncomputable def divisionPolynomialZSq (n : ℤ) : R[X] := W.divisionPolynomial n ^ 2 * if Even n.natAbs then W.divisionPolynomial2Sq else 1 @@ -629,7 +633,7 @@ private lemma natDegree_coeff_divisionPolynomialZSq_ofNat (n : ℕ) : simp only [divisionPolynomialZSq, divisionPolynomial_ofNat, Int.natAbs_even, Int.even_coe_nat, Nat.even_add_one, ite_not] at h ⊢ by_cases hn : Even n - · rcases (even_iff_exists_two_mul n).mp hn with ⟨m, rfl⟩ + · rcases even_iff_exists_two_mul.mp hn with ⟨m, rfl⟩ rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos] at h rw [if_pos hn, mul_one, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, @@ -673,7 +677,6 @@ lemma coeff_divisionPolynomialZSq (n : ℤ) : /-- The univariate polynomials congruent under $R[W]$ to the bivariate polynomials $\phi_n$ defined in terms of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -@[pp_dot] noncomputable def divisionPolynomialX (n : ℤ) : R[X] := X * W.divisionPolynomialZSq n - W.divisionPolynomial (n + 1) * W.divisionPolynomial (n - 1) * if Even n.natAbs then 1 else W.divisionPolynomial2Sq @@ -685,7 +688,7 @@ lemma divisionPolynomialX_ofNat (n : ℕ) : W.divisionPolynomialX (n + 1) = W.divisionPolynomial' (n + 2) * W.divisionPolynomial' n * (if Even n then W.divisionPolynomial2Sq else 1) := by erw [divisionPolynomialX, divisionPolynomialZSq_ofNat, ← mul_assoc, divisionPolynomial_ofNat, - add_sub_cancel, divisionPolynomial_ofNat, Int.natAbs_cast] + add_sub_cancel_right, divisionPolynomial_ofNat, Int.natAbs_cast] simp only [Nat.even_add_one, ite_not] @[simp] @@ -737,10 +740,10 @@ private lemma natDegree_coeff_divisionPolynomialX_ofNat (n : ℕ) : have h2 := W.natDegree_coeff_divisionPolynomialZSq (n + 2 : ℕ) have h3 := W.natDegree_coeff_divisionPolynomial' <| n + 3 simp only [divisionPolynomialX, Int.natAbs_cast, Nat.even_add_one, ite_not] at h1 h3 ⊢ - erw [← Nat.cast_add, divisionPolynomial_ofNat, ← Nat.cast_sub <| by linarith only, - divisionPolynomial_ofNat, ← Nat.add_one, Nat.add_sub_cancel] + erw [divisionPolynomial_ofNat, ← Nat.cast_sub <| by linarith only, divisionPolynomial_ofNat, + Nat.add_sub_cancel] by_cases hn : Even n - · rcases (even_iff_exists_two_mul n).mp hn with ⟨m, rfl⟩ + · rcases even_iff_exists_two_mul.mp hn with ⟨m, rfl⟩ rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos] at h1 rw [Int.natAbs_cast, show (2 * m + 2) ^ 2 = 4 * m * (m + 2) + 4 by ring1, Nat.succ_sub_one] @@ -797,7 +800,6 @@ lemma coeff_divisionPolynomialX (n : ℤ) : (W.divisionPolynomialX n).coeff (n.n /-! ### The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$ -/ /-- The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -@[pp_dot] noncomputable def divisionPolynomialZ (n : ℤ) : R[X][Y] := C (W.divisionPolynomial n) * if Even n.natAbs then Y - W.toAffine.negPolynomial else 1 From 6a76f32db87ebc423ac10a87e7b85aa11e03cb2e Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 23 May 2024 17:42:05 +0100 Subject: [PATCH 23/47] Fix upstream --- .../EllipticCurve/DivisionPolynomial.lean | 523 ------------------ 1 file changed, 523 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index e79339bf2586fc..eece929fadf709 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -6,7 +6,6 @@ Authors: David Kurniadi Angdinata import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.Data.Int.Parity import Mathlib.NumberTheory.EllipticDivisibilitySequence -import Mathlib.Tactic.ComputeDegree /-! # Division polynomials of Weierstrass curves @@ -44,17 +43,6 @@ if $n$ is even or $\tilde{\psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ otherwise. * `WeierstrassCurve.divisionPolynomialZ`: $\psi_n$ viewed as a bivariate polynomial. * TODO: $\omega$ viewed as a bivariate polynomial. -## Main statements - - * `WeierstrassCurve.natDegree_divisionPolynomial2Sq`: the degree of $\psi_2^2$. - * `WeierstrassCurve.coeff_divisionPolynomial2Sq`: the leading coefficient of $\psi_2^2$. - * `WeierstrassCurve.natDegree_divisionPolynomial`: the degree of $\tilde{\psi}_n$. - * `WeierstrassCurve.coeff_divisionPolynomial`: the leading coefficient of $\tilde{\psi}_n$. - * `WeierstrassCurve.natDegree_divisionPolynomialX`: the degree of $\phi_n$. - * `WeierstrassCurve.coeff_divisionPolynomialX`: the leading coefficient of $\phi_n$. - * `WeierstrassCurve.natDegree_divisionPolynomialZSq`: the degree of $\psi_n^2$. - * `WeierstrassCurve.coeff_divisionPolynomialZSq`: the leading coefficient of $\psi_n^2$. - ## Implementation notes Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate @@ -96,15 +84,6 @@ noncomputable def divisionPolynomial2Sq : R[X] := lemma divisionPolynomial2Sq_eq : W.divisionPolynomial2Sq = W.twoTorsionPolynomial.toPoly := rfl -lemma natDegree_divisionPolynomial2Sq : W.divisionPolynomial2Sq.natDegree ≤ 3 := by - rw [divisionPolynomial2Sq] - compute_degree - -@[simp] -lemma coeff_divisionPolynomial2Sq : W.divisionPolynomial2Sq.coeff 3 = 4 := by - rw [divisionPolynomial2Sq] - compute_degree! - /-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{N}$ -/ /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials @@ -152,361 +131,6 @@ lemma divisionPolynomial'_even (m : ℕ) : W.divisionPolynomial' (2 * (m + 3)) = W.divisionPolynomial' (m + 4) ^ 2 := preNormEDS'_even .. -private lemma natDegree_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : - (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).natDegree ≤ n := by - nth_rw 1 [← max_self n, habc, hdef] - convert natDegree_sub_le_of_le .. using 1 - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| - natDegree_pow_le_of_le 2 hc - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| - natDegree_pow_le_of_le 2 hf - -private lemma coeff_odd {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = da + 3 * db + 2 * dc) (hdef : n = dd + 3 * de + 2 * df) : - (a * b ^ 3 * c ^ 2 - d * e ^ 3 * f ^ 2).coeff n = a.coeff da * b.coeff db ^ 3 * c.coeff dc ^ 2 - - d.coeff dd * e.coeff de ^ 3 * f.coeff df ^ 2 := by - rw [coeff_sub, habc, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le ha <| natDegree_pow_le_of_le 3 hb) <| natDegree_pow_le_of_le 2 hc, - coeff_pow_of_natDegree_le hc, coeff_mul_of_natDegree_le ha <| natDegree_pow_le_of_le 3 hb, - coeff_pow_of_natDegree_le hb, ← habc, hdef, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le hd <| natDegree_pow_le_of_le 3 he) <| natDegree_pow_le_of_le 2 hf, - coeff_mul_of_natDegree_le hd <| natDegree_pow_le_of_le 3 he, coeff_pow_of_natDegree_le he, - coeff_pow_of_natDegree_le hf] - -private lemma natDegree_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : - (a ^ 2 * b * c - d * e * f ^ 2).natDegree ≤ n := by - nth_rw 1 [← max_self n, habc, hdef] - convert natDegree_sub_le_of_le .. using 1 - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc - · exact natDegree_mul_le_of_le (natDegree_mul_le_of_le hd he) <| natDegree_pow_le_of_le 2 hf - -private lemma coeff_even {a b c d e f : R[X]} {da db dc dd de df n : ℕ} - (ha : a.natDegree ≤ da) (hb : b.natDegree ≤ db) (hc : c.natDegree ≤ dc) - (hd : d.natDegree ≤ dd) (he : e.natDegree ≤ de) (hf : f.natDegree ≤ df) - (habc : n = 2 * da + db + dc) (hdef : n = dd + de + 2 * df) : - (a ^ 2 * b * c - d * e * f ^ 2).coeff n = - a.coeff da ^ 2 * b.coeff db * c.coeff dc - d.coeff dd * e.coeff de * f.coeff df ^ 2 := by - rw [coeff_sub, habc, coeff_mul_of_natDegree_le - (natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 ha) hb) hc, - coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 ha) hb, coeff_pow_of_natDegree_le ha, - ← habc, hdef, coeff_mul_of_natDegree_le (natDegree_mul_le_of_le hd he) <| - natDegree_pow_le_of_le 2 hf, coeff_mul_of_natDegree_le hd he, coeff_pow_of_natDegree_le hf] - -private lemma natDegree_divisionPolynomial'_zero : (W.divisionPolynomial' 0).natDegree = 0 := by - rw [divisionPolynomial'_zero, natDegree_zero] - -private lemma coeff_divisionPolynomial'_zero : (W.divisionPolynomial' 0).coeff 0 = 0 := by - rw [divisionPolynomial'_zero, coeff_zero] - -private lemma natDegree_divisionPolynomial'_one : (W.divisionPolynomial' 1).natDegree = 0 := by - rw [divisionPolynomial'_one, natDegree_one] - -private lemma coeff_divisionPolynomial'_one : (W.divisionPolynomial' 1).coeff 0 = 1 := by - rw [divisionPolynomial'_one, coeff_one_zero] - -private lemma natDegree_divisionPolynomial'_two : (W.divisionPolynomial' 2).natDegree = 0 := by - rw [divisionPolynomial'_two, natDegree_one] - -private lemma coeff_divisionPolynomial'_two : (W.divisionPolynomial' 2).coeff 0 = 1 := by - rw [divisionPolynomial'_two, coeff_one_zero] - -private lemma natDegree_divisionPolynomial'_three : (W.divisionPolynomial' 3).natDegree ≤ 4 := by - rw [divisionPolynomial'_three] - compute_degree - -private lemma coeff_divisionPolynomial'_three : (W.divisionPolynomial' 3).coeff 4 = 3 := by - rw [divisionPolynomial'_three] - compute_degree! - -private lemma natDegree_divisionPolynomial'_four : (W.divisionPolynomial' 4).natDegree ≤ 6 := by - rw [divisionPolynomial'_four] - compute_degree - -private lemma coeff_divisionPolynomial'_four : (W.divisionPolynomial' 4).coeff 6 = 2 := by - rw [divisionPolynomial'_four] - compute_degree! - -private lemma natDegree_divisionPolynomial'_five : (W.divisionPolynomial' 5).natDegree ≤ 12 := by - rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, if_pos even_zero, if_pos even_zero, - ← @one_pow R[X]] - exact natDegree_odd W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_two.le - W.natDegree_divisionPolynomial2Sq W.natDegree_divisionPolynomial'_one.le - W.natDegree_divisionPolynomial'_three natDegree_one.le rfl rfl - -private lemma coeff_divisionPolynomial'_five : (W.divisionPolynomial' 5).coeff 12 = 5 := by - rw [show 5 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, if_pos even_zero, if_pos even_zero, - ← @one_pow R[X], coeff_odd W.natDegree_divisionPolynomial'_four - W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial2Sq - W.natDegree_divisionPolynomial'_one.le W.natDegree_divisionPolynomial'_three natDegree_one.le - rfl rfl, coeff_divisionPolynomial'_four, coeff_divisionPolynomial'_two, - coeff_divisionPolynomial2Sq, coeff_divisionPolynomial'_one, coeff_divisionPolynomial'_three, - coeff_one_zero] - norm_num1 - -private lemma natDegree_divisionPolynomial'_six : (W.divisionPolynomial' 6).natDegree ≤ 16 := by - rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even] - exact natDegree_even W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial'_three - W.natDegree_divisionPolynomial'_five W.natDegree_divisionPolynomial'_one.le - W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four rfl rfl - -private lemma coeff_divisionPolynomial'_six : (W.divisionPolynomial' 6).coeff 16 = 3 := by - rw [show 6 = 2 * 3 by rfl, divisionPolynomial'_even, coeff_even - W.natDegree_divisionPolynomial'_two.le W.natDegree_divisionPolynomial'_three - W.natDegree_divisionPolynomial'_five W.natDegree_divisionPolynomial'_one.le - W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four rfl rfl, - coeff_divisionPolynomial'_two, coeff_divisionPolynomial'_three, coeff_divisionPolynomial'_five, - coeff_divisionPolynomial'_one, coeff_divisionPolynomial'_four] - norm_num1 - -private lemma natDegree_divisionPolynomial'_eight : (W.divisionPolynomial' 8).natDegree ≤ 30 := by - rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even] - exact natDegree_even W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four - W.natDegree_divisionPolynomial'_six W.natDegree_divisionPolynomial'_two.le - W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_five rfl rfl - -private lemma coeff_divisionPolynomial'_eight : (W.divisionPolynomial' 8).coeff 30 = 4 := by - rw [show 8 = 2 * 4 by rfl, divisionPolynomial'_even, coeff_even - W.natDegree_divisionPolynomial'_three W.natDegree_divisionPolynomial'_four - W.natDegree_divisionPolynomial'_six W.natDegree_divisionPolynomial'_two.le - W.natDegree_divisionPolynomial'_four W.natDegree_divisionPolynomial'_five rfl rfl, - coeff_divisionPolynomial'_three, coeff_divisionPolynomial'_four, coeff_divisionPolynomial'_six, - coeff_divisionPolynomial'_two, coeff_divisionPolynomial'_five] - norm_num1 - -section Inductive - -variable {W} (m : ℕ) (ih : ∀ k < 2 * m + 2, - ((W.divisionPolynomial' <| 2 * k).natDegree ≤ 2 * k ^ 2 - 2 ∧ - (W.divisionPolynomial' <| 2 * k).coeff (2 * k ^ 2 - 2) = k) ∧ - ((W.divisionPolynomial' <| 2 * k + 1).natDegree ≤ 2 * k ^ 2 + 2 * k ∧ - (W.divisionPolynomial' <| 2 * k + 1).coeff (2 * k ^ 2 + 2 * k) = (2 * k + 1 : ℕ))) - -private lemma natDegree_divisionPolynomial'_add_one : - (W.divisionPolynomial' <| 2 * m + 1).natDegree ≤ 2 * m * (m + 1) := by - convert (ih m <| by linarith only).right.left using 1 - ring1 - -private lemma coeff_divisionPolynomial'_add_one : - (W.divisionPolynomial' <| 2 * m + 1).coeff (2 * m * (m + 1)) = (2 * m + 1 : ℕ) := by - convert (ih m <| by linarith only).right.right using 2 - ring1 - -private lemma natDegree_divisionPolynomial'_add_two : - (W.divisionPolynomial' <| 2 * m + 2).natDegree ≤ 2 * m * (m + 2) := by - convert (ih (m + 1) <| by linarith only).left.left using 1 - rw [show 2 * (m + 1) ^ 2 = 2 * m * (m + 2) + 2 by ring1, Nat.add_sub_cancel] - -private lemma coeff_divisionPolynomial'_add_two : - (W.divisionPolynomial' <| 2 * m + 2).coeff (2 * m * (m + 2)) = (m + 1 : ℕ) := by - convert (ih (m + 1) <| by linarith only).left.right using 2 - rw [show 2 * (m + 1) ^ 2 = 2 * m * (m + 2) + 2 by ring1, Nat.add_sub_cancel] - -private lemma natDegree_divisionPolynomial'_add_three : - (W.divisionPolynomial' <| 2 * m + 3).natDegree ≤ 2 * (m + 1) * (m + 2) := by - convert (ih (m + 1) <| by linarith only).right.left using 1 - ring1 - -private lemma coeff_divisionPolynomial'_add_three : - (W.divisionPolynomial' <| 2 * m + 3).coeff (2 * (m + 1) * (m + 2)) = (2 * m + 3 : ℕ) := by - convert (ih (m + 1) <| by linarith only).right.right using 2 - ring1 - -private lemma natDegree_divisionPolynomial'_add_four : - (W.divisionPolynomial' <| 2 * m + 4).natDegree ≤ 2 * (m + 1) * (m + 3) := by - rcases m with _ | m - · exact W.natDegree_divisionPolynomial'_four - · convert (ih (m + 3) <| by linarith only).left.left using 1 - rw [show 2 * (m + 3) ^ 2 = 2 * (m + 2) * (m + 4) + 2 by ring1, Nat.add_sub_cancel] - -private lemma coeff_divisionPolynomial'_add_four : - (W.divisionPolynomial' <| 2 * m + 4).coeff (2 * (m + 1) * (m + 3)) = (m + 2 : ℕ) := by - rcases m with _ | m - · exact W.coeff_divisionPolynomial'_four - · convert (ih (m + 3) <| by linarith only).left.right using 2 - rw [show 2 * (m + 3) ^ 2 = 2 * (m + 2) * (m + 4) + 2 by ring1, Nat.add_sub_cancel] - -private lemma natDegree_divisionPolynomial'_add_five : - (W.divisionPolynomial' <| 2 * m + 5).natDegree ≤ 2 * (m + 2) * (m + 3) := by - rcases m with _ | m - · exact W.natDegree_divisionPolynomial'_five - · convert (ih (m + 3) <| by linarith only).right.left using 1 - ring1 - -private lemma coeff_divisionPolynomial'_add_five : - (W.divisionPolynomial' <| 2 * m + 5).coeff (2 * (m + 2) * (m + 3)) = (2 * m + 5 : ℕ) := by - rcases m with _ | m - · exact W.coeff_divisionPolynomial'_five - · convert (ih (m + 3) <| by linarith only).right.right using 2 - ring1 - -private lemma natDegree_divisionPolynomial'_add_six : - (W.divisionPolynomial' <| 2 * m + 6).natDegree ≤ 2 * (m + 2) * (m + 4) := by - rcases m with _ | _ | m - · exact W.natDegree_divisionPolynomial'_six - · exact W.natDegree_divisionPolynomial'_eight - · convert (ih (m + 5) <| by linarith only).left.left using 1 - rw [show 2 * (m + 5) ^ 2 = 2 * (m + 4) * (m + 6) + 2 by ring1, Nat.add_sub_cancel] - -private lemma coeff_divisionPolynomial'_add_six : - (W.divisionPolynomial' <| 2 * m + 6).coeff (2 * (m + 2) * (m + 4)) = (m + 3 : ℕ) := by - rcases m with _ | _ | m - · exact W.coeff_divisionPolynomial'_six - · exact W.coeff_divisionPolynomial'_eight - · convert (ih (m + 5) <| by linarith only).left.right using 2 - rw [show 2 * (m + 5) ^ 2 = 2 * (m + 4) * (m + 6) + 2 by ring1, Nat.add_sub_cancel] - -private lemma natDegree_odd_divisionPolynomial'_even : - (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 3) := by - rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, ← @one_pow R[X]] - exact natDegree_odd (natDegree_divisionPolynomial'_add_four m ih) - (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq - (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) - natDegree_one.le (by ring1) (by ring1) - -private lemma natDegree_odd_divisionPolynomial'_odd : - (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).natDegree ≤ 2 * (2 * m + 3) * (2 * m + 4) := by - rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X], - if_neg m.not_even_two_mul_add_one] - exact natDegree_odd (natDegree_divisionPolynomial'_add_five m ih) - (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le - (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) - W.natDegree_divisionPolynomial2Sq (by ring1) (by ring1) - -private lemma coeff_odd_divisionPolynomial'_even : - (W.divisionPolynomial' <| 2 * (2 * m + 2) + 1).coeff (2 * (2 * m + 2) * (2 * m + 3)) = - (2 * (2 * m + 2) + 1 : ℕ) := by - rw [divisionPolynomial'_odd, if_pos <| even_two_mul m, if_pos <| even_two_mul m, ← @one_pow R[X], - coeff_odd (natDegree_divisionPolynomial'_add_four m ih) - (natDegree_divisionPolynomial'_add_two m ih) W.natDegree_divisionPolynomial2Sq - (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) - natDegree_one.le (by ring1) (by ring1), coeff_divisionPolynomial'_add_four m ih, - coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial2Sq, - coeff_divisionPolynomial'_add_one m ih, coeff_divisionPolynomial'_add_three m ih, - coeff_one_zero] - push_cast - ring1 - -private lemma coeff_odd_divisionPolynomial'_odd : - (W.divisionPolynomial' <| 2 * (2 * m + 3) + 1).coeff (2 * (2 * m + 3) * (2 * m + 4)) = - (2 * (2 * m + 3) + 1 : ℕ) := by - rw [divisionPolynomial'_odd, if_neg m.not_even_two_mul_add_one, ← @one_pow R[X], - if_neg m.not_even_two_mul_add_one, coeff_odd (natDegree_divisionPolynomial'_add_five m ih) - (natDegree_divisionPolynomial'_add_three m ih) natDegree_one.le - (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) - W.natDegree_divisionPolynomial2Sq (by ring1) (by ring1), - coeff_divisionPolynomial'_add_five m ih, coeff_divisionPolynomial'_add_three m ih, - coeff_one_zero, coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial'_add_four m ih, - coeff_divisionPolynomial2Sq] - push_cast - ring1 - -private lemma natDegree_even_divisionPolynomial'_even : - (W.divisionPolynomial' <| 2 * (2 * m + 3)).natDegree ≤ 2 * (2 * m + 2) * (2 * m + 4) := by - rw [divisionPolynomial'_even] - exact natDegree_even (natDegree_divisionPolynomial'_add_two m ih) - (natDegree_divisionPolynomial'_add_three m ih) (natDegree_divisionPolynomial'_add_five m ih) - (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) - (natDegree_divisionPolynomial'_add_four m ih) (by ring1) (by ring1) - -private lemma natDegree_even_divisionPolynomial'_odd : - (W.divisionPolynomial' <| 2 * (2 * m + 4)).natDegree ≤ 2 * (2 * m + 3) * (2 * m + 5) := by - rw [divisionPolynomial'_even] - exact natDegree_even (natDegree_divisionPolynomial'_add_three m ih) - (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_six m ih) - (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) - (natDegree_divisionPolynomial'_add_five m ih) (by ring1) (by ring1) - -private lemma coeff_even_divisionPolynomial'_even : - (W.divisionPolynomial' <| 2 * (2 * m + 3)).coeff (2 * (2 * m + 2) * (2 * m + 4)) = - (2 * m + 3 : ℕ) := by - rw [divisionPolynomial'_even, coeff_even (natDegree_divisionPolynomial'_add_two m ih) - (natDegree_divisionPolynomial'_add_three m ih) (natDegree_divisionPolynomial'_add_five m ih) - (natDegree_divisionPolynomial'_add_one m ih) (natDegree_divisionPolynomial'_add_three m ih) - (natDegree_divisionPolynomial'_add_four m ih) (by ring1) (by ring1), - coeff_divisionPolynomial'_add_two m ih, coeff_divisionPolynomial'_add_three m ih, - coeff_divisionPolynomial'_add_five m ih, coeff_divisionPolynomial'_add_one m ih, - coeff_divisionPolynomial'_add_four m ih] - push_cast - ring1 - -private lemma coeff_even_divisionPolynomial'_odd : - (W.divisionPolynomial' <| 2 * (2 * m + 4)).coeff (2 * (2 * m + 3) * (2 * m + 5)) = - (2 * m + 4 : ℕ) := by - rw [divisionPolynomial'_even, coeff_even (natDegree_divisionPolynomial'_add_three m ih) - (natDegree_divisionPolynomial'_add_four m ih) (natDegree_divisionPolynomial'_add_six m ih) - (natDegree_divisionPolynomial'_add_two m ih) (natDegree_divisionPolynomial'_add_four m ih) - (natDegree_divisionPolynomial'_add_five m ih) (by ring1) (by ring1), - coeff_divisionPolynomial'_add_three m ih, coeff_divisionPolynomial'_add_four m ih, - coeff_divisionPolynomial'_add_six m ih, coeff_divisionPolynomial'_add_two m ih, - coeff_divisionPolynomial'_add_five m ih] - push_cast - ring1 - -end Inductive - -private lemma natDegree_coeff_divisionPolynomial' (n : ℕ) : - if Even n then (W.divisionPolynomial' n).natDegree ≤ (n ^ 2 - 4) / 2 ∧ - (W.divisionPolynomial' n).coeff ((n ^ 2 - 4) / 2) = (n / 2 : ℕ) - else (W.divisionPolynomial' n).natDegree ≤ (n ^ 2 - 1) / 2 ∧ - (W.divisionPolynomial' n).coeff ((n ^ 2 - 1) / 2) = n := by - induction n using normEDSRec' with - | base0 => exact ⟨W.natDegree_divisionPolynomial'_zero.le, - by erw [coeff_divisionPolynomial'_zero, Nat.cast_zero]⟩ - | base1 => exact ⟨W.natDegree_divisionPolynomial'_one.le, - by erw [coeff_divisionPolynomial'_one, Nat.cast_one]⟩ - | base2 => exact ⟨W.natDegree_divisionPolynomial'_two.le, - by erw [coeff_divisionPolynomial'_two, Nat.cast_one]⟩ - | base3 => exact ⟨W.natDegree_divisionPolynomial'_three, W.coeff_divisionPolynomial'_three⟩ - | base4 => exact ⟨W.natDegree_divisionPolynomial'_four, W.coeff_divisionPolynomial'_four⟩ - | _ m ih => - replace ih (k : ℕ) (hk : k < m + 2) : - ((W.divisionPolynomial' <| 2 * k).natDegree ≤ 2 * k ^ 2 - 2 ∧ - (W.divisionPolynomial' <| 2 * k).coeff (2 * k ^ 2 - 2) = k) ∧ - ((W.divisionPolynomial' <| 2 * k + 1).natDegree ≤ 2 * k ^ 2 + 2 * k ∧ - (W.divisionPolynomial' <| 2 * k + 1).coeff (2 * k ^ 2 + 2 * k) = (2 * k + 1 : ℕ)) := by - rw [← show ((2 * k) ^ 2 - 4) / 2 = 2 * k ^ 2 - 2 by - exact Nat.div_eq_of_eq_mul_right two_pos <| by rw [Nat.mul_sub_left_distrib]; ring_nf, - ← show ((2 * k + 1) ^ 2 - 1) / 2 = 2 * k ^ 2 + 2 * k by - exact Nat.div_eq_of_eq_mul_right two_pos <| by erw [add_sq, Nat.add_sub_cancel]; ring1] - nth_rw 5 [← k.mul_div_cancel_left two_pos] - exact ⟨imp_of_if_pos (ih (2 * k) <| by linarith only [hk]) <| even_two_mul k, - imp_of_if_neg (ih (2 * k + 1) <| by linarith only [hk]) k.not_even_two_mul_add_one⟩ - simp only [if_pos <| even_two_mul _, if_neg (m + 2).not_even_two_mul_add_one, - show (2 * (m + 3)) ^ 2 = 2 * (2 * (m + 2) * (m + 4)) + 4 by ring1, - show (2 * (m + 2) + 1) ^ 2 = 2 * (2 * (m + 2) * (m + 3)) + 1 by ring1, Nat.add_sub_cancel, - Nat.mul_div_cancel_left _ two_pos] - by_cases hm : Even m - · rcases even_iff_exists_two_mul.mp hm with ⟨m, rfl⟩ - simp only [natDegree_odd_divisionPolynomial'_even m ih, - coeff_odd_divisionPolynomial'_even m ih, natDegree_even_divisionPolynomial'_even m ih, - coeff_even_divisionPolynomial'_even m ih, and_self] - · replace ih (k : ℕ) (hk : k < m + 1) := ih k <| Nat.lt.step hk - rcases Nat.odd_iff_not_even.mpr hm with ⟨m, rfl⟩ - simp only [natDegree_odd_divisionPolynomial'_odd m ih, coeff_odd_divisionPolynomial'_odd m ih, - natDegree_even_divisionPolynomial'_odd m ih, coeff_even_divisionPolynomial'_odd m ih, - and_self] - -lemma natDegree_divisionPolynomial' (n : ℕ) : (W.divisionPolynomial' n).natDegree ≤ - if Even n then (n ^ 2 - 4) / 2 else (n ^ 2 - 1) / 2 := by - by_cases hn : Even n - · simpa only [if_pos hn] using (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial' n) hn).left - · simpa only [if_neg hn] using (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial' n) hn).left - -lemma coeff_divisionPolynomial' (n : ℕ) : - if Even n then (W.divisionPolynomial' n).coeff ((n ^ 2 - 4) / 2) = (n / 2 : ℕ) - else (W.divisionPolynomial' n).coeff ((n ^ 2 - 1) / 2) = n := by - by_cases hn : Even n - · rw [if_pos hn, (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial' n) hn).right] - · rw [if_neg hn, (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial' n) hn).right] - /-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials @@ -556,35 +180,6 @@ lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) = lemma divisionPolynomial_neg (n : ℤ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := preNormEDS_neg .. -private lemma natDegree_coeff_divisionPolynomial (n : ℤ) : - if Even n then (W.divisionPolynomial n).natDegree ≤ (n.natAbs ^ 2 - 4) / 2 ∧ - (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 4) / 2) = (n / 2 : ℤ) - else (W.divisionPolynomial n).natDegree ≤ (n.natAbs ^ 2 - 1) / 2 ∧ - (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 1) / 2) = n := by - rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ - · norm_cast - exact W.divisionPolynomial_ofNat n ▸ W.natDegree_coeff_divisionPolynomial' n - · simp only [even_neg, Int.even_coe_nat, divisionPolynomial_neg, natDegree_neg, - divisionPolynomial_ofNat, Int.natAbs_neg, Int.natAbs_cast, coeff_neg, neg_eq_iff_eq_neg, - Int.cast_neg, Int.cast_natCast, neg_neg] - convert W.natDegree_coeff_divisionPolynomial' n using 3 with hn - rcases even_iff_exists_two_mul.mp hn with ⟨_, rfl⟩ - rw [Nat.cast_mul, neg_mul_eq_mul_neg, Nat.cast_ofNat, Int.mul_ediv_cancel_left _ two_ne_zero, - Int.cast_neg, neg_neg, Int.cast_natCast, Nat.mul_div_cancel_left _ two_pos] - -lemma natDegree_divisionPolynomial (n : ℤ) : (W.divisionPolynomial n).natDegree ≤ - if Even n then (n.natAbs ^ 2 - 4) / 2 else (n.natAbs ^ 2 - 1) / 2 := by - by_cases hn : Even n - · simpa only [if_pos hn] using (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial n) hn).left - · simpa only [if_neg hn] using (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial n) hn).left - -lemma coeff_divisionPolynomial (n : ℤ) : - if Even n then (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 4) / 2) = (n / 2 : ℤ) - else (W.divisionPolynomial n).coeff ((n.natAbs ^ 2 - 1) / 2) = n := by - by_cases hn : Even n - · rw [if_pos hn, (imp_of_if_pos (W.natDegree_coeff_divisionPolynomial n) hn).right] - · rw [if_neg hn, (imp_of_if_neg (W.natDegree_coeff_divisionPolynomial n) hn).right] - /-! ### The univariate polynomials congruent to $\psi_n^2$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials congruent under $R[W]$ to the bivariate squares $\psi_n^2$ of the @@ -623,56 +218,6 @@ lemma divisionPolynomialZSq_neg (n : ℤ) : W.divisionPolynomialZSq (-n) = W.divisionPolynomialZSq n := by rw [divisionPolynomialZSq, divisionPolynomial_neg, neg_sq, Int.natAbs_neg, divisionPolynomialZSq] -private lemma natDegree_coeff_divisionPolynomialZSq_ofNat (n : ℕ) : - (W.divisionPolynomialZSq n).natDegree ≤ n ^ 2 - 1 ∧ - (W.divisionPolynomialZSq n).coeff (n ^ 2 - 1) = (n ^ 2 : ℕ) := by - rcases n with _ | n - · erw [divisionPolynomialZSq_zero] - exact ⟨natDegree_zero.le, Nat.cast_zero.symm⟩ - · have h := W.natDegree_coeff_divisionPolynomial' <| n + 1 - simp only [divisionPolynomialZSq, divisionPolynomial_ofNat, Int.natAbs_even, Int.even_coe_nat, - Nat.even_add_one, ite_not] at h ⊢ - by_cases hn : Even n - · rcases even_iff_exists_two_mul.mp hn with ⟨m, rfl⟩ - rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, - Nat.mul_div_cancel_left _ two_pos] at h - rw [if_pos hn, mul_one, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, - Nat.add_sub_cancel] - constructor - · exact natDegree_pow_le_of_le 2 h.left - · rw [coeff_pow_of_natDegree_le h.left, h.right] - push_cast - ring1 - · rcases Nat.odd_iff_not_even.mpr hn with ⟨m, rfl⟩ - rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.add_sub_cancel, - Nat.mul_div_cancel_left _ two_pos, show (2 * m + 2) / 2 = 2 * (m + 1) / 2 by rfl, - Nat.mul_div_cancel_left _ two_pos] at h - rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.succ_sub_one] - constructor - · exact natDegree_mul_le_of_le (natDegree_pow_le_of_le 2 h.left) - W.natDegree_divisionPolynomial2Sq - · rw [coeff_mul_of_natDegree_le (natDegree_pow_le_of_le 2 h.left) - W.natDegree_divisionPolynomial2Sq, coeff_pow_of_natDegree_le h.left, h.right, - coeff_divisionPolynomial2Sq] - push_cast - ring1 - -private lemma natDegree_coeff_divisionPolynomialZSq (n : ℤ) : - (W.divisionPolynomialZSq n).natDegree ≤ n.natAbs ^ 2 - 1 ∧ - (W.divisionPolynomialZSq n).coeff (n.natAbs ^ 2 - 1) = (n.natAbs ^ 2 : ℕ) := by - rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ - · exact W.natDegree_coeff_divisionPolynomialZSq_ofNat n - · simp only [divisionPolynomialZSq_neg, Int.natAbs_neg] - exact W.natDegree_coeff_divisionPolynomialZSq_ofNat n - -lemma natDegree_divisionPolynomialZSq (n : ℤ) : - (W.divisionPolynomialZSq n).natDegree ≤ n.natAbs ^ 2 - 1 := - (W.natDegree_coeff_divisionPolynomialZSq n).left - -lemma coeff_divisionPolynomialZSq (n : ℤ) : - (W.divisionPolynomialZSq n).coeff (n.natAbs ^ 2 - 1) = (n.natAbs ^ 2 : ℕ) := - (W.natDegree_coeff_divisionPolynomialZSq n).right - /-! ### The univariate polynomials congruent to $\phi_n$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials congruent under $R[W]$ to the bivariate polynomials $\phi_n$ defined @@ -729,74 +274,6 @@ lemma divisionPolynomialX_neg (n : ℤ) : W.divisionPolynomialX (-n) = W.divisio divisionPolynomial_neg, ← neg_add', divisionPolynomial_neg, neg_mul_neg, mul_comm <| W.divisionPolynomial _, Int.natAbs_neg, divisionPolynomialX] -private lemma natDegree_coeff_divisionPolynomialX_ofNat (n : ℕ) : - (W.divisionPolynomialX n).natDegree ≤ n ^ 2 ∧ (W.divisionPolynomialX n).coeff (n ^ 2) = 1 := by - rcases n with _ | _ | n - · erw [divisionPolynomialX_zero] - exact ⟨natDegree_one.le, coeff_one_zero⟩ - · erw [divisionPolynomialX_one] - exact ⟨natDegree_X_le, coeff_X⟩ - · have h1 := W.natDegree_coeff_divisionPolynomial' <| n + 1 - have h2 := W.natDegree_coeff_divisionPolynomialZSq (n + 2 : ℕ) - have h3 := W.natDegree_coeff_divisionPolynomial' <| n + 3 - simp only [divisionPolynomialX, Int.natAbs_cast, Nat.even_add_one, ite_not] at h1 h3 ⊢ - erw [divisionPolynomial_ofNat, ← Nat.cast_sub <| by linarith only, divisionPolynomial_ofNat, - Nat.add_sub_cancel] - by_cases hn : Even n - · rcases even_iff_exists_two_mul.mp hn with ⟨m, rfl⟩ - rw [if_pos hn, show (2 * m + 1) ^ 2 = 2 * (2 * m * (m + 1)) + 1 by ring1, Nat.add_sub_cancel, - Nat.mul_div_cancel_left _ two_pos] at h1 - rw [Int.natAbs_cast, show (2 * m + 2) ^ 2 = 4 * m * (m + 2) + 4 by ring1, Nat.succ_sub_one] - at h2 - rw [if_pos hn, show (2 * m + 3) ^ 2 = 2 * (2 * (m + 1) * (m + 2)) + 1 by ring1, - Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos] at h3 - rw [if_pos hn, mul_one] - constructor - · convert natDegree_sub_le_of_le (natDegree_mul_le_of_le natDegree_X_le h2.left) <| - natDegree_mul_le_of_le h3.left h1.left using 1 - convert (max_self _).symm using 2 <;> ring1 - · rw [coeff_sub, show (2 * m + 2) ^ 2 = 4 * m * (m + 2) + 4 by ring1, coeff_X_mul, h2.right, - show _ + _ = (2 * (m + 1) * (m + 2)) + (2 * m * (m + 1)) by ring1, - coeff_mul_of_natDegree_le h3.left h1.left, h3.right, h1.right] - push_cast - ring1 - · rcases Nat.odd_iff_not_even.mpr hn with ⟨m, rfl⟩ - rw [if_neg hn, show (2 * m + 2) ^ 2 = 2 * (2 * m * (m + 2)) + 4 by ring1, Nat.add_sub_cancel, - Nat.mul_div_cancel_left _ two_pos, show (2 * m + 2) / 2 = 2 * (m + 1) / 2 by rfl, - Nat.mul_div_cancel_left _ two_pos] at h1 - rw [Int.natAbs_cast, show (2 * m + 3) ^ 2 = 4 * (m + 1) * (m + 2) + 1 by ring1, - Nat.add_sub_cancel] at h2 - rw [if_neg hn, show (2 * m + 4) ^ 2 = 2 * (2 * (m + 1) * (m + 3)) + 4 by ring1, - Nat.add_sub_cancel, Nat.mul_div_cancel_left _ two_pos, - show (2 * m + 4) / 2 = 2 * (m + 2) / 2 by rfl, Nat.mul_div_cancel_left _ two_pos] at h3 - rw [if_neg hn] - constructor - · convert natDegree_sub_le_of_le (natDegree_mul_le_of_le natDegree_X_le h2.left) <| - natDegree_mul_le_of_le (natDegree_mul_le_of_le h3.left h1.left) - W.natDegree_divisionPolynomial2Sq using 1 - convert (max_self _).symm using 2 <;> ring1 - · rw [coeff_sub, show (2 * m + 3) ^ 2 = 4 * (m + 1) * (m + 2) + 1 by ring1, coeff_X_mul, - h2.right, show _ + _ = (2 * (m + 1) * (m + 3)) + (2 * m * (m + 2)) + 3 by ring1, - coeff_mul_of_natDegree_le (natDegree_mul_le_of_le h3.left h1.left) - W.natDegree_divisionPolynomial2Sq, coeff_mul_of_natDegree_le h3.left h1.left, h3.right, - h1.right, coeff_divisionPolynomial2Sq] - push_cast - ring1 - -private lemma natDegree_coeff_divisionPolynomialX (n : ℤ) : - (W.divisionPolynomialX n).natDegree ≤ n.natAbs ^ 2 ∧ - (W.divisionPolynomialX n).coeff (n.natAbs ^ 2) = 1 := by - rcases n.eq_nat_or_neg with ⟨n, rfl | rfl⟩ - · exact W.natDegree_coeff_divisionPolynomialX_ofNat n - · rw [divisionPolynomialX_neg, Int.natAbs_neg] - exact W.natDegree_coeff_divisionPolynomialX_ofNat n - -lemma natDegree_divisionPolynomialX (n : ℤ) : (W.divisionPolynomialX n).natDegree ≤ n.natAbs ^ 2 := - (W.natDegree_coeff_divisionPolynomialX n).left - -lemma coeff_divisionPolynomialX (n : ℤ) : (W.divisionPolynomialX n).coeff (n.natAbs ^ 2) = 1 := - (W.natDegree_coeff_divisionPolynomialX n).right - /-! ### The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$ -/ /-- The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ From 343fc96f8dc27bd89b635858e6cfa604c3841572 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 27 May 2024 16:16:30 +0100 Subject: [PATCH 24/47] Rename division polynomials and add further documentation --- .../EllipticCurve/DivisionPolynomial.lean | 465 ++++++++++++------ .../EllipticDivisibilitySequence.lean | 12 +- 2 files changed, 309 insertions(+), 168 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index eece929fadf709..1956b8d0818925 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -17,41 +17,67 @@ elliptic divisibility sequences defined in `Mathlib.NumberTheory.EllipticDivisib ## Mathematical background Let $W$ be a Weierstrass curve over a commutative ring $R$. The sequence of $n$-division polynomials -of $W$ is a normalised elliptic divisibility sequence $\psi_n \in R[X][Y]$ with initial values - * $\psi_0 = 0$, - * $\psi_1 = 1$, - * $\psi_2 = 2Y + a_1X + a_3$, - * $\psi_3 = 3X^4 + b_2X^3 + 3b_4X ^ 2 + 3b_6X + b_8$, and - * $\psi_4 = \psi_2 \cdot +$\psi_n \in R[X, Y]$ of $W$ is the normalised elliptic divisibility sequence with initial values + * $\psi_0 := 0$, + * $\psi_1 := 1$, + * $\psi_2 := 2Y + a_1X + a_3$, + * $\psi_3 := 3X^4 + b_2X^3 + 3b_4X ^ 2 + 3b_6X + b_8$, and + * $\psi_4 := \psi_2 \cdot (2X^6 + b_2X^5 + 5b_4X^4 + 10b_6X^3 + 10b_8X^2 + (b_2b_8 - b_4b_6)X + (b_4b_8 - b_6^2))$. -Furthermore, define associated sequences $\phi_n, \omega_n, \widetilde{\psi_n} \in R[X][Y]$ by - * $\phi_n := x\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, - * $\omega_n := \tfrac{1}{2\psi_n} \cdot (\psi_{2n} - \psi_n^2(a_1\phi_n + a_3\psi_n^2)$, and - * $\tilde{\psi}_n := \psi_n / \psi_2$ if $n$ is even and $\tilde{\psi}_n := \psi_n$ otherwise. -Under the coordinate ring $R[W]$ of $W$, the $2$-division polynomial $\psi_2$ satisfies the identity -$\psi_2^2 \equiv 4X^3 + b_2X^2 + 2b_4X + b_6$, so that $\phi_n, \psi_n^2, \tilde{\psi}_n \in R[X]$. -Furthermore, their leading terms can be computed via induction to be $\phi_n = X^{n^2} + \dots$ -and $\psi_n^2 = n^2X^{n^2 - 1} + \dots$, and $\tilde{\psi}_n = nX^{\tfrac{n^2 - 4}{2}} + \dots$ -if $n$ is even or $\tilde{\psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ otherwise. + +Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by + * $\phi_n := X\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, and + * $\omega_n := \tfrac{1}{2\psi_n} \cdot (\psi_{2n} - \psi_n^2(a_1\phi_n + a_3\psi_n^2)$. + +Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial +$\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences +associated to a normalised elliptic divisibility sequence show that $\psi_n / \psi_2$ are congruent +to certain polynomials in $R[X]$. In particular, define $\tilde{\Psi}_n \in R[X]$ as the auxiliary +sequence for a normalised elliptic divisibility sequence with initial values + * $\tilde{\Psi}_0 := 0$, + * $\tilde{\Psi}_1 := 1$, + * $\tilde{\Psi}_2 := 1$, + * $\tilde{\Psi}_3 := \psi_3$, and + * $\tilde{\Psi}_4 := \psi_4 / \psi_2$. + +The corresponding normalised elliptic divisibility sequence $\Psi_n \in R[X, Y]$ is then given by + * $\Psi_n := \tilde{\Psi}_n\psi_2$ if $n$ is even, and + * $\Psi_n := \tilde{\Psi}_n$ if $n$ is odd. + +Furthermore, define the associated sequences $\Psi_n^{[2]}, \Phi_n \in R[X]$ by + * $\Psi_n^{[2]} := \tilde{\Psi}_n^2\Psi_2^{[2]}$ if $n$ is even, + * $\Psi_n^{[2]} := \tilde{\Psi}_n^2$ if $n$ is odd, + * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1}\tilde{\Psi}_{n - 1}$ if $n$ is even, and + * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1}\tilde{\Psi}_{n - 1}\Psi_2^{[2]}$ if $n$ is odd. + +With these definitions in mind, $\psi_n \in R[X, Y]$ and $\phi_n \in R[X, Y]$ are congruent in +$R[W]$ to $\Psi_n \in R[X, Y]$ and $\Phi_n \in R[X]$ respectively, which are defined in terms of +$\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. By induction, their leading terms are + * $\tilde{\Psi}_n = nX^{\tfrac{n^2 - 4}{2}} + \dots$ if $n$ is even, + * $\tilde{\Psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ if $n$ is odd, + * $\Psi_n^{[2]} = n^2X^{n^2 - 1} + \dots$, and + * $\Phi_n = X^{n^2} + \dots$. ## Main definitions - * `WeierstrassCurve.divisionPolynomial2Sq`: $\psi_2^2$ viewed as a univariate polynomial. - * `WeierstrassCurve.divisionPolynomial`: $\tilde{\psi}_n$ viewed as a univariate polynomial. - * `WeierstrassCurve.divisionPolynomialZSq`: $\psi_n^2$ viewed as a univariate polynomial. - * `WeierstrassCurve.divisionPolynomialX`: $\phi_n$ viewed as a univariate polynomial. - * `WeierstrassCurve.divisionPolynomialZ`: $\psi_n$ viewed as a bivariate polynomial. - * TODO: $\omega$ viewed as a bivariate polynomial. + * `WeierstrassCurve.Ψ₂Sq`: the univariate polynomial $\Psi_2^{[2]}$. + * `WeierstrassCurve.Ψ'`: the univariate polynomials $\tilde{\Psi}_n$. + * `WeierstrassCurve.ΨSq`: the univariate polynomials $\Psi_n^{[2]}$. + * `WeierstrassCurve.Ψ`: the bivariate polynomials $\Psi_n$. + * `WeierstrassCurve.Φ`: the univariate polynomials $\Phi_n$. + * `WeierstrassCurve.ψ`: the bivariate $n$-division polynomials $\psi_n$. + * `WeierstrassCurve.φ`: the bivariate polynomials $\phi_n$. + * TODO: the bivariate polynomials $\omega_n$. ## Implementation notes Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate -polynomials $\psi_n$ are defined in terms of the univariate polynomials $\tilde{\psi}_n$. This is -done partially to avoid ring division, but more crucially to allow the definition of $\phi_n$ and -$\psi_n^2$ as univariate polynomials without needing to work under the coordinate ring, and to allow -the computation of their leading terms without ambiguity. Furthermore, evaluating these polynomials -at a rational point $(x, y)$ on $W$ recovers their original definition up to linear combinations of -the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring. +polynomials $\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is +done partially to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[2]}$ +and $\Phi_n$ as univariate polynomials without needing to work under the coordinate ring, and to +allow the computation of their leading terms without ambiguity. Furthermore, evaluating these +polynomials at a rational point on $W$ recovers their original definition up to linear combinations +of the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring. ## References @@ -67,248 +93,363 @@ open Polynomial PolynomialPolynomial local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) -universe u +universe u v namespace WeierstrassCurve -variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] (W : WeierstrassCurve R) (f : R →+* S) + +/-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ -/-! ### The univariate polynomial congruent to $\psi_2^2$ -/ +/-- The $2$-division polynomial $\psi_2 = \Psi_2$. -/ +protected noncomputable def ψ₂ : R[X][Y] := + W.toAffine.polynomialY -/-- The polynomial $4X^3 + b_2X^2 + 2b_4X + b_6$ congruent to the square $\psi_2^2$ of the -$2$-division polynomial $\psi_2 = 2Y + a_1X + a_3$ under $R[W]$. -/ -noncomputable def divisionPolynomial2Sq : R[X] := +/-- The univariate polynomial $\Psi_2^{[2]}$ congruent to $\psi_2^2$. -/ +protected noncomputable def Ψ₂Sq : R[X] := C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ --- TODO: remove `twoTorsionPolynomial` in favour of `divisionPolynomial2Sq` -lemma divisionPolynomial2Sq_eq : W.divisionPolynomial2Sq = W.twoTorsionPolynomial.toPoly := +lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by + rw [WeierstrassCurve.Ψ₂Sq, WeierstrassCurve.ψ₂, b₂, b₄, b₆, Affine.polynomialY, Affine.polynomial] + C_simp + ring1 + +-- TODO: remove `twoTorsionPolynomial` in favour of `Ψ₂Sq` +lemma Ψ₂Sq_eq : W.Ψ₂Sq = W.twoTorsionPolynomial.toPoly := rfl -/-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{N}$ -/ +/-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$ -/ -/-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials -$\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{N}$. -/ -noncomputable def divisionPolynomial' (n : ℕ) : R[X] := - preNormEDS' (W.divisionPolynomial2Sq ^ 2) - (3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈) - (2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + - C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2)) n +/-- The $3$-division polynomial $\psi_3 = \Psi_3$. -/ +protected noncomputable def Ψ₃ : R[X] := + 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ + +/-- The univariate polynomial $\tilde{\Psi}_4$, which is auxiliary to the $4$-division polynomial +$\psi_4 = \Psi_4 = \tilde{\Psi}_4\psi_2$. -/ +protected noncomputable def Ψ₄' : R[X] := + 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) + +/-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$, which are auxiliary to the +bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ +protected noncomputable def Ψ'' (n : ℕ) : R[X] := + preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n @[simp] -lemma divisionPolynomial'_zero : W.divisionPolynomial' 0 = 0 := +lemma Ψ''_zero : W.Ψ'' 0 = 0 := preNormEDS'_zero .. @[simp] -lemma divisionPolynomial'_one : W.divisionPolynomial' 1 = 1 := +lemma Ψ''_one : W.Ψ'' 1 = 1 := preNormEDS'_one .. @[simp] -lemma divisionPolynomial'_two : W.divisionPolynomial' 2 = 1 := +lemma Ψ''_two : W.Ψ'' 2 = 1 := preNormEDS'_two .. @[simp] -lemma divisionPolynomial'_three : W.divisionPolynomial' 3 = - 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ := +lemma Ψ''_three : W.Ψ'' 3 = W.Ψ₃ := preNormEDS'_three .. @[simp] -lemma divisionPolynomial'_four : W.divisionPolynomial' 4 = - 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + - C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) := +lemma Ψ''_four : W.Ψ'' 4 = W.Ψ₄' := preNormEDS'_four .. -lemma divisionPolynomial'_odd (m : ℕ) : W.divisionPolynomial' (2 * (m + 2) + 1) = - W.divisionPolynomial' (m + 4) * W.divisionPolynomial' (m + 2) ^ 3 * - (if Even m then W.divisionPolynomial2Sq ^ 2 else 1) - - W.divisionPolynomial' (m + 1) * W.divisionPolynomial' (m + 3) ^ 3 * - (if Even m then 1 else W.divisionPolynomial2Sq ^ 2) := +lemma Ψ''_odd (m : ℕ) : W.Ψ'' (2 * (m + 2) + 1) = + W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := preNormEDS'_odd .. -lemma divisionPolynomial'_even (m : ℕ) : W.divisionPolynomial' (2 * (m + 3)) = - W.divisionPolynomial' (m + 2) ^ 2 * W.divisionPolynomial' (m + 3) * - W.divisionPolynomial' (m + 5) - - W.divisionPolynomial' (m + 1) * W.divisionPolynomial' (m + 3) * - W.divisionPolynomial' (m + 4) ^ 2 := +lemma Ψ''_even (m : ℕ) : W.Ψ'' (2 * (m + 3)) = + W.Ψ'' (m + 2) ^ 2 * W.Ψ'' (m + 3) * W.Ψ'' (m + 5) - + W.Ψ'' (m + 1) * W.Ψ'' (m + 3) * W.Ψ'' (m + 4) ^ 2 := preNormEDS'_even .. -/-! ### The univariate polynomials congruent to $\tilde{\psi}_n$ for $n \in \mathbb{Z}$ -/ +/-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$ -/ -/-- The univariate polynomials congruent under $R[W]$ to the bivariate auxiliary polynomials -$\tilde{\psi}_n$ associated to the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -noncomputable def divisionPolynomial (n : ℤ) : R[X] := - n.sign * W.divisionPolynomial' n.natAbs +/-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the +bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ +protected noncomputable def Ψ' (n : ℤ) : R[X] := + preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n @[simp] -lemma divisionPolynomial_ofNat (n : ℕ) : W.divisionPolynomial n = W.divisionPolynomial' n := +lemma Ψ'_ofNat (n : ℕ) : W.Ψ' n = W.Ψ'' n := preNormEDS_ofNat .. @[simp] -lemma divisionPolynomial_zero : W.divisionPolynomial 0 = 0 := +lemma Ψ'_zero : W.Ψ' 0 = 0 := preNormEDS_zero .. @[simp] -lemma divisionPolynomial_one : W.divisionPolynomial 1 = 1 := +lemma Ψ'_one : W.Ψ' 1 = 1 := preNormEDS_one .. @[simp] -lemma divisionPolynomial_two : W.divisionPolynomial 2 = 1 := +lemma Ψ'_two : W.Ψ' 2 = 1 := preNormEDS_two .. @[simp] -lemma divisionPolynomial_three : W.divisionPolynomial 3 = W.divisionPolynomial' 3 := +lemma Ψ'_three : W.Ψ' 3 = W.Ψ₃ := preNormEDS_three .. @[simp] -lemma divisionPolynomial_four : W.divisionPolynomial 4 = W.divisionPolynomial' 4 := +lemma Ψ'_four : W.Ψ' 4 = W.Ψ₄' := preNormEDS_four .. -lemma divisionPolynomial_odd (m : ℕ) : W.divisionPolynomial (2 * (m + 2) + 1) = - W.divisionPolynomial (m + 4) * W.divisionPolynomial (m + 2) ^ 3 * - (if Even m then W.divisionPolynomial2Sq ^ 2 else 1) - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) ^ 3 * - (if Even m then 1 else W.divisionPolynomial2Sq ^ 2) := +lemma Ψ'_odd (m : ℕ) : W.Ψ' (2 * (m + 2) + 1) = + W.Ψ' (m + 4) * W.Ψ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.Ψ' (m + 1) * W.Ψ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := preNormEDS_odd .. -lemma divisionPolynomial_even (m : ℕ) : W.divisionPolynomial (2 * (m + 3)) = - W.divisionPolynomial (m + 2) ^ 2 * W.divisionPolynomial (m + 3) * - W.divisionPolynomial (m + 5) - - W.divisionPolynomial (m + 1) * W.divisionPolynomial (m + 3) * - W.divisionPolynomial (m + 4) ^ 2 := +lemma Ψ'_even (m : ℕ) : W.Ψ' (2 * (m + 3)) = + W.Ψ' (m + 2) ^ 2 * W.Ψ' (m + 3) * W.Ψ' (m + 5) - + W.Ψ' (m + 1) * W.Ψ' (m + 3) * W.Ψ' (m + 4) ^ 2 := preNormEDS_even .. @[simp] -lemma divisionPolynomial_neg (n : ℤ) : W.divisionPolynomial (-n) = -W.divisionPolynomial n := +lemma Ψ'_neg (n : ℤ) : W.Ψ' (-n) = -W.Ψ' n := preNormEDS_neg .. -/-! ### The univariate polynomials congruent to $\psi_n^2$ for $n \in \mathbb{Z}$ -/ +/-! ### The univariate polynomials $\Psi_n^{[2]}$ -/ + +/-- The univariate polynomials $\Psi_n^{[2]}$ congruent to $\psi_n^2$. -/ +protected noncomputable def ΨSq (n : ℤ) : R[X] := + W.Ψ' n ^ 2 * if Even n.natAbs then W.Ψ₂Sq else 1 + +@[simp] +lemma ΨSq_ofNat (n : ℕ) : W.ΨSq n = W.Ψ'' n ^ 2 * if Even n then W.Ψ₂Sq else 1 := by + rw [WeierstrassCurve.ΨSq, Ψ'_ofNat, Int.natAbs_cast] + +@[simp] +lemma ΨSq_zero : W.ΨSq 0 = 0 := by + erw [ΨSq_ofNat, zero_pow two_ne_zero, zero_mul] + +@[simp] +lemma ΨSq_one : W.ΨSq 1 = 1 := by + erw [ΨSq_ofNat, one_pow, mul_one] + +@[simp] +lemma ΨSq_two : W.ΨSq 2 = W.Ψ₂Sq := by + erw [ΨSq_ofNat, one_pow, one_mul, if_pos even_two] + +@[simp] +lemma ΨSq_three : W.ΨSq 3 = W.Ψ₃ ^ 2 := by + erw [ΨSq_ofNat, Ψ''_three, mul_one] + +@[simp] +lemma ΨSq_four : W.ΨSq 4 = W.Ψ₄' ^ 2 * W.Ψ₂Sq := by + erw [ΨSq_ofNat, Ψ''_four, if_pos <| by decide] + +lemma ΨSq_odd (m : ℕ) : W.ΨSq (2 * (m + 2) + 1) = + (W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by + erw [ΨSq_ofNat, Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one, mul_one] -/-- The univariate polynomials congruent under $R[W]$ to the bivariate squares $\psi_n^2$ of the -$n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -noncomputable def divisionPolynomialZSq (n : ℤ) : R[X] := - W.divisionPolynomial n ^ 2 * if Even n.natAbs then W.divisionPolynomial2Sq else 1 +lemma ΨSq_even (m : ℕ) : W.ΨSq (2 * (m + 3)) = + (W.Ψ'' (m + 2) ^ 2 * W.Ψ'' (m + 3) * W.Ψ'' (m + 5) - + W.Ψ'' (m + 1) * W.Ψ'' (m + 3) * W.Ψ'' (m + 4) ^ 2) ^ 2 * W.Ψ₂Sq := by + erw [ΨSq_ofNat, Ψ''_even, if_pos <| even_two_mul _] @[simp] -lemma divisionPolynomialZSq_ofNat (n : ℕ) : W.divisionPolynomialZSq n = - W.divisionPolynomial' n ^ 2 * if Even n then W.divisionPolynomial2Sq else 1 := by - rw [divisionPolynomialZSq, divisionPolynomial_ofNat, Int.natAbs_cast] +lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by + rw [WeierstrassCurve.ΨSq, Ψ'_neg, neg_sq, Int.natAbs_neg, WeierstrassCurve.ΨSq] + +/-! ### The bivariate polynomials $\Psi_n$ -/ + +/-- The bivariate polynomials $\Psi_n$ congruent to the $n$-division polynomials $\psi_n$. -/ +protected noncomputable def Ψ (n : ℤ) : R[X][Y] := + C (W.Ψ' n) * if Even n.natAbs then W.ψ₂ else 1 @[simp] -lemma divisionPolynomialZSq_zero : W.divisionPolynomialZSq 0 = 0 := by - erw [divisionPolynomialZSq_ofNat, zero_pow two_ne_zero, zero_mul] +lemma Ψ_ofNat (n : ℕ) : W.Ψ n = C (W.Ψ'' n) * if Even n then W.ψ₂ else 1 := by + rw [WeierstrassCurve.Ψ, Ψ'_ofNat, Int.natAbs_cast] @[simp] -lemma divisionPolynomialZSq_one : W.divisionPolynomialZSq 1 = 1 := by - erw [divisionPolynomialZSq_ofNat, one_pow, mul_one] +lemma Ψ_zero : W.Ψ 0 = 0 := by + erw [Ψ_ofNat, C_0, zero_mul] @[simp] -lemma divisionPolynomialZSq_two : W.divisionPolynomialZSq 2 = W.divisionPolynomial2Sq := by - erw [divisionPolynomialZSq_ofNat, one_pow, one_mul, if_pos even_two] +lemma Ψ_one : W.Ψ 1 = 1 := by + erw [Ψ_ofNat, C_1, mul_one] @[simp] -lemma divisionPolynomialZSq_three : W.divisionPolynomialZSq 3 = W.divisionPolynomial' 3 ^ 2 := by - erw [divisionPolynomialZSq_ofNat, mul_one] +lemma Ψ_two : W.Ψ 2 = W.ψ₂ := by + erw [Ψ_ofNat, one_mul, if_pos even_two] @[simp] -lemma divisionPolynomialZSq_four : W.divisionPolynomialZSq 4 = - W.divisionPolynomial' 4 ^ 2 * W.divisionPolynomial2Sq := by - erw [divisionPolynomialZSq_ofNat, if_pos <| by decide] +lemma Ψ_three : W.Ψ 3 = C W.Ψ₃ := by + erw [Ψ_ofNat, Ψ''_three, mul_one] @[simp] -lemma divisionPolynomialZSq_neg (n : ℤ) : - W.divisionPolynomialZSq (-n) = W.divisionPolynomialZSq n := by - rw [divisionPolynomialZSq, divisionPolynomial_neg, neg_sq, Int.natAbs_neg, divisionPolynomialZSq] +lemma Ψ_four : W.Ψ 4 = C W.Ψ₄' * W.ψ₂ := by + erw [Ψ_ofNat, Ψ''_four, if_pos <| by decide] -/-! ### The univariate polynomials congruent to $\phi_n$ for $n \in \mathbb{Z}$ -/ +lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = + W.Ψ (m + 4) * W.Ψ (m + 2) ^ 3 - W.Ψ (m + 1) * W.Ψ (m + 3) ^ 3 + + W.toAffine.polynomial * (16 * W.toAffine.polynomial - 8 * W.ψ₂ ^ 2) * C + (if Even m then W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 + else -W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3) := by + repeat erw [Ψ_ofNat] + rw [Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one] + by_cases hm : Even m + · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm + have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 + have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 + have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 + rw [if_pos hm, if_pos hm, if_pos hm4, if_pos hm2, if_neg hm1, if_neg hm3, if_pos hm] + C_simp + rw [C_Ψ₂Sq_eq] + ring1 + · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm + have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 + have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 + have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 + rw [if_neg hm, if_neg hm, if_neg hm4, if_neg hm2, if_pos hm1, if_pos hm3, if_neg hm] + C_simp + rw [C_Ψ₂Sq_eq] + ring1 -/-- The univariate polynomials congruent under $R[W]$ to the bivariate polynomials $\phi_n$ defined -in terms of the $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -noncomputable def divisionPolynomialX (n : ℤ) : R[X] := - X * W.divisionPolynomialZSq n - W.divisionPolynomial (n + 1) * W.divisionPolynomial (n - 1) * - if Even n.natAbs then 1 else W.divisionPolynomial2Sq +lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = + W.Ψ (m + 2) ^ 2 * W.Ψ (m + 3) * W.Ψ (m + 5) - W.Ψ (m + 1) * W.Ψ (m + 3) * W.Ψ (m + 4) ^ 2 := by + repeat erw [Ψ_ofNat] + erw [Ψ''_even, if_pos <| even_two_mul _] + by_cases hm : Even m + · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm + have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 + have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 + have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 + have hm5 : ¬Even (m + 5) := fun h => Nat.even_add_one.mp h hm4 + rw [if_pos hm2, if_neg hm3, if_neg hm5, if_neg hm1, if_pos hm4] + C_simp + ring1 + · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm + have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 + have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 + have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 + have hm5 : Even (m + 5) := Nat.even_add_one.mpr hm4 + rw [if_neg hm2, if_pos hm3, if_pos hm5, if_pos hm1, if_neg hm4] + C_simp + ring1 @[simp] -lemma divisionPolynomialX_ofNat (n : ℕ) : W.divisionPolynomialX (n + 1) = - X * W.divisionPolynomial' (n + 1) ^ 2 * - (if Even n then 1 else W.divisionPolynomial2Sq) - - W.divisionPolynomial' (n + 2) * W.divisionPolynomial' n * - (if Even n then W.divisionPolynomial2Sq else 1) := by - erw [divisionPolynomialX, divisionPolynomialZSq_ofNat, ← mul_assoc, divisionPolynomial_ofNat, - add_sub_cancel_right, divisionPolynomial_ofNat, Int.natAbs_cast] +lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by + rw [WeierstrassCurve.Ψ, Ψ'_neg, C_neg, neg_mul (α := R[X][Y]), Int.natAbs_neg, WeierstrassCurve.Ψ] + +/-! ### The univariate polynomials $\Phi_n$ -/ + +/-- The univariate polynomials $\Phi_n$ congruent to $\phi_n$. -/ +protected noncomputable def Φ (n : ℤ) : R[X] := + X * W.ΨSq n - W.Ψ' (n + 1) * W.Ψ' (n - 1) * if Even n.natAbs then 1 else W.Ψ₂Sq + +@[simp] +lemma Φ_ofNat (n : ℕ) : W.Φ (n + 1) = + X * W.Ψ'' (n + 1) ^ 2 * (if Even n then 1 else W.Ψ₂Sq) - + W.Ψ'' (n + 2) * W.Ψ'' n * (if Even n then W.Ψ₂Sq else 1) := by + erw [WeierstrassCurve.Φ, ΨSq_ofNat, ← mul_assoc, Ψ'_ofNat, add_sub_cancel_right, Ψ'_ofNat, + Int.natAbs_cast] simp only [Nat.even_add_one, ite_not] @[simp] -lemma divisionPolynomialX_zero : W.divisionPolynomialX 0 = 1 := by - rw [divisionPolynomialX, divisionPolynomialZSq_zero, mul_zero, zero_sub, zero_add, - divisionPolynomial_one, one_mul, zero_sub, divisionPolynomial_neg, divisionPolynomial_one, - neg_one_mul, neg_neg, if_pos even_zero.natAbs] +lemma Φ_zero : W.Φ 0 = 1 := by + rw [WeierstrassCurve.Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, Ψ'_one, one_mul, zero_sub, Ψ'_neg, + Ψ'_one, neg_one_mul, neg_neg, if_pos even_zero.natAbs] @[simp] -lemma divisionPolynomialX_one : W.divisionPolynomialX 1 = X := by - erw [divisionPolynomialX_ofNat, one_pow, mul_one, mul_one, mul_zero, zero_mul, sub_zero] +lemma Φ_one : W.Φ 1 = X := by + erw [Φ_ofNat, Ψ''_one, one_pow, mul_one, mul_one, mul_zero, zero_mul, sub_zero] @[simp] -lemma divisionPolynomialX_two : W.divisionPolynomialX 2 = - X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by - erw [divisionPolynomialX_ofNat, divisionPolynomial'_two, if_neg Nat.not_even_one, - divisionPolynomial2Sq, divisionPolynomial'_three, if_neg Nat.not_even_one] +lemma Φ_two : W.Φ 2 = X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by + erw [Φ_ofNat, Ψ''_two, if_neg Nat.not_even_one, WeierstrassCurve.Ψ₂Sq, Ψ''_three, mul_one, + WeierstrassCurve.Ψ₃] C_simp ring1 @[simp] -lemma divisionPolynomialX_three : W.divisionPolynomialX 3 = - X * W.divisionPolynomial' 3 ^ 2 - W.divisionPolynomial' 4 * W.divisionPolynomial2Sq := by - erw [divisionPolynomialX_ofNat, divisionPolynomial'_three, mul_one, mul_one, if_pos even_two] +lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.Ψ₄' * W.Ψ₂Sq := by + erw [Φ_ofNat, Ψ''_three, mul_one, Ψ''_four, mul_one, if_pos even_two] @[simp] -lemma divisionPolynomialX_four : W.divisionPolynomialX 4 = - X * W.divisionPolynomial' 4 ^ 2 * W.divisionPolynomial2Sq - W.divisionPolynomial' 3 * - (W.divisionPolynomial' 4 * W.divisionPolynomial2Sq ^ 2 - W.divisionPolynomial' 3 ^ 3) := by - erw [divisionPolynomialX_ofNat, if_neg <| by decide, if_neg <| by decide, - show 3 + 2 = 2 * 2 + 1 by rfl, divisionPolynomial'_odd, divisionPolynomial'_two, - if_pos even_zero, if_pos even_zero] +lemma Φ_four : W.Φ 4 = X * W.Ψ₄' ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.Ψ₄' * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by + erw [Φ_ofNat, Ψ''_four, if_neg <| by decide, show 3 + 2 = 2 * 2 + 1 by rfl, Ψ''_odd, Ψ''_four, + Ψ''_two, if_pos even_zero, Ψ''_three, if_pos even_zero] ring1 @[simp] -lemma divisionPolynomialX_neg (n : ℤ) : W.divisionPolynomialX (-n) = W.divisionPolynomialX n := by - rw [divisionPolynomialX, divisionPolynomialZSq_neg, neg_add_eq_sub, ← neg_sub n, - divisionPolynomial_neg, ← neg_add', divisionPolynomial_neg, neg_mul_neg, - mul_comm <| W.divisionPolynomial _, Int.natAbs_neg, divisionPolynomialX] +lemma Φ_neg (n : ℤ) : W.Φ (-n) = W.Φ n := by + rw [WeierstrassCurve.Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, Ψ'_neg, ← neg_add', Ψ'_neg, + neg_mul_neg, mul_comm <| W.Ψ' _, Int.natAbs_neg, WeierstrassCurve.Φ] + +/-! ### The bivariate polynomials $\psi_n$ -/ + +/-- The bivariate $n$-division polynomials $\psi_n$. -/ +protected noncomputable def ψ (n : ℤ) : R[X][Y] := + normEDS W.ψ₂ (C W.Ψ₃) (C W.Ψ₄') n + +@[simp] +lemma ψ_zero : W.ψ 0 = 0 := + normEDS_zero .. + +@[simp] +lemma ψ_one : W.ψ 1 = 1 := + normEDS_one .. + +@[simp] +lemma ψ_two : W.ψ 2 = W.ψ₂ := + normEDS_two .. + +@[simp] +lemma ψ_three : W.ψ 3 = C W.Ψ₃ := + normEDS_three .. -/-! ### The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$ -/ +@[simp] +lemma ψ_four : W.ψ 4 = C W.Ψ₄' * W.ψ₂ := + normEDS_four .. + +lemma ψ_odd (m : ℕ) : W.ψ (2 * (m + 2) + 1) = + W.ψ (m + 4) * W.ψ (m + 2) ^ 3 - W.ψ (m + 1) * W.ψ (m + 3) ^ 3 := + normEDS_odd .. -/-- The bivariate $n$-division polynomials $\psi_n$ for $n \in \mathbb{Z}$. -/ -noncomputable def divisionPolynomialZ (n : ℤ) : R[X][Y] := - C (W.divisionPolynomial n) * if Even n.natAbs then Y - W.toAffine.negPolynomial else 1 +lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = + W.ψ (m + 2) ^ 2 * W.ψ (m + 3) * W.ψ (m + 5) - W.ψ (m + 1) * W.ψ (m + 3) * W.ψ (m + 4) ^ 2 := + normEDS_even .. @[simp] -lemma divisionPolynomialZ_ofNat (n : ℕ) : W.divisionPolynomialZ n = - C (W.divisionPolynomial' n) * if Even n then Y - W.toAffine.negPolynomial else 1 := by - rw [divisionPolynomialZ, divisionPolynomial_ofNat, Int.natAbs_cast] +lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n := + normEDS_neg .. + +/-! ### The bivariate polynomials $\phi_n$ -/ + +/-- The bivariate polynomials $\phi_n$. -/ +protected noncomputable def φ (n : ℤ) : R[X][Y] := + C X * W.ψ n ^ 2 - W.ψ (n + 1) * W.ψ (n - 1) @[simp] -lemma divisionPolynomialZ_zero : W.divisionPolynomialZ 0 = 0 := by - erw [divisionPolynomialZ_ofNat, C_0, zero_mul] +lemma φ_zero : W.φ 0 = 1 := by + erw [WeierstrassCurve.φ, ψ_zero, zero_pow two_ne_zero, mul_zero, zero_sub, ψ_one, one_mul, + zero_sub, ψ_neg, neg_neg, ψ_one] @[simp] -lemma divisionPolynomialZ_one : W.divisionPolynomialZ 1 = 1 := by - erw [divisionPolynomialZ_ofNat, C_1, mul_one] +lemma φ_one : W.φ 1 = C X := by + erw [WeierstrassCurve.φ, ψ_one, one_pow, mul_one, ψ_zero, mul_zero, sub_zero] @[simp] -lemma divisionPolynomialZ_two : W.divisionPolynomialZ 2 = Y - W.toAffine.negPolynomial := by - erw [divisionPolynomialZ_ofNat, one_mul, if_pos even_two] +lemma φ_two : W.φ 2 = C X * W.ψ₂ ^ 2 - C W.Ψ₃ := by + erw [WeierstrassCurve.φ, ψ_two, ψ_three, ψ_one, mul_one] @[simp] -lemma divisionPolynomialZ_three : W.divisionPolynomialZ 3 = C (W.divisionPolynomial' 3) := by - erw [divisionPolynomialZ_ofNat, mul_one] +lemma φ_three : W.φ 3 = C X * C W.Ψ₃ ^ 2 - C W.Ψ₄' * W.ψ₂ ^ 2 := by + erw [WeierstrassCurve.φ, ψ_three, ψ_four, mul_assoc, ψ_two, ← sq] @[simp] -lemma divisionPolynomialZ_four : W.divisionPolynomialZ 4 = - C (W.divisionPolynomial' 4) * (Y - W.toAffine.negPolynomial) := by - erw [divisionPolynomialZ_ofNat, if_pos <| by decide] +lemma φ_four : W.φ 4 = C X * C W.Ψ₄' ^ 2 * W.ψ₂ ^ 2 - C W.Ψ₄' * W.ψ₂ ^ 4 * C W.Ψ₃ + C W.Ψ₃ ^ 4 := by + erw [WeierstrassCurve.φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, ψ_four, ψ_two, ψ_one, + ψ_three] + ring1 @[simp] -lemma divisionPolynomialZ_neg (n : ℤ) : W.divisionPolynomialZ (-n) = -W.divisionPolynomialZ n := by - rw [divisionPolynomialZ, divisionPolynomial_neg, C_neg, neg_mul (α := R[X][Y]), Int.natAbs_neg, - divisionPolynomialZ] +lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by + rw [WeierstrassCurve.φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, + ← neg_add', ψ_neg, neg_mul_neg, mul_comm <| W.ψ _, WeierstrassCurve.φ] end WeierstrassCurve diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index dfc40431a25c37..d2bb171e8ff630 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -281,7 +281,7 @@ lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = ring1 @[simp] -lemma normEDS_neg (n : ℕ) : normEDS b c d (-n) = -normEDS b c d n := by +lemma normEDS_neg (n : ℤ) : normEDS b c d (-n) = -normEDS b c d n := by rw [normEDS, preNormEDS_neg, Int.natAbs_neg, neg_mul, normEDS] /-- Strong recursion principle for a normalised EDS: if we have @@ -291,11 +291,11 @@ lemma normEDS_neg (n : ℕ) : normEDS b c d (-n) = -normEDS b c d n := by then we have `P n` for all `n : ℕ`. -/ @[elab_as_elim] noncomputable def normEDSRec' {P : ℕ → Sort u} - (base0 : P 0) (base1 : P 1) (base2 : P 2) (base3 : P 3) (base4 : P 4) + (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) (even : ∀ m : ℕ, (∀ k < 2 * (m + 3), P k) → P (2 * (m + 3))) (odd : ∀ m : ℕ, (∀ k < 2 * (m + 2) + 1, P k) → P (2 * (m + 2) + 1)) (n : ℕ) : P n := - n.evenOddStrongRec (by rintro (_ | _ | _ | _) h; exacts [base0, base2, base4, even _ h]) - (by rintro (_ | _ | _) h; exacts [base1, base3, odd _ h]) + n.evenOddStrongRec (by rintro (_ | _ | _ | _) h; exacts [zero, two, four, even _ h]) + (by rintro (_ | _ | _) h; exacts [one, three, odd _ h]) /-- Recursion principle for a normalised EDS: if we have * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, @@ -306,10 +306,10 @@ noncomputable def normEDSRec' {P : ℕ → Sort u} then we have `P n` for all `n : ℕ`. -/ @[elab_as_elim] noncomputable def normEDSRec {P : ℕ → Sort u} - (base0 : P 0) (base1 : P 1) (base2 : P 2) (base3 : P 3) (base4 : P 4) + (zero : P 0) (one : P 1) (two : P 2) (three : P 3) (four : P 4) (even : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (m + 5) → P (2 * (m + 3))) (odd : ∀ m : ℕ, P (m + 1) → P (m + 2) → P (m + 3) → P (m + 4) → P (2 * (m + 2) + 1)) (n : ℕ) : P n := - normEDSRec' base0 base1 base2 base3 base4 + normEDSRec' zero one two three four (fun _ ih => by apply even <;> exact ih _ <| by linarith only) (fun _ ih => by apply odd <;> exact ih _ <| by linarith only) n From 087072fbf17b5ab2029950bde2d78a77a68c4af5 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 29 May 2024 23:12:06 +0100 Subject: [PATCH 25/47] Add further map lemmas --- .../EllipticCurve/Affine.lean | 234 ++++++++++++------ 1 file changed, 154 insertions(+), 80 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 93c6c4e79a2563..874a64194851c5 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -78,15 +78,11 @@ The group law on this set is then uniquely determined by these constructions. elliptic curve, rational point, affine coordinates -/ -local macro "map_simp" : tactic => - `(tactic| simp only [map_ofNat, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀, - WeierstrassCurve.map]) - /-- The notation `Y` for `X` in the `PolynomialPolynomial` scope. -/ -scoped[PolynomialPolynomial] notation "Y" => Polynomial.X +scoped[PolynomialPolynomial] notation3:max "Y" => Polynomial.X (R := Polynomial _) /-- The notation `R[X][Y]` for `R[X][X]` in the `PolynomialPolynomial` scope. -/ -scoped[PolynomialPolynomial] notation R "[X][Y]" => Polynomial (Polynomial R) +scoped[PolynomialPolynomial] notation3:max R "[X][Y]" => Polynomial (Polynomial R) open Polynomial PolynomialPolynomial @@ -95,11 +91,17 @@ local macro "C_simp" : tactic => local macro "derivative_simp" : tactic => `(tactic| simp only [derivative_C, derivative_X, derivative_X_pow, derivative_neg, derivative_add, - derivative_sub, derivative_mul, derivative_sq]) + derivative_sub, derivative_mul, derivative_sq]) local macro "eval_simp" : tactic => `(tactic| simp only [eval_C, eval_X, eval_neg, eval_add, eval_sub, eval_mul, eval_pow]) +local macro "map_simp" : tactic => + `(tactic| simp only [map_ofNat, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀, + Polynomial.map_ofNat, map_C, map_X, Polynomial.map_neg, Polynomial.map_add, Polynomial.map_sub, + Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_div, coe_mapRingHom, + WeierstrassCurve.map]) + universe r s u v w /-! ## Weierstrass curves -/ @@ -286,8 +288,8 @@ lemma nonsingular_iff (x y : R) : W.Nonsingular x y ↔ @[simp] lemma nonsingular_zero : W.Nonsingular 0 0 ↔ W.a₆ = 0 ∧ (W.a₃ ≠ 0 ∨ W.a₄ ≠ 0) := by - rw [Nonsingular, equation_zero, C_0, eval_polynomialX_zero, neg_ne_zero, - eval_polynomialY_zero, or_comm] + rw [Nonsingular, equation_zero, C_0, eval_polynomialX_zero, neg_ne_zero, eval_polynomialY_zero, + or_comm] #align weierstrass_curve.nonsingular_zero WeierstrassCurve.Affine.nonsingular_zero lemma nonsingular_iff_variableChange (x y : R) : @@ -320,7 +322,7 @@ section Ring /-- The polynomial $-Y - a_1X - a_3$ associated to negation. -/ noncomputable def negPolynomial : R[X][Y] := - -Y - C (C W.a₁ * X + C W.a₃) + -(Y : R[X][Y]) - C (C W.a₁ * X + C W.a₃) #align weierstrass_curve.neg_polynomial WeierstrassCurve.Affine.negPolynomial /-- The $Y$-coordinate of the negation of an affine point in `W`. @@ -634,7 +636,7 @@ inductive Point #align weierstrass_curve.point WeierstrassCurve.Affine.Point /-- For an algebraic extension `S` of `R`, the type of nonsingular `S`-rational points on `W`. -/ -scoped[WeierstrassCurve] notation3 W "⟮" S "⟯" => Affine.Point <| baseChange W S +scoped notation3 W "⟮" S "⟯" => Affine.Point <| baseChange W S namespace Point @@ -770,140 +772,218 @@ section BaseChange /-! ### Maps and base changes -/ -variable {A : Type v} [CommRing A] (φ : R →+* A) +variable {S : Type v} [CommRing S] (f : R →+* S) + +lemma map_polynomial : (W.map f).toAffine.polynomial = W.polynomial.map (mapRingHom f) := by + simp only [polynomial] + map_simp + +lemma map_eval_polynomial (x : R[X]) (y : R) : + ((W.map f).toAffine.polynomial.eval <| x.map f).eval (f y) = + f ((W.polynomial.eval x).eval y) := by + rw [map_polynomial, eval_map, ← coe_mapRingHom, eval₂_hom, coe_mapRingHom, eval_map, eval₂_hom] -lemma map_equation {φ : R →+* A} (hφ : Function.Injective φ) (x y : R) : - (W.map φ).toAffine.Equation (φ x) (φ y) ↔ W.Equation x y := by - simpa only [equation_iff] using - ⟨fun h => hφ <| by map_simp; exact h, fun h => by convert congr_arg φ h <;> map_simp⟩ +variable {f} in +lemma map_equation (hf : Function.Injective f) (x y : R) : + (W.map f).toAffine.Equation (f x) (f y) ↔ W.Equation x y := by + simp only [Equation, ← map_C, map_eval_polynomial, map_eq_zero_iff f hf] #align weierstrass_curve.equation_iff_base_change WeierstrassCurve.Affine.map_equation -lemma map_nonsingular {φ : R →+* A} (hφ : Function.Injective φ) (x y : R) : - (W.map φ).toAffine.Nonsingular (φ x) (φ y) ↔ W.Nonsingular x y := by - rw [nonsingular_iff, nonsingular_iff, and_congr <| W.map_equation hφ x y] - refine ⟨Or.imp (not_imp_not.mpr fun h => ?_) (not_imp_not.mpr fun h => ?_), - Or.imp (not_imp_not.mpr fun h => ?_) (not_imp_not.mpr fun h => ?_)⟩ - any_goals apply hφ; map_simp; exact h - any_goals convert congr_arg φ h <;> map_simp +lemma map_polynomialX : (W.map f).toAffine.polynomialX = W.polynomialX.map (mapRingHom f) := by + simp only [polynomialX] + map_simp + +lemma map_eval_polynomialX (x : R[X]) (y : R) : + ((W.map f).toAffine.polynomialX.eval <| x.map f).eval (f y) = + f ((W.polynomialX.eval x).eval y) := by + rw [map_polynomialX, eval_map, ← coe_mapRingHom, eval₂_hom, coe_mapRingHom, eval_map, eval₂_hom] + +lemma map_polynomialY : (W.map f).toAffine.polynomialY = W.polynomialY.map (mapRingHom f) := by + simp only [polynomialY] + map_simp + +lemma map_eval_polynomialY (x : R[X]) (y : R) : + ((W.map f).toAffine.polynomialY.eval <| x.map f).eval (f y) = + f ((W.polynomialY.eval x).eval y) := by + rw [map_polynomialY, eval_map, ← coe_mapRingHom, eval₂_hom, coe_mapRingHom, eval_map, eval₂_hom] + +variable {f} in +lemma map_nonsingular (hf : Function.Injective f) (x y : R) : + (W.map f).toAffine.Nonsingular (f x) (f y) ↔ W.Nonsingular x y := by + simp only [Nonsingular, W.map_equation hf, ← map_C, map_eval_polynomialX, map_eval_polynomialY, + map_ne_zero_iff f hf] #align weierstrass_curve.nonsingular_iff_base_change WeierstrassCurve.Affine.map_nonsingular -lemma map_negY (x y : R) : (W.map φ).toAffine.negY (φ x) (φ y) = φ (W.negY x y) := by +lemma map_negPolynomial : + (W.map f).toAffine.negPolynomial = W.negPolynomial.map (mapRingHom f) := by + simp only [negPolynomial] + map_simp + +lemma map_negY (x y : R) : (W.map f).toAffine.negY (f x) (f y) = f (W.negY x y) := by simp only [negY] map_simp set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_neg_Y WeierstrassCurve.Affine.map_negY +lemma map_linePolynomial (x y L : R) : + linePolynomial (f x) (f y) (f L) = (linePolynomial x y L).map f := by + simp only [linePolynomial] + map_simp + +lemma map_addPolynomial (x y L : R) : + (W.map f).toAffine.addPolynomial (f x) (f y) (f L) = (W.addPolynomial x y L).map f := by + rw [addPolynomial, map_polynomial, eval_map, linePolynomial, addPolynomial, ← coe_mapRingHom, + ← eval₂_hom, linePolynomial] + map_simp + lemma map_addX (x₁ x₂ L : R) : - (W.map φ).toAffine.addX (φ x₁) (φ x₂) (φ L) = φ (W.addX x₁ x₂ L) := by + (W.map f).toAffine.addX (f x₁) (f x₂) (f L) = f (W.addX x₁ x₂ L) := by simp only [addX] map_simp set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X WeierstrassCurve.Affine.map_addX lemma map_addY' (x₁ x₂ y₁ L : R) : - (W.map φ).toAffine.addY' (φ x₁) (φ x₂) (φ y₁) (φ L) = φ (W.addY' x₁ x₂ y₁ L) := by + (W.map f).toAffine.addY' (f x₁) (f x₂) (f y₁) (f L) = f (W.addY' x₁ x₂ y₁ L) := by simp only [addY', map_addX] map_simp set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y' WeierstrassCurve.Affine.map_addY' lemma map_addY (x₁ x₂ y₁ L : R) : - (W.map φ).toAffine.addY (φ x₁) (φ x₂) (φ y₁) (φ L) = φ (W.toAffine.addY x₁ x₂ y₁ L) := by + (W.map f).toAffine.addY (f x₁) (f x₂) (f y₁) (f L) = f (W.toAffine.addY x₁ x₂ y₁ L) := by simp only [addY, map_addY', map_addX, map_negY] set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y WeierstrassCurve.Affine.map_addY -lemma map_slope {F : Type u} [Field F] (W : Affine F) {K : Type v} [Field K] (φ : F →+* K) - (x₁ x₂ y₁ y₂ : F) : (W.map φ).toAffine.slope (φ x₁) (φ x₂) (φ y₁) (φ y₂) = - φ (W.slope x₁ x₂ y₁ y₂) := by +lemma map_slope {F : Type u} [Field F] (W : Affine F) {K : Type v} [Field K] (f : F →+* K) + (x₁ x₂ y₁ y₂ : F) : (W.map f).toAffine.slope (f x₁) (f x₂) (f y₁) (f y₂) = + f (W.slope x₁ x₂ y₁ y₂) := by by_cases hx : x₁ = x₂ · by_cases hy : y₁ = W.negY x₂ y₂ - · rw [slope_of_Yeq hx hy, slope_of_Yeq <| congr_arg _ hx, map_zero] - rw [hy, map_negY] - · rw [slope_of_Yne hx hy, slope_of_Yne <| congr_arg _ hx] - · simp only [negY] - map_simp - · rw [map_negY] - contrapose! hy - exact φ.injective hy - · rw [slope_of_Xne hx, slope_of_Xne] - · map_simp - · contrapose! hx - exact φ.injective hx + · rw [slope_of_Yeq (congr_arg f hx) <| by rw [hy, map_negY], slope_of_Yeq hx hy, map_zero] + · rw [slope_of_Yne (congr_arg f hx) <| W.map_negY f x₂ y₂ ▸ fun h => hy <| f.injective h, + map_negY, slope_of_Yne hx hy] + map_simp + · rw [slope_of_Xne fun h => hx <| f.injective h, slope_of_Xne hx] + map_simp #align weierstrass_curve.base_change_slope WeierstrassCurve.Affine.map_slope variable {R : Type r} [CommRing R] (W : Affine R) {S : Type s} [CommRing S] [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] - {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (ψ : A →ₐ[S] B) + {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (φ : A →ₐ[S] B) + +lemma baseChange_polynomial : (W.baseChange B).toAffine.polynomial = + (W.baseChange A).toAffine.polynomial.map (mapRingHom φ) := by + rw [← map_polynomial, map_baseChange] + +lemma baseChange_eval_polynomial (x : A[X]) (y : A) : + ((W.baseChange B).toAffine.polynomial.eval <| x.map φ).eval (φ y) = + φ (((W.baseChange A).toAffine.polynomial.eval x).eval y) := by + erw [← map_eval_polynomial, map_baseChange] + rfl -lemma baseChange_equation {ψ : A →ₐ[S] B} (hψ : Function.Injective ψ) (x y : A) : - (W.baseChange B).toAffine.Equation (ψ x) (ψ y) ↔ (W.baseChange A).toAffine.Equation x y := by - erw [← map_equation _ hψ, map_baseChange] +variable {φ} in +lemma baseChange_equation (hφ : Function.Injective φ) (x y : A) : + (W.baseChange B).toAffine.Equation (φ x) (φ y) ↔ (W.baseChange A).toAffine.Equation x y := by + erw [← map_equation _ hφ, map_baseChange] rfl #align weierstrass_curve.equation_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_equation -lemma baseChange_nonsingular {ψ : A →ₐ[S] B} (hψ : Function.Injective ψ) (x y : A) : - (W.baseChange B).toAffine.Nonsingular (ψ x) (ψ y) ↔ +lemma baseChange_polynomialX : (W.baseChange B).toAffine.polynomialX = + (W.baseChange A).toAffine.polynomialX.map (mapRingHom φ) := by + rw [← map_polynomialX, map_baseChange] + +lemma baseChange_eval_polynomialX (x : A[X]) (y : A) : + ((W.baseChange B).toAffine.polynomialX.eval <| x.map φ).eval (φ y) = + φ (((W.baseChange A).toAffine.polynomialX.eval x).eval y) := by + erw [← map_eval_polynomialX, map_baseChange] + rfl + +lemma baseChange_polynomialY : (W.baseChange B).toAffine.polynomialY = + (W.baseChange A).toAffine.polynomialY.map (mapRingHom φ) := by + rw [← map_polynomialY, map_baseChange] + +lemma baseChange_eval_polynomialY (x : A[X]) (y : A) : + ((W.baseChange B).toAffine.polynomialY.eval <| x.map φ).eval (φ y) = + φ (((W.baseChange A).toAffine.polynomialY.eval x).eval y) := by + erw [← map_eval_polynomialY, map_baseChange] + rfl + +variable {φ} in +lemma baseChange_nonsingular (hφ : Function.Injective φ) (x y : A) : + (W.baseChange B).toAffine.Nonsingular (φ x) (φ y) ↔ (W.baseChange A).toAffine.Nonsingular x y := by - erw [← map_nonsingular _ hψ, map_baseChange] + erw [← map_nonsingular _ hφ, map_baseChange] rfl #align weierstrass_curve.nonsingular_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_nonsingular +lemma baseChange_negPolynomial : + (W.baseChange B).toAffine.negPolynomial = + (W.baseChange A).toAffine.negPolynomial.map (mapRingHom φ) := by + rw [← map_negPolynomial, map_baseChange] + lemma baseChange_negY (x y : A) : - (W.baseChange B).toAffine.negY (ψ x) (ψ y) = ψ ((W.baseChange A).toAffine.negY x y) := by + (W.baseChange B).toAffine.negY (φ x) (φ y) = φ ((W.baseChange A).toAffine.negY x y) := by erw [← map_negY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_neg_Y_of_base_change WeierstrassCurve.Affine.baseChange_negY +lemma baseChange_addPolynomial (x y L : A) : + (W.baseChange B).toAffine.addPolynomial (φ x) (φ y) (φ L) = + ((W.baseChange A).toAffine.addPolynomial x y L).map φ := by + rw [← map_addPolynomial, map_baseChange] + rfl + lemma baseChange_addX (x₁ x₂ L : A) : - (W.baseChange B).toAffine.addX (ψ x₁) (ψ x₂) (ψ L) = - ψ ((W.baseChange A).toAffine.addX x₁ x₂ L) := by + (W.baseChange B).toAffine.addX (φ x₁) (φ x₂) (φ L) = + φ ((W.baseChange A).toAffine.addX x₁ x₂ L) := by erw [← map_addX, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X_of_base_change WeierstrassCurve.Affine.baseChange_addX lemma baseChange_addY' (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.addY' (ψ x₁) (ψ x₂) (ψ y₁) (ψ L) = - ψ ((W.baseChange A).toAffine.addY' x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.addY' (φ x₁) (φ x₂) (φ y₁) (φ L) = + φ ((W.baseChange A).toAffine.addY' x₁ x₂ y₁ L) := by erw [← map_addY', map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y'_of_base_change WeierstrassCurve.Affine.baseChange_addY' lemma baseChange_addY (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.addY (ψ x₁) (ψ x₂) (ψ y₁) (ψ L) = - ψ ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.addY (φ x₁) (φ x₂) (φ y₁) (φ L) = + φ ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by erw [← map_addY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y_of_base_change WeierstrassCurve.Affine.baseChange_addY variable {F : Type u} [Field F] [Algebra R F] [Algebra S F] [IsScalarTower R S F] - {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (ψ : F →ₐ[S] K) + {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (φ : F →ₐ[S] K) {L : Type w} [Field L] [Algebra R L] [Algebra S L] [IsScalarTower R S L] (χ : K →ₐ[S] L) lemma baseChange_slope (x₁ x₂ y₁ y₂ : F) : - (W.baseChange K).toAffine.slope (ψ x₁) (ψ x₂) (ψ y₁) (ψ y₂) = - ψ ((W.baseChange F).toAffine.slope x₁ x₂ y₁ y₂) := by + (W.baseChange K).toAffine.slope (φ x₁) (φ x₂) (φ y₁) (φ y₂) = + φ ((W.baseChange F).toAffine.slope x₁ x₂ y₁ y₂) := by erw [← map_slope, map_baseChange] rfl #align weierstrass_curve.base_change_slope_of_base_change WeierstrassCurve.Affine.baseChange_slope namespace Point -/-- The function from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `ψ : F →ₐ[S] K`, +/-- The function from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `φ : F →ₐ[S] K`, where `W` is defined over a subring of a ring `S`, and `F` and `K` are field extensions of `S`. -/ def mapFun : W⟮F⟯ → W⟮K⟯ | 0 => 0 - | some h => some <| (W.baseChange_nonsingular ψ.injective ..).mpr h + | some h => some <| (W.baseChange_nonsingular φ.injective ..).mpr h #align weierstrass_curve.point.of_base_change_fun WeierstrassCurve.Affine.Point.mapFun -/-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `ψ : F →ₐ[S] K`, +/-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `φ : F →ₐ[S] K`, where `W` is defined over a subring of a ring `S`, and `F` and `K` are field extensions of `S`. -/ def map : W⟮F⟯ →+ W⟮K⟯ where - toFun := mapFun W ψ + toFun := mapFun W φ map_zero' := rfl map_add' := by rintro (_ | @⟨x₁, y₁, _⟩) (_ | @⟨x₂, y₂, _⟩) @@ -911,39 +991,33 @@ def map : W⟮F⟯ →+ W⟮K⟯ where by_cases hx : x₁ = x₂ · by_cases hy : y₁ = (W.baseChange F).toAffine.negY x₂ y₂ · simp only [some_add_some_of_Yeq hx hy, mapFun] - rw [some_add_some_of_Yeq <| congr_arg _ hx] - rw [hy, baseChange_negY] + rw [some_add_some_of_Yeq (congr_arg _ hx) <| by rw [hy, baseChange_negY]] · simp only [some_add_some_of_Yne hx hy, mapFun] - rw [some_add_some_of_Yne <| congr_arg _ hx] - · simp only [some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] - · rw [baseChange_negY] - contrapose! hy - exact ψ.injective hy - · simp only [some_add_some_of_Xne hx, mapFun] - rw [some_add_some_of_Xne] - · simp only [some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] - · contrapose! hx - exact ψ.injective hx + rw [some_add_some_of_Yne (congr_arg _ hx) <| + by simpa only [baseChange_negY] using fun h => hy <| φ.injective h] + simp only [some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] + · simp only [some_add_some_of_Xne hx, mapFun, some_add_some_of_Xne fun h => hx <| φ.injective h, + some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] #align weierstrass_curve.point.of_base_change WeierstrassCurve.Affine.Point.map -lemma map_zero : map W ψ (0 : W⟮F⟯) = 0 := +lemma map_zero : map W φ (0 : W⟮F⟯) = 0 := rfl lemma map_some {x y : F} (h : (W.baseChange F).toAffine.Nonsingular x y) : - map W ψ (some h) = some ((W.baseChange_nonsingular ψ.injective ..).mpr h) := + map W φ (some h) = some ((W.baseChange_nonsingular φ.injective ..).mpr h) := rfl lemma map_id (P : W⟮F⟯) : map W (Algebra.ofId F F) P = P := by cases P <;> rfl -lemma map_map (P : W⟮F⟯) : map W χ (map W ψ P) = map W (χ.comp ψ) P := by +lemma map_map (P : W⟮F⟯) : map W χ (map W φ P) = map W (χ.comp φ) P := by cases P <;> rfl -lemma map_injective : Function.Injective <| map W ψ := by +lemma map_injective : Function.Injective <| map W φ := by rintro (_ | _) (_ | _) h any_goals contradiction · rfl - · simpa only [some.injEq] using ⟨ψ.injective (some.inj h).left, ψ.injective (some.inj h).right⟩ + · simpa only [some.injEq] using ⟨φ.injective (some.inj h).left, φ.injective (some.inj h).right⟩ #align weierstrass_curve.point.of_base_change_injective WeierstrassCurve.Affine.Point.map_injective variable (F K) From f38bd5e6b817d414c257045675c685a5aa1a9289 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 29 May 2024 23:26:15 +0100 Subject: [PATCH 26/47] Fix Group --- Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 3d1bdebc86f31c..0e1a764919e6f4 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -175,8 +175,9 @@ set_option linter.uppercaseLean3 false in lemma XYIdeal_add_eq (x₁ x₂ y₁ L : R) : XYIdeal W (W.addX x₁ x₂ L) (C <| W.addY x₁ x₂ y₁ L) = span {mk W <| W.negPolynomial - C (linePolynomial x₁ y₁ L)} ⊔ XIdeal W (W.addX x₁ x₂ L) := by simp only [XYIdeal, XIdeal, XClass, YClass, addY, addY', negY, negPolynomial, linePolynomial] - rw [sub_sub <| -Y, neg_sub_left Y, map_neg, span_singleton_neg, sup_comm, ← span_insert, - ← span_pair_add_mul_right <| mk W <| C <| C <| W.a₁ + L, ← _root_.map_mul, ← map_add] + rw [sub_sub <| -(Y : R[X][Y]), neg_sub_left (Y : R[X][Y]), map_neg, span_singleton_neg, sup_comm, + ← span_insert, ← span_pair_add_mul_right <| mk W <| C <| C <| W.a₁ + L, ← _root_.map_mul, + ← map_add] apply congr_arg (_ ∘ _ ∘ _ ∘ _) C_simp ring1 @@ -465,7 +466,7 @@ lemma norm_smul_basis (p q : R[X]) : lemma coe_norm_smul_basis (p q : R[X]) : Algebra.norm R[X] (p • (1 : W.CoordinateRing) + q • mk W Y) = - mk W ((C p + C q * X) * (C p + C q * (-Y - C (C W.a₁ * X + C W.a₃)))) := + mk W ((C p + C q * X) * (C p + C q * (-(Y : R[X][Y]) - C (C W.a₁ * X + C W.a₃)))) := AdjoinRoot.mk_eq_mk.mpr ⟨C q ^ 2, by simp only [norm_smul_basis, polynomial]; C_simp; ring1⟩ #align weierstrass_curve.coordinate_ring.coe_norm_smul_basis WeierstrassCurve.Affine.CoordinateRing.coe_norm_smul_basis From 092eeabb083a103161569f01b832cfb0bf0e816c Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 30 May 2024 11:45:30 +0100 Subject: [PATCH 27/47] Fix names --- .../EllipticCurve/Affine.lean | 195 +++++++++--------- .../EllipticCurve/Group.lean | 26 +-- 2 files changed, 111 insertions(+), 110 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 874a64194851c5..a777cd2c1d75e7 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -268,7 +268,10 @@ set_option linter.uppercaseLean3 false in #align weierstrass_curve.eval_polynomial_Y_zero WeierstrassCurve.Affine.eval_polynomialY_zero /-- The proposition that an affine point $(x, y)$ in `W` is nonsingular. -In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. -/ +In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. + +Note that this definition is only mathematically accurate for fields. +TODO: generalise this definition to be mathematically accurate a larger class of rings. -/ def Nonsingular (x y : R) : Prop := W.Equation x y ∧ ((W.polynomialX.eval <| C y).eval x ≠ 0 ∨ (W.polynomialY.eval <| C y).eval x ≠ 0) #align weierstrass_curve.nonsingular WeierstrassCurve.Affine.Nonsingular @@ -394,15 +397,15 @@ def addX (x₁ x₂ L : R) : R := set_option linter.uppercaseLean3 false in #align weierstrass_curve.add_X WeierstrassCurve.Affine.addX -/-- The $Y$-coordinate, before applying the final negation, of the addition of two affine points -$(x_1, y_1)$ and $(x_2, y_2)$, where the line through them is not vertical and has a slope of $L$. +/-- The $Y$-coordinate of the negated addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$, +where the line through them is not vertical and has a slope of $L$. This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ @[simp] -def addY' (x₁ x₂ y₁ L : R) : R := +def negAddY (x₁ x₂ y₁ L : R) : R := L * (W.addX x₁ x₂ L - x₁) + y₁ set_option linter.uppercaseLean3 false in -#align weierstrass_curve.add_Y' WeierstrassCurve.Affine.addY' +#align weierstrass_curve.add_Y' WeierstrassCurve.Affine.negAddY /-- The $Y$-coordinate of the addition of two affine points $(x_1, y_1)$ and $(x_2, y_2)$ in `W`, where the line through them is not vertical and has a slope of $L$. @@ -410,7 +413,7 @@ where the line through them is not vertical and has a slope of $L$. This depends on `W`, and has argument order: $x_1$, $x_2$, $y_1$, $L$. -/ @[simp] def addY (x₁ x₂ y₁ L : R) : R := - W.negY (W.addX x₁ x₂ L) (W.addY' x₁ x₂ y₁ L) + W.negY (W.addX x₁ x₂ L) (W.negAddY x₁ x₂ y₁ L) set_option linter.uppercaseLean3 false in #align weierstrass_curve.add_Y WeierstrassCurve.Affine.addY @@ -427,9 +430,9 @@ lemma nonsingular_neg_iff (x y : R) : W.Nonsingular x (W.negY x y) ↔ W.Nonsing #align weierstrass_curve.nonsingular_neg_iff WeierstrassCurve.Affine.nonsingular_neg_iff lemma equation_add_iff (x₁ x₂ y₁ L : R) : - W.Equation (W.addX x₁ x₂ L) (W.addY' x₁ x₂ y₁ L) ↔ + W.Equation (W.addX x₁ x₂ L) (W.negAddY x₁ x₂ y₁ L) ↔ (W.addPolynomial x₁ y₁ L).eval (W.addX x₁ x₂ L) = 0 := by - rw [Equation, addY', addPolynomial, linePolynomial, polynomial] + rw [Equation, negAddY, addPolynomial, linePolynomial, polynomial] eval_simp #align weierstrass_curve.equation_add_iff WeierstrassCurve.Affine.equation_add_iff @@ -453,11 +456,11 @@ lemma nonsingular_neg {x y : R} (h : W.Nonsingular x y) : W.Nonsingular x <| W.n (W.nonsingular_neg_iff ..).mpr h #align weierstrass_curve.nonsingular_neg WeierstrassCurve.Affine.nonsingular_neg -lemma nonsingular_add_of_eval_derivative_ne_zero {x₁ x₂ y₁ L : R} - (hx' : W.Equation (W.addX x₁ x₂ L) (W.addY' x₁ x₂ y₁ L)) +lemma nonsingular_negAdd_of_eval_derivative_ne_zero {x₁ x₂ y₁ L : R} + (hx' : W.Equation (W.addX x₁ x₂ L) (W.negAddY x₁ x₂ y₁ L)) (hx : (W.addPolynomial x₁ y₁ L).derivative.eval (W.addX x₁ x₂ L) ≠ 0) : - W.Nonsingular (W.addX x₁ x₂ L) (W.addY' x₁ x₂ y₁ L) := by - rw [Nonsingular, and_iff_right hx', addY', polynomialX, polynomialY] + W.Nonsingular (W.addX x₁ x₂ L) (W.negAddY x₁ x₂ y₁ L) := by + rw [Nonsingular, and_iff_right hx', negAddY, polynomialX, polynomialY] eval_simp contrapose! hx rw [addPolynomial, linePolynomial, polynomial] @@ -466,7 +469,7 @@ lemma nonsingular_add_of_eval_derivative_ne_zero {x₁ x₂ y₁ L : R} simp only [zero_add, add_zero, sub_zero, zero_mul, mul_one] eval_simp linear_combination (norm := (norm_num1; ring1)) hx.left + L * hx.right -#align weierstrass_curve.nonsingular_add_of_eval_derivative_ne_zero WeierstrassCurve.Affine.nonsingular_add_of_eval_derivative_ne_zero +#align weierstrass_curve.nonsingular_add_of_eval_derivative_ne_zero WeierstrassCurve.Affine.nonsingular_negAdd_of_eval_derivative_ne_zero end Ring @@ -493,58 +496,59 @@ noncomputable def slope {F : Type u} [Field F] (W : Affine F) (x₁ x₂ y₁ y variable {F : Type u} [Field F] {W : Affine F} @[simp] -lemma slope_of_Yeq {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ = W.negY x₂ y₂) : +lemma slope_of_Y_eq {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ = W.negY x₂ y₂) : W.slope x₁ x₂ y₁ y₂ = 0 := by rw [slope, if_pos hx, if_pos hy] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.slope_of_Yeq WeierstrassCurve.Affine.slope_of_Yeq +#align weierstrass_curve.slope_of_Yeq WeierstrassCurve.Affine.slope_of_Y_eq @[simp] -lemma slope_of_Yne {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : W.slope x₁ x₂ y₁ y₂ = - (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄ - W.a₁ * y₁) / (y₁ - W.negY x₁ y₁) := by +lemma slope_of_Y_ne {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : + W.slope x₁ x₂ y₁ y₂ = + (3 * x₁ ^ 2 + 2 * W.a₂ * x₁ + W.a₄ - W.a₁ * y₁) / (y₁ - W.negY x₁ y₁) := by rw [slope, if_pos hx, if_neg hy] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.slope_of_Yne WeierstrassCurve.Affine.slope_of_Yne +#align weierstrass_curve.slope_of_Yne WeierstrassCurve.Affine.slope_of_Y_ne @[simp] -lemma slope_of_Xne {x₁ x₂ y₁ y₂ : F} (hx : x₁ ≠ x₂) : +lemma slope_of_X_ne {x₁ x₂ y₁ y₂ : F} (hx : x₁ ≠ x₂) : W.slope x₁ x₂ y₁ y₂ = (y₁ - y₂) / (x₁ - x₂) := by rw [slope, if_neg hx] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.slope_of_Xne WeierstrassCurve.Affine.slope_of_Xne +#align weierstrass_curve.slope_of_Xne WeierstrassCurve.Affine.slope_of_X_ne -lemma slope_of_Yne_eq_eval {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : +lemma slope_of_Y_ne_eq_eval {x₁ x₂ y₁ y₂ : F} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : W.slope x₁ x₂ y₁ y₂ = -(W.polynomialX.eval <| C y₁).eval x₁ / (W.polynomialY.eval <| C y₁).eval x₁ := by - rw [slope_of_Yne hx hy, eval_polynomialX, neg_sub] + rw [slope_of_Y_ne hx hy, eval_polynomialX, neg_sub] congr 1 rw [negY, eval_polynomialY] ring1 set_option linter.uppercaseLean3 false in -#align weierstrass_curve.slope_of_Yne_eq_eval WeierstrassCurve.Affine.slope_of_Yne_eq_eval +#align weierstrass_curve.slope_of_Yne_eq_eval WeierstrassCurve.Affine.slope_of_Y_ne_eq_eval -lemma Yeq_of_Xeq {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) (hx : x₁ = x₂) : - y₁ = y₂ ∨ y₁ = W.negY x₂ y₂ := by +lemma Y_eq_of_X_eq {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) + (hx : x₁ = x₂) : y₁ = y₂ ∨ y₁ = W.negY x₂ y₂ := by rw [equation_iff] at h₁ h₂ rw [← sub_eq_zero, ← sub_eq_zero (a := y₁), ← mul_eq_zero, negY] linear_combination (norm := (rw [hx]; ring1)) h₁ - h₂ set_option linter.uppercaseLean3 false in -#align weierstrass_curve.Yeq_of_Xeq WeierstrassCurve.Affine.Yeq_of_Xeq +#align weierstrass_curve.Yeq_of_Xeq WeierstrassCurve.Affine.Y_eq_of_X_eq -lemma Yeq_of_Yne {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) (hx : x₁ = x₂) +lemma Y_eq_of_Y_ne {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : y₁ = y₂ := - (Yeq_of_Xeq h₁ h₂ hx).resolve_right hy + (Y_eq_of_X_eq h₁ h₂ hx).resolve_right hy set_option linter.uppercaseLean3 false in -#align weierstrass_curve.Yeq_of_Yne WeierstrassCurve.Affine.Yeq_of_Yne +#align weierstrass_curve.Yeq_of_Yne WeierstrassCurve.Affine.Y_eq_of_Y_ne lemma addPolynomial_slope {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : W.addPolynomial x₁ y₁ (W.slope x₁ x₂ y₁ y₂) = -((X - C x₁) * (X - C x₂) * (X - C (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂))) := by rw [addPolynomial_eq, neg_inj, Cubic.prod_X_sub_C_eq, Cubic.toPoly_injective] by_cases hx : x₁ = x₂ - · rcases hx, Yeq_of_Yne h₁ h₂ hx (hxy hx) with ⟨rfl, rfl⟩ + · rcases hx, Y_eq_of_Y_ne h₁ h₂ hx (hxy hx) with ⟨rfl, rfl⟩ rw [equation_iff] at h₁ h₂ - rw [slope_of_Yne rfl <| hxy rfl] + rw [slope_of_Y_ne rfl <| hxy rfl] rw [negY, ← sub_ne_zero] at hxy ext · rfl @@ -554,7 +558,7 @@ lemma addPolynomial_slope {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ ring1 · linear_combination (norm := (field_simp [hxy rfl]; ring1)) -h₁ · rw [equation_iff] at h₁ h₂ - rw [slope_of_Xne hx] + rw [slope_of_X_ne hx] rw [← sub_eq_zero] at hx ext · rfl @@ -566,21 +570,20 @@ lemma addPolynomial_slope {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ linear_combination (norm := (field_simp [hx]; ring1)) x₂ * h₁ - x₁ * h₂ #align weierstrass_curve.add_polynomial_slope WeierstrassCurve.Affine.addPolynomial_slope -/-- The addition of two affine points in `W` on a sloped line, -before applying the final negation that maps $Y$ to $-Y - a_1X - a_3$, lies in `W`. -/ -lemma equation_add' {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) - (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : - W.Equation (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.addY' x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := by +/-- The negated addition of two affine points in `W` on a sloped line lies in `W`. -/ +lemma equation_negAdd {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) + (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : W.Equation + (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.negAddY x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := by rw [equation_add_iff, addPolynomial_slope h₁ h₂ hxy] eval_simp rw [neg_eq_zero, sub_self, mul_zero] -#align weierstrass_curve.equation_add' WeierstrassCurve.Affine.equation_add' +#align weierstrass_curve.equation_add' WeierstrassCurve.Affine.equation_negAdd /-- The addition of two affine points in `W` on a sloped line lies in `W`. -/ lemma equation_add {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) (h₂ : W.Equation x₂ y₂) (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : W.Equation (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.addY x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := - equation_neg <| equation_add' h₁ h₂ hxy + equation_neg <| equation_negAdd h₁ h₂ hxy #align weierstrass_curve.equation_add WeierstrassCurve.Affine.equation_add lemma derivative_addPolynomial_slope {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) @@ -593,32 +596,30 @@ lemma derivative_addPolynomial_slope {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equatio ring1 #align weierstrass_curve.derivative_add_polynomial_slope WeierstrassCurve.Affine.derivative_addPolynomial_slope -/-- The addition of two nonsingular affine points in `W` on a sloped line, -before applying the final negation that maps $Y$ to $-Y - a_1X - a_3$, is nonsingular. -/ -lemma nonsingular_add' {x₁ x₂ y₁ y₂ : F} (h₁ : W.Nonsingular x₁ y₁) (h₂ : W.Nonsingular x₂ y₂) - (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : - W.Nonsingular (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) - (W.addY' x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := by +/-- The negated addition of two nonsingular affine points in `W` on a sloped line is nonsingular. -/ +lemma nonsingular_negAdd {x₁ x₂ y₁ y₂ : F} (h₁ : W.Nonsingular x₁ y₁) (h₂ : W.Nonsingular x₂ y₂) + (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : W.Nonsingular + (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.negAddY x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := by by_cases hx₁ : W.addX x₁ x₂ (W.slope x₁ x₂ y₁ y₂) = x₁ - · rwa [addY', hx₁, sub_self, mul_zero, zero_add] + · rwa [negAddY, hx₁, sub_self, mul_zero, zero_add] · by_cases hx₂ : W.addX x₁ x₂ (W.slope x₁ x₂ y₁ y₂) = x₂ · by_cases hx : x₁ = x₂ · subst hx contradiction - · rwa [addY', ← neg_sub, mul_neg, hx₂, slope_of_Xne hx, + · rwa [negAddY, ← neg_sub, mul_neg, hx₂, slope_of_X_ne hx, div_mul_cancel₀ _ <| sub_ne_zero_of_ne hx, neg_sub, sub_add_cancel] - · apply nonsingular_add_of_eval_derivative_ne_zero <| equation_add' h₁.1 h₂.1 hxy + · apply nonsingular_negAdd_of_eval_derivative_ne_zero <| equation_negAdd h₁.1 h₂.1 hxy rw [derivative_addPolynomial_slope h₁.left h₂.left hxy] eval_simp simpa only [neg_ne_zero, sub_self, mul_zero, add_zero] using mul_ne_zero (sub_ne_zero_of_ne hx₁) (sub_ne_zero_of_ne hx₂) -#align weierstrass_curve.nonsingular_add' WeierstrassCurve.Affine.nonsingular_add' +#align weierstrass_curve.nonsingular_add' WeierstrassCurve.Affine.nonsingular_negAdd /-- The addition of two nonsingular affine points in `W` on a sloped line is nonsingular. -/ lemma nonsingular_add {x₁ x₂ y₁ y₂ : F} (h₁ : W.Nonsingular x₁ y₁) (h₂ : W.Nonsingular x₂ y₂) (hxy : x₁ = x₂ → y₁ ≠ W.negY x₂ y₂) : W.Nonsingular (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.addY x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := - nonsingular_neg <| nonsingular_add' h₁ h₂ hxy + nonsingular_neg <| nonsingular_negAdd h₁ h₂ hxy #align weierstrass_curve.nonsingular_add WeierstrassCurve.Affine.nonsingular_add end Field @@ -710,59 +711,59 @@ noncomputable instance instAddZeroClassPoint : AddZeroClass W.Point := ⟨by rintro (_ | _) <;> rfl, by rintro (_ | _) <;> rfl⟩ @[simp] -lemma some_add_some_of_Yeq {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} +lemma add_of_Y_eq {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} (hx : x₁ = x₂) (hy : y₁ = W.negY x₂ y₂) : some h₁ + some h₂ = 0 := by simp only [← add_def, add, dif_pos hx, dif_pos hy] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_some_of_Yeq WeierstrassCurve.Affine.Point.some_add_some_of_Yeq +#align weierstrass_curve.point.some_add_some_of_Yeq WeierstrassCurve.Affine.Point.add_of_Y_eq @[simp] -lemma some_add_self_of_Yeq {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ = W.negY x₁ y₁) : +lemma add_self_of_Y_eq {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ = W.negY x₁ y₁) : some h₁ + some h₁ = 0 := - some_add_some_of_Yeq rfl hy + add_of_Y_eq rfl hy set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_self_of_Yeq WeierstrassCurve.Affine.Point.some_add_self_of_Yeq +#align weierstrass_curve.point.some_add_self_of_Yeq WeierstrassCurve.Affine.Point.add_self_of_Y_eq @[simp] -lemma some_add_some_of_Yne {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} +lemma add_of_Y_ne {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : some h₁ + some h₂ = some (nonsingular_add h₁ h₂ fun _ => hy) := by simp only [← add_def, add, dif_pos hx, dif_neg hy] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_some_of_Yne WeierstrassCurve.Affine.Point.some_add_some_of_Yne +#align weierstrass_curve.point.some_add_some_of_Yne WeierstrassCurve.Affine.Point.add_of_Y_ne -lemma some_add_some_of_Yne' {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} +lemma add_of_Y_ne' {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} (hx : x₁ = x₂) (hy : y₁ ≠ W.negY x₂ y₂) : - some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ fun _ => hy) := - some_add_some_of_Yne hx hy + some h₁ + some h₂ = -some (nonsingular_negAdd h₁ h₂ fun _ => hy) := + add_of_Y_ne hx hy set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_some_of_Yne' WeierstrassCurve.Affine.Point.some_add_some_of_Yne' +#align weierstrass_curve.point.some_add_some_of_Yne' WeierstrassCurve.Affine.Point.add_of_Y_ne' @[simp] -lemma some_add_self_of_Yne {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ ≠ W.negY x₁ y₁) : +lemma add_self_of_Y_ne {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ ≠ W.negY x₁ y₁) : some h₁ + some h₁ = some (nonsingular_add h₁ h₁ fun _ => hy) := - some_add_some_of_Yne rfl hy + add_of_Y_ne rfl hy set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_self_of_Yne WeierstrassCurve.Affine.Point.some_add_self_of_Yne +#align weierstrass_curve.point.some_add_self_of_Yne WeierstrassCurve.Affine.Point.add_self_of_Y_ne -lemma some_add_self_of_Yne' {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ ≠ W.negY x₁ y₁) : - some h₁ + some h₁ = -some (nonsingular_add' h₁ h₁ fun _ => hy) := - some_add_some_of_Yne rfl hy +lemma add_self_of_Y_ne' {x₁ y₁ : F} {h₁ : W.Nonsingular x₁ y₁} (hy : y₁ ≠ W.negY x₁ y₁) : + some h₁ + some h₁ = -some (nonsingular_negAdd h₁ h₁ fun _ => hy) := + add_of_Y_ne rfl hy set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_self_of_Yne' WeierstrassCurve.Affine.Point.some_add_self_of_Yne' +#align weierstrass_curve.point.some_add_self_of_Yne' WeierstrassCurve.Affine.Point.add_self_of_Y_ne' @[simp] -lemma some_add_some_of_Xne {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} +lemma add_of_X_ne {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} (hx : x₁ ≠ x₂) : some h₁ + some h₂ = some (nonsingular_add h₁ h₂ fun h => (hx h).elim) := by simp only [← add_def, add, dif_neg hx] set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_some_of_Xne WeierstrassCurve.Affine.Point.some_add_some_of_Xne +#align weierstrass_curve.point.some_add_some_of_Xne WeierstrassCurve.Affine.Point.add_of_X_ne -lemma some_add_some_of_Xne' {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} - (hx : x₁ ≠ x₂) : some h₁ + some h₂ = -some (nonsingular_add' h₁ h₂ fun h => (hx h).elim) := - some_add_some_of_Xne hx +lemma add_of_X_ne' {x₁ x₂ y₁ y₂ : F} {h₁ : W.Nonsingular x₁ y₁} {h₂ : W.Nonsingular x₂ y₂} + (hx : x₁ ≠ x₂) : some h₁ + some h₂ = -some (nonsingular_negAdd h₁ h₂ fun h => (hx h).elim) := + add_of_X_ne hx set_option linter.uppercaseLean3 false in -#align weierstrass_curve.point.some_add_some_of_Xne' WeierstrassCurve.Affine.Point.some_add_some_of_Xne' +#align weierstrass_curve.point.some_add_some_of_Xne' WeierstrassCurve.Affine.Point.add_of_X_ne' end Point @@ -843,16 +844,16 @@ lemma map_addX (x₁ x₂ L : R) : set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X WeierstrassCurve.Affine.map_addX -lemma map_addY' (x₁ x₂ y₁ L : R) : - (W.map f).toAffine.addY' (f x₁) (f x₂) (f y₁) (f L) = f (W.addY' x₁ x₂ y₁ L) := by - simp only [addY', map_addX] +lemma map_negAddY (x₁ x₂ y₁ L : R) : + (W.map f).toAffine.negAddY (f x₁) (f x₂) (f y₁) (f L) = f (W.negAddY x₁ x₂ y₁ L) := by + simp only [negAddY, map_addX] map_simp set_option linter.uppercaseLean3 false in -#align weierstrass_curve.base_change_add_Y' WeierstrassCurve.Affine.map_addY' +#align weierstrass_curve.base_change_add_Y' WeierstrassCurve.Affine.map_negAddY lemma map_addY (x₁ x₂ y₁ L : R) : (W.map f).toAffine.addY (f x₁) (f x₂) (f y₁) (f L) = f (W.toAffine.addY x₁ x₂ y₁ L) := by - simp only [addY, map_addY', map_addX, map_negY] + simp only [addY, map_negAddY, map_addX, map_negY] set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y WeierstrassCurve.Affine.map_addY @@ -861,11 +862,11 @@ lemma map_slope {F : Type u} [Field F] (W : Affine F) {K : Type v} [Field K] (f f (W.slope x₁ x₂ y₁ y₂) := by by_cases hx : x₁ = x₂ · by_cases hy : y₁ = W.negY x₂ y₂ - · rw [slope_of_Yeq (congr_arg f hx) <| by rw [hy, map_negY], slope_of_Yeq hx hy, map_zero] - · rw [slope_of_Yne (congr_arg f hx) <| W.map_negY f x₂ y₂ ▸ fun h => hy <| f.injective h, - map_negY, slope_of_Yne hx hy] + · rw [slope_of_Y_eq (congr_arg f hx) <| by rw [hy, map_negY], slope_of_Y_eq hx hy, map_zero] + · rw [slope_of_Y_ne (congr_arg f hx) <| W.map_negY f x₂ y₂ ▸ fun h => hy <| f.injective h, + map_negY, slope_of_Y_ne hx hy] map_simp - · rw [slope_of_Xne fun h => hx <| f.injective h, slope_of_Xne hx] + · rw [slope_of_X_ne fun h => hx <| f.injective h, slope_of_X_ne hx] map_simp #align weierstrass_curve.base_change_slope WeierstrassCurve.Affine.map_slope @@ -944,13 +945,13 @@ lemma baseChange_addX (x₁ x₂ L : A) : set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X_of_base_change WeierstrassCurve.Affine.baseChange_addX -lemma baseChange_addY' (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.addY' (φ x₁) (φ x₂) (φ y₁) (φ L) = - φ ((W.baseChange A).toAffine.addY' x₁ x₂ y₁ L) := by - erw [← map_addY', map_baseChange] +lemma baseChange_negAddY (x₁ x₂ y₁ L : A) : + (W.baseChange B).toAffine.negAddY (φ x₁) (φ x₂) (φ y₁) (φ L) = + φ ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by + erw [← map_negAddY, map_baseChange] rfl set_option linter.uppercaseLean3 false in -#align weierstrass_curve.base_change_add_Y'_of_base_change WeierstrassCurve.Affine.baseChange_addY' +#align weierstrass_curve.base_change_add_Y'_of_base_change WeierstrassCurve.Affine.baseChange_negAddY lemma baseChange_addY (x₁ x₂ y₁ L : A) : (W.baseChange B).toAffine.addY (φ x₁) (φ x₂) (φ y₁) (φ L) = @@ -962,7 +963,7 @@ set_option linter.uppercaseLean3 false in variable {F : Type u} [Field F] [Algebra R F] [Algebra S F] [IsScalarTower R S F] {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (φ : F →ₐ[S] K) - {L : Type w} [Field L] [Algebra R L] [Algebra S L] [IsScalarTower R S L] (χ : K →ₐ[S] L) + {L : Type w} [Field L] [Algebra R L] [Algebra S L] [IsScalarTower R S L] (ψ : K →ₐ[S] L) lemma baseChange_slope (x₁ x₂ y₁ y₂ : F) : (W.baseChange K).toAffine.slope (φ x₁) (φ x₂) (φ y₁) (φ y₂) = @@ -990,13 +991,13 @@ def map : W⟮F⟯ →+ W⟮K⟯ where any_goals rfl by_cases hx : x₁ = x₂ · by_cases hy : y₁ = (W.baseChange F).toAffine.negY x₂ y₂ - · simp only [some_add_some_of_Yeq hx hy, mapFun] - rw [some_add_some_of_Yeq (congr_arg _ hx) <| by rw [hy, baseChange_negY]] - · simp only [some_add_some_of_Yne hx hy, mapFun] - rw [some_add_some_of_Yne (congr_arg _ hx) <| + · simp only [add_of_Y_eq hx hy, mapFun] + rw [add_of_Y_eq (congr_arg _ hx) <| by rw [hy, baseChange_negY]] + · simp only [add_of_Y_ne hx hy, mapFun] + rw [add_of_Y_ne (congr_arg _ hx) <| by simpa only [baseChange_negY] using fun h => hy <| φ.injective h] simp only [some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] - · simp only [some_add_some_of_Xne hx, mapFun, some_add_some_of_Xne fun h => hx <| φ.injective h, + · simp only [add_of_X_ne hx, mapFun, add_of_X_ne fun h => hx <| φ.injective h, some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] #align weierstrass_curve.point.of_base_change WeierstrassCurve.Affine.Point.map @@ -1010,7 +1011,7 @@ lemma map_some {x y : F} (h : (W.baseChange F).toAffine.Nonsingular x y) : lemma map_id (P : W⟮F⟯) : map W (Algebra.ofId F F) P = P := by cases P <;> rfl -lemma map_map (P : W⟮F⟯) : map W χ (map W φ P) = map W (χ.comp φ) P := by +lemma map_map (P : W⟮F⟯) : map W ψ (map W φ P) = map W (ψ.comp φ) P := by cases P <;> rfl lemma map_injective : Function.Injective <| map W φ := by @@ -1030,9 +1031,9 @@ abbrev baseChange [Algebra F K] [IsScalarTower R F K] : W⟮F⟯ →+ W⟮K⟯ : variable {F K} lemma map_baseChange [Algebra F K] [IsScalarTower R F K] [Algebra F L] [IsScalarTower R F L] - (χ : K →ₐ[F] L) (P : W⟮F⟯) : map W χ (baseChange W F K P) = baseChange W F L P := by + (ψ : K →ₐ[F] L) (P : W⟮F⟯) : map W ψ (baseChange W F K P) = baseChange W F L P := by have : Subsingleton (F →ₐ[F] L) := inferInstance - convert map_map W (Algebra.ofId F K) χ P + convert map_map W (Algebra.ofId F K) ψ P end Point diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 0e1a764919e6f4..4e2e7a8cd7eca4 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -174,7 +174,7 @@ set_option linter.uppercaseLean3 false in lemma XYIdeal_add_eq (x₁ x₂ y₁ L : R) : XYIdeal W (W.addX x₁ x₂ L) (C <| W.addY x₁ x₂ y₁ L) = span {mk W <| W.negPolynomial - C (linePolynomial x₁ y₁ L)} ⊔ XIdeal W (W.addX x₁ x₂ L) := by - simp only [XYIdeal, XIdeal, XClass, YClass, addY, addY', negY, negPolynomial, linePolynomial] + simp only [XYIdeal, XIdeal, XClass, YClass, addY, negAddY, negY, negPolynomial, linePolynomial] rw [sub_sub <| -(Y : R[X][Y]), neg_sub_left (Y : R[X][Y]), map_neg, span_singleton_neg, sup_comm, ← span_insert, ← span_pair_add_mul_right <| mk W <| C <| C <| W.a₁ + L, ← _root_.map_mul, ← map_add] @@ -205,9 +205,9 @@ lemma XYIdeal_eq₂ {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁) XYIdeal W x₂ (C y₂) = XYIdeal W x₂ (linePolynomial x₁ y₁ <| W.slope x₁ x₂ y₁ y₂) := by have hy₂ : y₂ = (linePolynomial x₁ y₁ <| W.slope x₁ x₂ y₁ y₂).eval x₂ := by by_cases hx : x₁ = x₂ - · rcases hx, Yeq_of_Yne h₁ h₂ hx <| hxy hx with ⟨rfl, rfl⟩ + · rcases hx, Y_eq_of_Y_ne h₁ h₂ hx <| hxy hx with ⟨rfl, rfl⟩ field_simp [linePolynomial, sub_ne_zero_of_ne <| hxy rfl] - · field_simp [linePolynomial, slope_of_Xne hx, sub_ne_zero_of_ne hx] + · field_simp [linePolynomial, slope_of_X_ne hx, sub_ne_zero_of_ne hx] ring1 nth_rw 1 [hy₂] simp only [XYIdeal, XClass, YClass, linePolynomial] @@ -282,7 +282,7 @@ lemma XYIdeal_mul_XYIdeal {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ mem_map_iff_of_surjective _ AdjoinRoot.mk_surjective, ← span_insert, mem_span_insert', mem_span_singleton'] by_cases hx : x₁ = x₂ - · rcases hx, Yeq_of_Yne h₁ h₂ hx (hxy hx) with ⟨rfl, rfl⟩ + · rcases hx, Y_eq_of_Y_ne h₁ h₂ hx (hxy hx) with ⟨rfl, rfl⟩ let y := (y₁ - W.negY x₁ y₁) ^ 2 replace hxy := pow_ne_zero 2 <| sub_ne_zero_of_ne <| hxy rfl refine ⟨1 + C (C <| y⁻¹ * 4) * W.polynomial, @@ -554,12 +554,12 @@ noncomputable def toClass : W.Point →+ Additive (ClassGroup W.CoordinateRing) by_cases hx : x₁ = x₂ · by_cases hy : y₁ = W.negY x₂ y₂ · substs hx hy - simpa only [some_add_some_of_Yeq rfl rfl] using - (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal'_of_Yeq h₂).symm - · simpa only [some_add_some_of_Yne hx hy] using - (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal' h₁ h₂ fun _ => hy).symm - · simpa only [some_add_some_of_Xne hx] using - (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal' h₁ h₂ fun h => (hx h).elim).symm + rw [add_of_Y_eq rfl rfl] + exact (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal'_of_Yeq h₂).symm + · rw [add_of_Y_ne hx hy] + exact (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal' h₁ h₂ fun _ => hy).symm + · rw [add_of_X_ne hx] + exact (CoordinateRing.mk_XYIdeal'_mul_mk_XYIdeal' h₁ h₂ fun h => (hx h).elim).symm #align weierstrass_curve.point.to_class WeierstrassCurve.Affine.Point.toClass -- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error @@ -583,11 +583,11 @@ lemma add_eq_zero (P Q : W.Point) : P + Q = 0 ↔ P = -Q := by by_cases hx : x₁ = x₂ · by_cases hy : y₁ = W.negY x₂ y₂ · exact ⟨hx, hy⟩ - · rw [some_add_some_of_Yne hx hy] at h + · rw [add_of_Y_ne hx hy] at h contradiction - · rw [some_add_some_of_Xne hx] at h + · rw [add_of_X_ne hx] at h contradiction - · exact fun ⟨hx, hy⟩ => some_add_some_of_Yeq hx hy + · exact fun ⟨hx, hy⟩ => add_of_Y_eq hx hy #align weierstrass_curve.point.add_eq_zero WeierstrassCurve.Affine.Point.add_eq_zero -- Porting note (#10619): removed `@[simp]` to avoid a `simpNF` linter error From de9d1708b0d604e4a39c4ca7242e4414ea2d3468 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 30 May 2024 11:49:42 +0100 Subject: [PATCH 28/47] Fix sentence --- Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index a777cd2c1d75e7..9b07f4674a8fc8 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -271,7 +271,7 @@ set_option linter.uppercaseLean3 false in In other words, either $W_X(x, y) \ne 0$ or $W_Y(x, y) \ne 0$. Note that this definition is only mathematically accurate for fields. -TODO: generalise this definition to be mathematically accurate a larger class of rings. -/ +TODO: generalise this definition to be mathematically accurate for a larger class of rings. -/ def Nonsingular (x y : R) : Prop := W.Equation x y ∧ ((W.polynomialX.eval <| C y).eval x ≠ 0 ∨ (W.polynomialY.eval <| C y).eval x ≠ 0) #align weierstrass_curve.nonsingular WeierstrassCurve.Affine.Nonsingular From 6495e45ff0fff6680448406a8efa71c32656c8e6 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 30 May 2024 14:54:22 +0100 Subject: [PATCH 29/47] Fix docstring --- Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 9b07f4674a8fc8..e12f6e78597b16 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -40,11 +40,13 @@ The set of nonsingular rational points forms an abelian group under a secant-and and has slope $\ell := (3x_1^2 + 2a_2x_1 + a_4 - a_1y_1) / (2y_1 + a_1x_1 + a_3)$. * Otherwise $x_1 \ne x_2$, then `L` is the secant of `W` through `P` and `Q`, and has slope $\ell := (y_1 - y_2) / (x_1 - x_2)$. + In the latter two cases, the $X$-coordinate of `P + Q` is then the unique third solution of the equation obtained by substituting the line $Y = \ell(X - x_1) + y_1$ into the Weierstrass equation, and can be written down explicitly as $x := \ell^2 + a_1\ell - a_2 - x_1 - x_2$ by inspecting the $X^2$ terms. The $Y$-coordinate of `P + Q`, after applying the final negation that maps $Y$ to $-Y - a_1X - a_3$, is precisely $y := -(\ell(x - x_1) + y_1) - a_1x - a_3$. + The group law on this set is then uniquely determined by these constructions. ## Main definitions From 4c05448d09df53a45988c0fb8d85696462b24004 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 30 May 2024 19:44:52 +0100 Subject: [PATCH 30/47] Add maps for division polynomials --- .../EllipticCurve/DivisionPolynomial.lean | 49 +++++++++++++++++++ .../EllipticDivisibilitySequence.lean | 19 ++++++- 2 files changed, 67 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index 1956b8d0818925..4ccc01d4d0c4ef 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -93,6 +93,12 @@ open Polynomial PolynomialPolynomial local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) +local macro "map_simp" : tactic => + `(tactic| simp only [map_ofNat, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀, + Polynomial.map_ofNat, Polynomial.map_one, map_C, map_X, Polynomial.map_neg, Polynomial.map_add, + Polynomial.map_sub, Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_div, coe_mapRingHom, + apply_ite <| mapRingHom _, WeierstrassCurve.map]) + universe u v namespace WeierstrassCurve @@ -452,4 +458,47 @@ lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by rw [WeierstrassCurve.φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, ← neg_add', ψ_neg, neg_mul_neg, mul_comm <| W.ψ _, WeierstrassCurve.φ] +lemma map_ψ₂ : (W.map f).ψ₂ = W.ψ₂.map (mapRingHom f) := by + simp only [WeierstrassCurve.ψ₂, Affine.map_polynomialY] + +lemma map_Ψ₂Sq : (W.map f).Ψ₂Sq = W.Ψ₂Sq.map f := by + simp only [WeierstrassCurve.Ψ₂Sq, map_b₂, map_b₄, map_b₆] + map_simp + +lemma map_Ψ₃ : (W.map f).Ψ₃ = W.Ψ₃.map f := by + simp only [WeierstrassCurve.Ψ₃, map_b₂, map_b₄, map_b₆, map_b₈] + map_simp + +lemma map_Ψ₄' : (W.map f).Ψ₄' = W.Ψ₄'.map f := by + simp only [WeierstrassCurve.Ψ₄', map_b₂, map_b₄, map_b₆, map_b₈] + map_simp + +lemma map_Ψ'' (n : ℕ) : (W.map f).Ψ'' n = (W.Ψ'' n).map f := by + simp only [WeierstrassCurve.Ψ'', map_Ψ₂Sq, map_Ψ₃, map_Ψ₄', ← coe_mapRingHom, map_preNormEDS'] + map_simp + +lemma map_Ψ' (n : ℤ) : (W.map f).Ψ' n = (W.Ψ' n).map f := by + simp only [WeierstrassCurve.Ψ', map_Ψ₂Sq, map_Ψ₃, map_Ψ₄', ← coe_mapRingHom, map_preNormEDS] + map_simp + +lemma map_ΨSq (n : ℤ) : (W.map f).ΨSq n = (W.ΨSq n).map f := by + simp only [WeierstrassCurve.ΨSq, map_Ψ', map_Ψ₂Sq, ← coe_mapRingHom] + map_simp + +lemma map_Ψ (n : ℤ) : (W.map f).Ψ n = (W.Ψ n).map (mapRingHom f) := by + simp only [WeierstrassCurve.Ψ, map_Ψ', map_ψ₂, ← coe_mapRingHom] + map_simp + +lemma map_Φ (n : ℤ) : (W.map f).Φ n = (W.Φ n).map f := by + simp only [WeierstrassCurve.Φ, map_ΨSq, map_Ψ', map_Ψ₂Sq, ← coe_mapRingHom] + map_simp + +lemma map_ψ (n : ℤ) : (W.map f).ψ n = (W.ψ n).map (mapRingHom f) := by + simp only [WeierstrassCurve.ψ, map_ψ₂, map_Ψ₃, map_Ψ₄', ← coe_mapRingHom, map_normEDS] + map_simp + +lemma map_φ (n : ℤ) : (W.map f).φ n = (W.φ n).map (mapRingHom f) := by + simp only [WeierstrassCurve.φ, map_ψ] + map_simp + end WeierstrassCurve diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index d2bb171e8ff630..1bf2b294282357 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -65,7 +65,7 @@ elliptic, divisibility, sequence universe u v w -variable {R : Type u} [CommRing R] (W : ℤ → R) +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] (W : ℤ → R) (f : R →+* S) /-- The proposition that a sequence indexed by integers is an elliptic sequence. -/ def IsEllSequence : Prop := @@ -313,3 +313,20 @@ noncomputable def normEDSRec {P : ℕ → Sort u} normEDSRec' zero one two three four (fun _ ih => by apply even <;> exact ih _ <| by linarith only) (fun _ ih => by apply odd <;> exact ih _ <| by linarith only) n + +lemma map_preNormEDS' (n : ℕ) : f (preNormEDS' b c d n) = preNormEDS' (f b) (f c) (f d) n := by + induction n using normEDSRec' with + | zero => exact map_zero f + | one => exact map_one f + | two => exact map_one f + | three => rfl + | four => rfl + | _ _ ih => + simp only [preNormEDS'_odd, preNormEDS'_even, map_one, map_sub, map_mul, map_pow, apply_ite f] + repeat rw [ih _ <| by linarith only] + +lemma map_preNormEDS (n : ℤ) : f (preNormEDS b c d n) = preNormEDS (f b) (f c) (f d) n := by + rw [preNormEDS, map_mul, map_intCast, map_preNormEDS', preNormEDS] + +lemma map_normEDS (n : ℤ) : f (normEDS b c d n) = normEDS (f b) (f c) (f d) n := by + rw [normEDS, map_mul, map_preNormEDS, map_pow, apply_ite f, map_one, normEDS] From 0eeb05e817d3339f853feac0815358f146228c77 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 31 May 2024 21:54:24 +0100 Subject: [PATCH 31/47] Golf odd/even lemmas --- .../EllipticCurve/DivisionPolynomial.lean | 40 ++----------------- .../EllipticDivisibilitySequence.lean | 34 ++-------------- 2 files changed, 8 insertions(+), 66 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index 1956b8d0818925..d572ceaac68148 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -290,46 +290,14 @@ lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = (if Even m then W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 else -W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3) := by repeat erw [Ψ_ofNat] - rw [Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one] - by_cases hm : Even m - · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm - have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 - have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 - have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 - rw [if_pos hm, if_pos hm, if_pos hm4, if_pos hm2, if_neg hm1, if_neg hm3, if_pos hm] - C_simp - rw [C_Ψ₂Sq_eq] - ring1 - · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm - have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 - have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 - have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 - rw [if_neg hm, if_neg hm, if_neg hm4, if_neg hm2, if_pos hm1, if_pos hm3, if_neg hm] - C_simp - rw [C_Ψ₂Sq_eq] - ring1 + simp_rw [Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] + split_ifs <;> C_simp <;> rw [C_Ψ₂Sq_eq] <;> ring1 lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = W.Ψ (m + 2) ^ 2 * W.Ψ (m + 3) * W.Ψ (m + 5) - W.Ψ (m + 1) * W.Ψ (m + 3) * W.Ψ (m + 4) ^ 2 := by repeat erw [Ψ_ofNat] - erw [Ψ''_even, if_pos <| even_two_mul _] - by_cases hm : Even m - · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm - have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 - have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 - have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 - have hm5 : ¬Even (m + 5) := fun h => Nat.even_add_one.mp h hm4 - rw [if_pos hm2, if_neg hm3, if_neg hm5, if_neg hm1, if_pos hm4] - C_simp - ring1 - · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm - have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 - have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 - have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 - have hm5 : Even (m + 5) := Nat.even_add_one.mpr hm4 - rw [if_neg hm2, if_pos hm3, if_pos hm5, if_pos hm1, if_neg hm4] - C_simp - ring1 + simp_rw [Ψ''_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] + split_ifs <;> C_simp <;> ring1 @[simp] lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index d2bb171e8ff630..d86085a3d61ed0 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -244,41 +244,15 @@ lemma normEDS_odd (m : ℕ) : normEDS b c d (2 * (m + 2) + 1) = normEDS b c d (m + 4) * normEDS b c d (m + 2) ^ 3 - normEDS b c d (m + 1) * normEDS b c d (m + 3) ^ 3 := by repeat erw [normEDS_ofNat] - simp only [preNormEDS'_odd, if_neg (m + 2).not_even_two_mul_add_one] - by_cases hm : Even m - · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm - have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 - have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 - have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 - rw [if_pos hm, if_pos hm, if_pos hm4, if_pos hm2, if_neg hm1, if_neg hm3] - ring1 - · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm - have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 - have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 - have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 - rw [if_neg hm, if_neg hm, if_neg hm4, if_neg hm2, if_pos hm1, if_pos hm3] - ring1 + simp_rw [preNormEDS'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] + split_ifs <;> ring1 lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = normEDS b c d (m + 2) ^ 2 * normEDS b c d (m + 3) * normEDS b c d (m + 5) - normEDS b c d (m + 1) * normEDS b c d (m + 3) * normEDS b c d (m + 4) ^ 2 := by repeat erw [normEDS_ofNat] - simp only [preNormEDS'_even, if_pos <| even_two_mul _] - by_cases hm : Even m - · have hm1 : ¬Even (m + 1) := fun h => Nat.even_add_one.mp h hm - have hm2 : Even (m + 2) := Nat.even_add_one.mpr hm1 - have hm3 : ¬Even (m + 3) := fun h => Nat.even_add_one.mp h hm2 - have hm4 : Even (m + 4) := Nat.even_add_one.mpr hm3 - have hm5 : ¬Even (m + 5) := fun h => Nat.even_add_one.mp h hm4 - rw [if_pos hm2, if_neg hm3, if_neg hm5, if_neg hm1, if_pos hm4] - ring1 - · have hm1 : Even (m + 1) := Nat.even_add_one.mpr hm - have hm2 : ¬Even (m + 2) := fun h => Nat.even_add_one.mp h hm1 - have hm3 : Even (m + 3) := Nat.even_add_one.mpr hm2 - have hm4 : ¬Even (m + 4) := fun h => Nat.even_add_one.mp h hm3 - have hm5 : Even (m + 5) := Nat.even_add_one.mpr hm4 - rw [if_neg hm2, if_pos hm3, if_pos hm5, if_pos hm1, if_neg hm4] - ring1 + simp only [preNormEDS'_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] + split_ifs <;> ring1 @[simp] lemma normEDS_neg (n : ℤ) : normEDS b c d (-n) = -normEDS b c d n := by From 18a53a3789b4ca227fb924094f767a0b7d8eea05 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 31 May 2024 23:27:56 +0100 Subject: [PATCH 32/47] Rename variable --- .../EllipticCurve/Affine.lean | 95 +++++++++---------- 1 file changed, 46 insertions(+), 49 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index e12f6e78597b16..98cc802ddfc85e 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -874,119 +874,119 @@ lemma map_slope {F : Type u} [Field F] (W : Affine F) {K : Type v} [Field K] (f variable {R : Type r} [CommRing R] (W : Affine R) {S : Type s} [CommRing S] [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] - {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (φ : A →ₐ[S] B) + {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (g : A →ₐ[S] B) lemma baseChange_polynomial : (W.baseChange B).toAffine.polynomial = - (W.baseChange A).toAffine.polynomial.map (mapRingHom φ) := by + (W.baseChange A).toAffine.polynomial.map (mapRingHom g) := by rw [← map_polynomial, map_baseChange] lemma baseChange_eval_polynomial (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomial.eval <| x.map φ).eval (φ y) = - φ (((W.baseChange A).toAffine.polynomial.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomial.eval <| x.map g).eval (g y) = + g (((W.baseChange A).toAffine.polynomial.eval x).eval y) := by erw [← map_eval_polynomial, map_baseChange] rfl -variable {φ} in -lemma baseChange_equation (hφ : Function.Injective φ) (x y : A) : - (W.baseChange B).toAffine.Equation (φ x) (φ y) ↔ (W.baseChange A).toAffine.Equation x y := by - erw [← map_equation _ hφ, map_baseChange] +variable {g} in +lemma baseChange_equation (hg : Function.Injective g) (x y : A) : + (W.baseChange B).toAffine.Equation (g x) (g y) ↔ (W.baseChange A).toAffine.Equation x y := by + erw [← map_equation _ hg, map_baseChange] rfl #align weierstrass_curve.equation_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_equation lemma baseChange_polynomialX : (W.baseChange B).toAffine.polynomialX = - (W.baseChange A).toAffine.polynomialX.map (mapRingHom φ) := by + (W.baseChange A).toAffine.polynomialX.map (mapRingHom g) := by rw [← map_polynomialX, map_baseChange] lemma baseChange_eval_polynomialX (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomialX.eval <| x.map φ).eval (φ y) = - φ (((W.baseChange A).toAffine.polynomialX.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomialX.eval <| x.map g).eval (g y) = + g (((W.baseChange A).toAffine.polynomialX.eval x).eval y) := by erw [← map_eval_polynomialX, map_baseChange] rfl lemma baseChange_polynomialY : (W.baseChange B).toAffine.polynomialY = - (W.baseChange A).toAffine.polynomialY.map (mapRingHom φ) := by + (W.baseChange A).toAffine.polynomialY.map (mapRingHom g) := by rw [← map_polynomialY, map_baseChange] lemma baseChange_eval_polynomialY (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomialY.eval <| x.map φ).eval (φ y) = - φ (((W.baseChange A).toAffine.polynomialY.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomialY.eval <| x.map g).eval (g y) = + g (((W.baseChange A).toAffine.polynomialY.eval x).eval y) := by erw [← map_eval_polynomialY, map_baseChange] rfl -variable {φ} in -lemma baseChange_nonsingular (hφ : Function.Injective φ) (x y : A) : - (W.baseChange B).toAffine.Nonsingular (φ x) (φ y) ↔ +variable {g} in +lemma baseChange_nonsingular (hg : Function.Injective g) (x y : A) : + (W.baseChange B).toAffine.Nonsingular (g x) (g y) ↔ (W.baseChange A).toAffine.Nonsingular x y := by - erw [← map_nonsingular _ hφ, map_baseChange] + erw [← map_nonsingular _ hg, map_baseChange] rfl #align weierstrass_curve.nonsingular_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_nonsingular lemma baseChange_negPolynomial : (W.baseChange B).toAffine.negPolynomial = - (W.baseChange A).toAffine.negPolynomial.map (mapRingHom φ) := by + (W.baseChange A).toAffine.negPolynomial.map (mapRingHom g) := by rw [← map_negPolynomial, map_baseChange] lemma baseChange_negY (x y : A) : - (W.baseChange B).toAffine.negY (φ x) (φ y) = φ ((W.baseChange A).toAffine.negY x y) := by + (W.baseChange B).toAffine.negY (g x) (g y) = g ((W.baseChange A).toAffine.negY x y) := by erw [← map_negY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_neg_Y_of_base_change WeierstrassCurve.Affine.baseChange_negY lemma baseChange_addPolynomial (x y L : A) : - (W.baseChange B).toAffine.addPolynomial (φ x) (φ y) (φ L) = - ((W.baseChange A).toAffine.addPolynomial x y L).map φ := by + (W.baseChange B).toAffine.addPolynomial (g x) (g y) (g L) = + ((W.baseChange A).toAffine.addPolynomial x y L).map g := by rw [← map_addPolynomial, map_baseChange] rfl lemma baseChange_addX (x₁ x₂ L : A) : - (W.baseChange B).toAffine.addX (φ x₁) (φ x₂) (φ L) = - φ ((W.baseChange A).toAffine.addX x₁ x₂ L) := by + (W.baseChange B).toAffine.addX (g x₁) (g x₂) (g L) = + g ((W.baseChange A).toAffine.addX x₁ x₂ L) := by erw [← map_addX, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X_of_base_change WeierstrassCurve.Affine.baseChange_addX lemma baseChange_negAddY (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.negAddY (φ x₁) (φ x₂) (φ y₁) (φ L) = - φ ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.negAddY (g x₁) (g x₂) (g y₁) (g L) = + g ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by erw [← map_negAddY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y'_of_base_change WeierstrassCurve.Affine.baseChange_negAddY lemma baseChange_addY (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.addY (φ x₁) (φ x₂) (φ y₁) (φ L) = - φ ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.addY (g x₁) (g x₂) (g y₁) (g L) = + g ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by erw [← map_addY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y_of_base_change WeierstrassCurve.Affine.baseChange_addY variable {F : Type u} [Field F] [Algebra R F] [Algebra S F] [IsScalarTower R S F] - {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (φ : F →ₐ[S] K) - {L : Type w} [Field L] [Algebra R L] [Algebra S L] [IsScalarTower R S L] (ψ : K →ₐ[S] L) + {K : Type v} [Field K] [Algebra R K] [Algebra S K] [IsScalarTower R S K] (f : F →ₐ[S] K) + {L : Type w} [Field L] [Algebra R L] [Algebra S L] [IsScalarTower R S L] (g : K →ₐ[S] L) lemma baseChange_slope (x₁ x₂ y₁ y₂ : F) : - (W.baseChange K).toAffine.slope (φ x₁) (φ x₂) (φ y₁) (φ y₂) = - φ ((W.baseChange F).toAffine.slope x₁ x₂ y₁ y₂) := by + (W.baseChange K).toAffine.slope (f x₁) (f x₂) (f y₁) (f y₂) = + f ((W.baseChange F).toAffine.slope x₁ x₂ y₁ y₂) := by erw [← map_slope, map_baseChange] rfl #align weierstrass_curve.base_change_slope_of_base_change WeierstrassCurve.Affine.baseChange_slope namespace Point -/-- The function from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `φ : F →ₐ[S] K`, +/-- The function from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `f : F →ₐ[S] K`, where `W` is defined over a subring of a ring `S`, and `F` and `K` are field extensions of `S`. -/ def mapFun : W⟮F⟯ → W⟮K⟯ | 0 => 0 - | some h => some <| (W.baseChange_nonsingular φ.injective ..).mpr h + | some h => some <| (W.baseChange_nonsingular f.injective ..).mpr h #align weierstrass_curve.point.of_base_change_fun WeierstrassCurve.Affine.Point.mapFun -/-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `φ : F →ₐ[S] K`, +/-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by an algebra homomorphism `f : F →ₐ[S] K`, where `W` is defined over a subring of a ring `S`, and `F` and `K` are field extensions of `S`. -/ def map : W⟮F⟯ →+ W⟮K⟯ where - toFun := mapFun W φ + toFun := mapFun W f map_zero' := rfl map_add' := by rintro (_ | @⟨x₁, y₁, _⟩) (_ | @⟨x₂, y₂, _⟩) @@ -997,45 +997,42 @@ def map : W⟮F⟯ →+ W⟮K⟯ where rw [add_of_Y_eq (congr_arg _ hx) <| by rw [hy, baseChange_negY]] · simp only [add_of_Y_ne hx hy, mapFun] rw [add_of_Y_ne (congr_arg _ hx) <| - by simpa only [baseChange_negY] using fun h => hy <| φ.injective h] + by simpa only [baseChange_negY] using fun h => hy <| f.injective h] simp only [some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] - · simp only [add_of_X_ne hx, mapFun, add_of_X_ne fun h => hx <| φ.injective h, + · simp only [add_of_X_ne hx, mapFun, add_of_X_ne fun h => hx <| f.injective h, some.injEq, ← baseChange_addX, ← baseChange_addY, ← baseChange_slope] #align weierstrass_curve.point.of_base_change WeierstrassCurve.Affine.Point.map -lemma map_zero : map W φ (0 : W⟮F⟯) = 0 := +lemma map_zero : map W f (0 : W⟮F⟯) = 0 := rfl lemma map_some {x y : F} (h : (W.baseChange F).toAffine.Nonsingular x y) : - map W φ (some h) = some ((W.baseChange_nonsingular φ.injective ..).mpr h) := + map W f (some h) = some ((W.baseChange_nonsingular f.injective ..).mpr h) := rfl lemma map_id (P : W⟮F⟯) : map W (Algebra.ofId F F) P = P := by cases P <;> rfl -lemma map_map (P : W⟮F⟯) : map W ψ (map W φ P) = map W (ψ.comp φ) P := by +lemma map_map (P : W⟮F⟯) : map W g (map W f P) = map W (g.comp f) P := by cases P <;> rfl -lemma map_injective : Function.Injective <| map W φ := by +lemma map_injective : Function.Injective <| map W f := by rintro (_ | _) (_ | _) h any_goals contradiction · rfl - · simpa only [some.injEq] using ⟨φ.injective (some.inj h).left, φ.injective (some.inj h).right⟩ + · simpa only [some.injEq] using ⟨f.injective (some.inj h).left, f.injective (some.inj h).right⟩ #align weierstrass_curve.point.of_base_change_injective WeierstrassCurve.Affine.Point.map_injective -variable (F K) - +variable (F K) in /-- The group homomorphism from `W⟮F⟯` to `W⟮K⟯` induced by the base change from `F` to `K`, where `W` is defined over a subring of a ring `S`, and `F` and `K` are field extensions of `S`. -/ abbrev baseChange [Algebra F K] [IsScalarTower R F K] : W⟮F⟯ →+ W⟮K⟯ := map W <| Algebra.ofId F K -variable {F K} - lemma map_baseChange [Algebra F K] [IsScalarTower R F K] [Algebra F L] [IsScalarTower R F L] - (ψ : K →ₐ[F] L) (P : W⟮F⟯) : map W ψ (baseChange W F K P) = baseChange W F L P := by + (g : K →ₐ[F] L) (P : W⟮F⟯) : map W g (baseChange W F K P) = baseChange W F L P := by have : Subsingleton (F →ₐ[F] L) := inferInstance - convert map_map W (Algebra.ofId F K) ψ P + convert map_map W (Algebra.ofId F K) g P end Point From c71e42ec2753de6f82687cfd40f0c54139b946ff Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 1 Jun 2024 11:42:18 +0100 Subject: [PATCH 33/47] Fix merge --- .../EllipticCurve/Jacobian.lean | 29 ++++++++++--------- 1 file changed, 15 insertions(+), 14 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index f9d3768438188b..aee82a86e4302d 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -534,7 +534,7 @@ private lemma toAffine_slope_of_eq {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : have hPy : P y - W.negY P ≠ 0 := sub_ne_zero_of_ne <| Y_ne_negY_of_Y_ne' hP hQ hx hy simp only [mul_comm <| P z ^ _, ne_eq, ← div_eq_div_iff (pow_ne_zero _ hPz) (pow_ne_zero _ hQz)] at hx hy - rw [Affine.slope_of_Yne hx <| negY_of_Z_ne_zero hQz ▸ hy, ← negY_of_Z_ne_zero hPz, dblZ] + rw [Affine.slope_of_Y_ne hx <| negY_of_Z_ne_zero hQz ▸ hy, ← negY_of_Z_ne_zero hPz, dblZ] field_simp [pow_ne_zero 2 hPz] ring1 @@ -597,22 +597,22 @@ lemma negDblY_of_Y_eq [NoZeroDivisors R] {P Q : Fin 3 → R} (hQz : Q z ≠ 0) rw [negDblY, dblX_of_Y_eq hQz hx hy hy', eval_polynomialX, Y_eq_negY_of_Y_eq hQz hx hy hy'] ring1 -private lemma toAffine_addY'_of_eq {P : Fin 3 → F} {n d : F} (hPz : P z ≠ 0) (hd : d ≠ 0) : - W.toAffine.addY' (P x / P z ^ 2) (P x / P z ^ 2) (P y / P z ^ 3) (n / (P z * d)) = +private lemma toAffine_negAddY_of_eq {P : Fin 3 → F} {n d : F} (hPz : P z ≠ 0) (hd : d ≠ 0) : + W.toAffine.negAddY (P x / P z ^ 2) (P x / P z ^ 2) (P y / P z ^ 3) (n / (P z * d)) = (n * (n ^ 2 + W.a₁ * n * P z * d - W.a₂ * P z ^ 2 * d ^ 2 - 2 * P x * d ^ 2 - P x * d ^ 2) + P y * d ^ 3) / (P z * d) ^ 3 := by - linear_combination (norm := (rw [Affine.addY', toAffine_addX_of_eq hPz hd]; ring1)) + linear_combination (norm := (rw [Affine.negAddY, toAffine_addX_of_eq hPz hd]; ring1)) n * P x / (P z ^ 3 * d) * div_self (pow_ne_zero 2 hd) - P y / P z ^ 3 * div_self (pow_ne_zero 3 hd) lemma negDblY_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 = Q x * P z ^ 2) (hy : P y * Q z ^ 3 ≠ W.negY Q * P z ^ 3) : W.negDblY P / W.dblZ P ^ 3 = - W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + W.toAffine.negAddY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by rw [negDblY, dblX, toAffine_slope_of_eq hP hQ hPz hQz hx hy, dblZ, ← (div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)).mpr hx, - toAffine_addY'_of_eq hPz <| sub_ne_zero_of_ne <| Y_ne_negY_of_Y_ne' hP hQ hx hy] + toAffine_negAddY_of_eq hPz <| sub_ne_zero_of_ne <| Y_ne_negY_of_Y_ne' hP hQ hx hy] variable (W') in /-- The $Y$-coordinate of the doubling of a point representative. -/ @@ -673,7 +673,8 @@ private lemma toAffine_slope_of_ne {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : (hx : P x * Q z ^ 2 ≠ Q x * P z ^ 2) : W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3) = (P y * Q z ^ 3 - Q y * P z ^ 3) / (P z * Q z * addZ P Q) := by - rw [Affine.slope_of_Xne <| by rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)], + rw [Affine.slope_of_X_ne <| by + rwa [ne_eq, div_eq_div_iff (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz)], div_sub_div _ _ (pow_ne_zero 2 hPz) (pow_ne_zero 2 hQz), mul_comm <| _ ^ 2, addZ] field_simp [mul_ne_zero (mul_ne_zero hPz hQz) <| sub_ne_zero_of_ne hx, mul_ne_zero (mul_ne_zero (pow_ne_zero 3 hPz) (pow_ne_zero 3 hQz)) <| sub_ne_zero_of_ne hx] @@ -797,22 +798,22 @@ lemma negAddY_of_X_eq {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q rw [div_pow, ← negAddY_of_X_eq' hP hQ hx, mul_div_cancel_right₀ _ <| pow_ne_zero 3 <| mul_ne_zero hPz hQz] -private lemma toAffine_addY'_of_ne {P Q : Fin 3 → F} {n d : F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) +private lemma toAffine_negAddY_of_ne {P Q : Fin 3 → F} {n d : F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) (hd : d ≠ 0) : - W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (n / (P z * Q z * d)) = + W.toAffine.negAddY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (n / (P z * Q z * d)) = (n * (n ^ 2 + W.a₁ * n * P z * Q z * d - W.a₂ * P z ^ 2 * Q z ^ 2 * d ^ 2 - P x * Q z ^ 2 * d ^ 2 - Q x * P z ^ 2 * d ^ 2 - P x * Q z ^ 2 * d ^ 2) + P y * Q z ^ 3 * d ^ 3) / (P z * Q z * d) ^ 3 := by - linear_combination (norm := (rw [Affine.addY', toAffine_addX_of_ne hPz hQz hd]; ring1)) + linear_combination (norm := (rw [Affine.negAddY, toAffine_addX_of_ne hPz hQz hd]; ring1)) n * P x / (P z ^ 3 * Q z * d) * div_self (pow_ne_zero 2 <| mul_ne_zero hQz hd) - P y / P z ^ 3 * div_self (pow_ne_zero 3 <| mul_ne_zero hQz hd) lemma negAddY_of_Z_ne_zero {P Q : Fin 3 → F} (hP : W.Equation P) (hQ : W.Equation Q) (hPz : P z ≠ 0) - (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ Q x * P z ^ 2) : - W.negAddY P Q / addZ P Q ^ 3 = W.toAffine.addY' (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) - (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by + (hQz : Q z ≠ 0) (hx : P x * Q z ^ 2 ≠ Q x * P z ^ 2) : W.negAddY P Q / addZ P Q ^ 3 = + W.toAffine.negAddY (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) + (W.toAffine.slope (P x / P z ^ 2) (Q x / Q z ^ 2) (P y / P z ^ 3) (Q y / Q z ^ 3)) := by rw [negAddY_eq hPz hQz, addX_eq' hP hQ, div_div, ← mul_pow _ _ 3, toAffine_slope_of_ne hPz hQz hx, - toAffine_addY'_of_ne hPz hQz <| addZ_ne_zero_of_X_ne hx] + toAffine_negAddY_of_ne hPz hQz <| addZ_ne_zero_of_X_ne hx] variable (W') in /-- The $Y$-coordinate of the addition of two distinct point representatives. -/ From 7717581daeade74e4b8893a880e6797d46024cb7 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 1 Jun 2024 14:29:01 +0100 Subject: [PATCH 34/47] Add sections --- .../EllipticCurve/Affine.lean | 68 ++++++++++--------- 1 file changed, 37 insertions(+), 31 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 98cc802ddfc85e..d8015947857376 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -771,9 +771,9 @@ end Point end Group -section BaseChange +section Map -/-! ### Maps and base changes -/ +/-! ### Maps across ring homomorphisms -/ variable {S : Type v} [CommRing S] (f : R →+* S) @@ -872,92 +872,98 @@ lemma map_slope {F : Type u} [Field F] (W : Affine F) {K : Type v} [Field K] (f map_simp #align weierstrass_curve.base_change_slope WeierstrassCurve.Affine.map_slope +end Map + +section BaseChange + +/-! ### Base changes across algebra homomorphisms -/ + variable {R : Type r} [CommRing R] (W : Affine R) {S : Type s} [CommRing S] [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] - {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (g : A →ₐ[S] B) + {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (f : A →ₐ[S] B) lemma baseChange_polynomial : (W.baseChange B).toAffine.polynomial = - (W.baseChange A).toAffine.polynomial.map (mapRingHom g) := by + (W.baseChange A).toAffine.polynomial.map (mapRingHom f) := by rw [← map_polynomial, map_baseChange] lemma baseChange_eval_polynomial (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomial.eval <| x.map g).eval (g y) = - g (((W.baseChange A).toAffine.polynomial.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomial.eval <| x.map f).eval (f y) = + f (((W.baseChange A).toAffine.polynomial.eval x).eval y) := by erw [← map_eval_polynomial, map_baseChange] rfl variable {g} in -lemma baseChange_equation (hg : Function.Injective g) (x y : A) : - (W.baseChange B).toAffine.Equation (g x) (g y) ↔ (W.baseChange A).toAffine.Equation x y := by - erw [← map_equation _ hg, map_baseChange] +lemma baseChange_equation (hf : Function.Injective f) (x y : A) : + (W.baseChange B).toAffine.Equation (f x) (f y) ↔ (W.baseChange A).toAffine.Equation x y := by + erw [← map_equation _ hf, map_baseChange] rfl #align weierstrass_curve.equation_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_equation lemma baseChange_polynomialX : (W.baseChange B).toAffine.polynomialX = - (W.baseChange A).toAffine.polynomialX.map (mapRingHom g) := by + (W.baseChange A).toAffine.polynomialX.map (mapRingHom f) := by rw [← map_polynomialX, map_baseChange] lemma baseChange_eval_polynomialX (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomialX.eval <| x.map g).eval (g y) = - g (((W.baseChange A).toAffine.polynomialX.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomialX.eval <| x.map f).eval (f y) = + f (((W.baseChange A).toAffine.polynomialX.eval x).eval y) := by erw [← map_eval_polynomialX, map_baseChange] rfl lemma baseChange_polynomialY : (W.baseChange B).toAffine.polynomialY = - (W.baseChange A).toAffine.polynomialY.map (mapRingHom g) := by + (W.baseChange A).toAffine.polynomialY.map (mapRingHom f) := by rw [← map_polynomialY, map_baseChange] lemma baseChange_eval_polynomialY (x : A[X]) (y : A) : - ((W.baseChange B).toAffine.polynomialY.eval <| x.map g).eval (g y) = - g (((W.baseChange A).toAffine.polynomialY.eval x).eval y) := by + ((W.baseChange B).toAffine.polynomialY.eval <| x.map f).eval (f y) = + f (((W.baseChange A).toAffine.polynomialY.eval x).eval y) := by erw [← map_eval_polynomialY, map_baseChange] rfl -variable {g} in -lemma baseChange_nonsingular (hg : Function.Injective g) (x y : A) : - (W.baseChange B).toAffine.Nonsingular (g x) (g y) ↔ +variable {f} in +lemma baseChange_nonsingular (hf : Function.Injective f) (x y : A) : + (W.baseChange B).toAffine.Nonsingular (f x) (f y) ↔ (W.baseChange A).toAffine.Nonsingular x y := by - erw [← map_nonsingular _ hg, map_baseChange] + erw [← map_nonsingular _ hf, map_baseChange] rfl #align weierstrass_curve.nonsingular_iff_base_change_of_base_change WeierstrassCurve.Affine.baseChange_nonsingular lemma baseChange_negPolynomial : (W.baseChange B).toAffine.negPolynomial = - (W.baseChange A).toAffine.negPolynomial.map (mapRingHom g) := by + (W.baseChange A).toAffine.negPolynomial.map (mapRingHom f) := by rw [← map_negPolynomial, map_baseChange] lemma baseChange_negY (x y : A) : - (W.baseChange B).toAffine.negY (g x) (g y) = g ((W.baseChange A).toAffine.negY x y) := by + (W.baseChange B).toAffine.negY (f x) (f y) = f ((W.baseChange A).toAffine.negY x y) := by erw [← map_negY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_neg_Y_of_base_change WeierstrassCurve.Affine.baseChange_negY lemma baseChange_addPolynomial (x y L : A) : - (W.baseChange B).toAffine.addPolynomial (g x) (g y) (g L) = - ((W.baseChange A).toAffine.addPolynomial x y L).map g := by + (W.baseChange B).toAffine.addPolynomial (f x) (f y) (f L) = + ((W.baseChange A).toAffine.addPolynomial x y L).map f := by rw [← map_addPolynomial, map_baseChange] rfl lemma baseChange_addX (x₁ x₂ L : A) : - (W.baseChange B).toAffine.addX (g x₁) (g x₂) (g L) = - g ((W.baseChange A).toAffine.addX x₁ x₂ L) := by + (W.baseChange B).toAffine.addX (f x₁) (f x₂) (f L) = + f ((W.baseChange A).toAffine.addX x₁ x₂ L) := by erw [← map_addX, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_X_of_base_change WeierstrassCurve.Affine.baseChange_addX lemma baseChange_negAddY (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.negAddY (g x₁) (g x₂) (g y₁) (g L) = - g ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.negAddY (f x₁) (f x₂) (f y₁) (f L) = + f ((W.baseChange A).toAffine.negAddY x₁ x₂ y₁ L) := by erw [← map_negAddY, map_baseChange] rfl set_option linter.uppercaseLean3 false in #align weierstrass_curve.base_change_add_Y'_of_base_change WeierstrassCurve.Affine.baseChange_negAddY lemma baseChange_addY (x₁ x₂ y₁ L : A) : - (W.baseChange B).toAffine.addY (g x₁) (g x₂) (g y₁) (g L) = - g ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by + (W.baseChange B).toAffine.addY (f x₁) (f x₂) (f y₁) (f L) = + f ((W.baseChange A).toAffine.addY x₁ x₂ y₁ L) := by erw [← map_addY, map_baseChange] rfl set_option linter.uppercaseLean3 false in @@ -1030,9 +1036,9 @@ abbrev baseChange [Algebra F K] [IsScalarTower R F K] : W⟮F⟯ →+ W⟮K⟯ : map W <| Algebra.ofId F K lemma map_baseChange [Algebra F K] [IsScalarTower R F K] [Algebra F L] [IsScalarTower R F L] - (g : K →ₐ[F] L) (P : W⟮F⟯) : map W g (baseChange W F K P) = baseChange W F L P := by + (f : K →ₐ[F] L) (P : W⟮F⟯) : map W f (baseChange W F K P) = baseChange W F L P := by have : Subsingleton (F →ₐ[F] L) := inferInstance - convert map_map W (Algebra.ofId F K) g P + convert map_map W (Algebra.ofId F K) f P end Point From d21d346b1cc196cbe9a1b9e4840e06c0f0bbdee7 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 1 Jun 2024 14:45:56 +0100 Subject: [PATCH 35/47] Add base changes --- .../EllipticCurve/DivisionPolynomial.lean | 82 ++++++++++++++++++- 1 file changed, 80 insertions(+), 2 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index 4ccc01d4d0c4ef..28d81c092f7740 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -99,11 +99,13 @@ local macro "map_simp" : tactic => Polynomial.map_sub, Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_div, coe_mapRingHom, apply_ite <| mapRingHom _, WeierstrassCurve.map]) -universe u v +universe r s u v namespace WeierstrassCurve -variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] (W : WeierstrassCurve R) (f : R →+* S) +variable {R : Type r} {S : Type s} [CommRing R] [CommRing S] (W : WeierstrassCurve R) + +section Ψ₂Sq /-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ @@ -124,6 +126,10 @@ lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by lemma Ψ₂Sq_eq : W.Ψ₂Sq = W.twoTorsionPolynomial.toPoly := rfl +end Ψ₂Sq + +section Ψ'' + /-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$ -/ /-- The $3$-division polynomial $\psi_3 = \Psi_3$. -/ @@ -171,6 +177,10 @@ lemma Ψ''_even (m : ℕ) : W.Ψ'' (2 * (m + 3)) = W.Ψ'' (m + 1) * W.Ψ'' (m + 3) * W.Ψ'' (m + 4) ^ 2 := preNormEDS'_even .. +end Ψ'' + +section Ψ' + /-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the @@ -216,6 +226,10 @@ lemma Ψ'_even (m : ℕ) : W.Ψ' (2 * (m + 3)) = lemma Ψ'_neg (n : ℤ) : W.Ψ' (-n) = -W.Ψ' n := preNormEDS_neg .. +end Ψ' + +section ΨSq + /-! ### The univariate polynomials $\Psi_n^{[2]}$ -/ /-- The univariate polynomials $\Psi_n^{[2]}$ congruent to $\psi_n^2$. -/ @@ -260,6 +274,10 @@ lemma ΨSq_even (m : ℕ) : W.ΨSq (2 * (m + 3)) = lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by rw [WeierstrassCurve.ΨSq, Ψ'_neg, neg_sq, Int.natAbs_neg, WeierstrassCurve.ΨSq] +end ΨSq + +section Ψ + /-! ### The bivariate polynomials $\Psi_n$ -/ /-- The bivariate polynomials $\Psi_n$ congruent to the $n$-division polynomials $\psi_n$. -/ @@ -341,6 +359,10 @@ lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by rw [WeierstrassCurve.Ψ, Ψ'_neg, C_neg, neg_mul (α := R[X][Y]), Int.natAbs_neg, WeierstrassCurve.Ψ] +end Ψ + +section Φ + /-! ### The univariate polynomials $\Phi_n$ -/ /-- The univariate polynomials $\Phi_n$ congruent to $\phi_n$. -/ @@ -386,6 +408,10 @@ lemma Φ_neg (n : ℤ) : W.Φ (-n) = W.Φ n := by rw [WeierstrassCurve.Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, Ψ'_neg, ← neg_add', Ψ'_neg, neg_mul_neg, mul_comm <| W.Ψ' _, Int.natAbs_neg, WeierstrassCurve.Φ] +end Φ + +section ψ + /-! ### The bivariate polynomials $\psi_n$ -/ /-- The bivariate $n$-division polynomials $\psi_n$. -/ @@ -424,6 +450,10 @@ lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n := normEDS_neg .. +end ψ + +section φ + /-! ### The bivariate polynomials $\phi_n$ -/ /-- The bivariate polynomials $\phi_n$. -/ @@ -458,6 +488,12 @@ lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by rw [WeierstrassCurve.φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, ← neg_add', ψ_neg, neg_mul_neg, mul_comm <| W.ψ _, WeierstrassCurve.φ] +end φ + +section Map + +variable (f : R →+* S) + lemma map_ψ₂ : (W.map f).ψ₂ = W.ψ₂.map (mapRingHom f) := by simp only [WeierstrassCurve.ψ₂, Affine.map_polynomialY] @@ -501,4 +537,46 @@ lemma map_φ (n : ℤ) : (W.map f).φ n = (W.φ n).map (mapRingHom f) := by simp only [WeierstrassCurve.φ, map_ψ] map_simp +end Map + +section BaseChange + +variable [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] + {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (f : A →ₐ[S] B) + +lemma baseChange_ψ₂ : (W.baseChange B).ψ₂ = (W.baseChange A).ψ₂.map (mapRingHom f) := by + rw [← map_ψ₂, map_baseChange] + +lemma baseChange_Ψ₂Sq : (W.baseChange B).Ψ₂Sq = (W.baseChange A).Ψ₂Sq.map f := by + rw [← map_Ψ₂Sq, map_baseChange] + +lemma baseChange_Ψ₃ : (W.baseChange B).Ψ₃ = (W.baseChange A).Ψ₃.map f := by + rw [← map_Ψ₃, map_baseChange] + +lemma baseChange_Ψ₄' : (W.baseChange B).Ψ₄' = (W.baseChange A).Ψ₄'.map f := by + rw [← map_Ψ₄', map_baseChange] + +lemma baseChange_Ψ'' (n : ℕ) : (W.baseChange B).Ψ'' n = ((W.baseChange A).Ψ'' n).map f := by + rw [← map_Ψ'', map_baseChange] + +lemma baseChange_Ψ' (n : ℤ) : (W.baseChange B).Ψ' n = ((W.baseChange A).Ψ' n).map f := by + rw [← map_Ψ', map_baseChange] + +lemma baseChange_ΨSq (n : ℤ) : (W.baseChange B).ΨSq n = ((W.baseChange A).ΨSq n).map f := by + rw [← map_ΨSq, map_baseChange] + +lemma baseChange_Ψ (n : ℤ) : (W.baseChange B).Ψ n = ((W.baseChange A).Ψ n).map (mapRingHom f) := by + rw [← map_Ψ, map_baseChange] + +lemma baseChange_Φ (n : ℤ) : (W.baseChange B).Φ n = ((W.baseChange A).Φ n).map f := by + rw [← map_Φ, map_baseChange] + +lemma baseChange_ψ (n : ℤ) : (W.baseChange B).ψ n = ((W.baseChange A).ψ n).map (mapRingHom f) := by + rw [← map_ψ, map_baseChange] + +lemma baseChange_φ (n : ℤ) : (W.baseChange B).φ n = ((W.baseChange A).φ n).map (mapRingHom f) := by + rw [← map_φ, map_baseChange] + +end BaseChange + end WeierstrassCurve From fa649869b30eea69f1734b953eaf92fe64a1aaee Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 8 Jun 2024 18:09:04 +0100 Subject: [PATCH 36/47] Fix merge --- .../EllipticCurve/DivisionPolynomial.lean | 157 +++++++++--------- 1 file changed, 83 insertions(+), 74 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean index d572ceaac68148..2868185a1b5b41 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ import Mathlib.AlgebraicGeometry.EllipticCurve.Affine -import Mathlib.Data.Int.Parity import Mathlib.NumberTheory.EllipticDivisibilitySequence /-! @@ -27,7 +26,16 @@ $\psi_n \in R[X, Y]$ of $W$ is the normalised elliptic divisibility sequence wit Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by * $\phi_n := X\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, and - * $\omega_n := \tfrac{1}{2\psi_n} \cdot (\psi_{2n} - \psi_n^2(a_1\phi_n + a_3\psi_n^2)$. + * $\omega_n := \tfrac{1}{2} \cdot (\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2))$. + +Note that $\omega_n$ is always well-defined as a polynomial. As a start, it can be shown by +induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ is always +well-defined as a polynomial, while division by 2 is well-defined when $R$ has characteristic +different from 2. In general, it can be shown that 2 always divides the polynomial +$\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero universal ring +$\mathcal{R}[X, Y] := \mathbb{Z}[a_i][X, Y] / \langle W(a_i, X, Y)\rangle$ associated to $W$, where +$W(a_i, X, Y)$ is the associated Weierstrass equation. Then $\omega_n$ can be equivalently defined +as its image under the associated universal morphism $\mathcal{R}[X, Y] \to R[X, Y]$. Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial $\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences @@ -61,7 +69,7 @@ $\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. By induction, their leadi ## Main definitions * `WeierstrassCurve.Ψ₂Sq`: the univariate polynomial $\Psi_2^{[2]}$. - * `WeierstrassCurve.Ψ'`: the univariate polynomials $\tilde{\Psi}_n$. + * `WeierstrassCurve.preΨ`: the univariate polynomials $\tilde{\Psi}_n$. * `WeierstrassCurve.ΨSq`: the univariate polynomials $\Psi_n^{[2]}$. * `WeierstrassCurve.Ψ`: the bivariate polynomials $\Psi_n$. * `WeierstrassCurve.Φ`: the univariate polynomials $\Phi_n$. @@ -79,6 +87,8 @@ allow the computation of their leading terms without ambiguity. Furthermore, eva polynomials at a rational point on $W$ recovers their original definition up to linear combinations of the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring. +TODO: implementation notes for the definition of $\omega_n$. + ## References [J Silverman, *The Arithmetic of Elliptic Curves*][silverman2009] @@ -132,221 +142,220 @@ protected noncomputable def Ψ₄' : R[X] := /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ -protected noncomputable def Ψ'' (n : ℕ) : R[X] := +protected noncomputable def preΨ' (n : ℕ) : R[X] := preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n @[simp] -lemma Ψ''_zero : W.Ψ'' 0 = 0 := +lemma preΨ'_zero : W.preΨ' 0 = 0 := preNormEDS'_zero .. @[simp] -lemma Ψ''_one : W.Ψ'' 1 = 1 := +lemma preΨ'_one : W.preΨ' 1 = 1 := preNormEDS'_one .. @[simp] -lemma Ψ''_two : W.Ψ'' 2 = 1 := +lemma preΨ'_two : W.preΨ' 2 = 1 := preNormEDS'_two .. @[simp] -lemma Ψ''_three : W.Ψ'' 3 = W.Ψ₃ := +lemma preΨ'_three : W.preΨ' 3 = W.Ψ₃ := preNormEDS'_three .. @[simp] -lemma Ψ''_four : W.Ψ'' 4 = W.Ψ₄' := +lemma preΨ'_four : W.preΨ' 4 = W.Ψ₄' := preNormEDS'_four .. -lemma Ψ''_odd (m : ℕ) : W.Ψ'' (2 * (m + 2) + 1) = - W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := +lemma preΨ'_odd (m : ℕ) : W.preΨ' (2 * (m + 2) + 1) = + W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := preNormEDS'_odd .. -lemma Ψ''_even (m : ℕ) : W.Ψ'' (2 * (m + 3)) = - W.Ψ'' (m + 2) ^ 2 * W.Ψ'' (m + 3) * W.Ψ'' (m + 5) - - W.Ψ'' (m + 1) * W.Ψ'' (m + 3) * W.Ψ'' (m + 4) ^ 2 := +lemma preΨ'_even (m : ℕ) : W.preΨ' (2 * (m + 3)) = + W.preΨ' (m + 2) ^ 2 * W.preΨ' (m + 3) * W.preΨ' (m + 5) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) * W.preΨ' (m + 4) ^ 2 := preNormEDS'_even .. /-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ -protected noncomputable def Ψ' (n : ℤ) : R[X] := +protected noncomputable def preΨ (n : ℤ) : R[X] := preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n @[simp] -lemma Ψ'_ofNat (n : ℕ) : W.Ψ' n = W.Ψ'' n := +lemma preΨ_ofNat (n : ℕ) : W.preΨ n = W.preΨ' n := preNormEDS_ofNat .. @[simp] -lemma Ψ'_zero : W.Ψ' 0 = 0 := +lemma preΨ_zero : W.preΨ 0 = 0 := preNormEDS_zero .. @[simp] -lemma Ψ'_one : W.Ψ' 1 = 1 := +lemma preΨ_one : W.preΨ 1 = 1 := preNormEDS_one .. @[simp] -lemma Ψ'_two : W.Ψ' 2 = 1 := +lemma preΨ_two : W.preΨ 2 = 1 := preNormEDS_two .. @[simp] -lemma Ψ'_three : W.Ψ' 3 = W.Ψ₃ := +lemma preΨ_three : W.preΨ 3 = W.Ψ₃ := preNormEDS_three .. @[simp] -lemma Ψ'_four : W.Ψ' 4 = W.Ψ₄' := +lemma preΨ_four : W.preΨ 4 = W.Ψ₄' := preNormEDS_four .. -lemma Ψ'_odd (m : ℕ) : W.Ψ' (2 * (m + 2) + 1) = - W.Ψ' (m + 4) * W.Ψ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.Ψ' (m + 1) * W.Ψ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := +lemma preΨ_odd (m : ℕ) : W.preΨ (2 * (m + 2) + 1) = + W.preΨ (m + 4) * W.preΨ (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ (m + 1) * W.preΨ (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2) := preNormEDS_odd .. -lemma Ψ'_even (m : ℕ) : W.Ψ' (2 * (m + 3)) = - W.Ψ' (m + 2) ^ 2 * W.Ψ' (m + 3) * W.Ψ' (m + 5) - - W.Ψ' (m + 1) * W.Ψ' (m + 3) * W.Ψ' (m + 4) ^ 2 := +lemma preΨ_even (m : ℕ) : W.preΨ (2 * (m + 3)) = + W.preΨ (m + 2) ^ 2 * W.preΨ (m + 3) * W.preΨ (m + 5) - + W.preΨ (m + 1) * W.preΨ (m + 3) * W.preΨ (m + 4) ^ 2 := preNormEDS_even .. @[simp] -lemma Ψ'_neg (n : ℤ) : W.Ψ' (-n) = -W.Ψ' n := +lemma preΨ_neg (n : ℤ) : W.preΨ (-n) = -W.preΨ n := preNormEDS_neg .. /-! ### The univariate polynomials $\Psi_n^{[2]}$ -/ /-- The univariate polynomials $\Psi_n^{[2]}$ congruent to $\psi_n^2$. -/ protected noncomputable def ΨSq (n : ℤ) : R[X] := - W.Ψ' n ^ 2 * if Even n.natAbs then W.Ψ₂Sq else 1 + W.preΨ n ^ 2 * if Even n then W.Ψ₂Sq else 1 @[simp] -lemma ΨSq_ofNat (n : ℕ) : W.ΨSq n = W.Ψ'' n ^ 2 * if Even n then W.Ψ₂Sq else 1 := by - rw [WeierstrassCurve.ΨSq, Ψ'_ofNat, Int.natAbs_cast] +lemma ΨSq_ofNat (n : ℕ) : W.ΨSq n = W.preΨ' n ^ 2 * if Even n then W.Ψ₂Sq else 1 := by + simp only [WeierstrassCurve.ΨSq, preΨ_ofNat, Int.even_coe_nat] @[simp] lemma ΨSq_zero : W.ΨSq 0 = 0 := by - erw [ΨSq_ofNat, zero_pow two_ne_zero, zero_mul] + erw [ΨSq_ofNat, preΨ'_zero, zero_pow two_ne_zero, zero_mul] @[simp] lemma ΨSq_one : W.ΨSq 1 = 1 := by - erw [ΨSq_ofNat, one_pow, mul_one] + erw [ΨSq_ofNat, preΨ'_one, one_pow, mul_one] @[simp] lemma ΨSq_two : W.ΨSq 2 = W.Ψ₂Sq := by - erw [ΨSq_ofNat, one_pow, one_mul, if_pos even_two] + erw [ΨSq_ofNat, preΨ'_two, one_pow, one_mul, if_pos even_two] @[simp] lemma ΨSq_three : W.ΨSq 3 = W.Ψ₃ ^ 2 := by - erw [ΨSq_ofNat, Ψ''_three, mul_one] + erw [ΨSq_ofNat, preΨ'_three, mul_one] @[simp] lemma ΨSq_four : W.ΨSq 4 = W.Ψ₄' ^ 2 * W.Ψ₂Sq := by - erw [ΨSq_ofNat, Ψ''_four, if_pos <| by decide] + erw [ΨSq_ofNat, preΨ'_four, if_pos <| by decide] lemma ΨSq_odd (m : ℕ) : W.ΨSq (2 * (m + 2) + 1) = - (W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - - W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by - erw [ΨSq_ofNat, Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one, mul_one] + (W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 * (if Even m then W.Ψ₂Sq ^ 2 else 1) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3 * (if Even m then 1 else W.Ψ₂Sq ^ 2)) ^ 2 := by + erw [ΨSq_ofNat, preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, mul_one] lemma ΨSq_even (m : ℕ) : W.ΨSq (2 * (m + 3)) = - (W.Ψ'' (m + 2) ^ 2 * W.Ψ'' (m + 3) * W.Ψ'' (m + 5) - - W.Ψ'' (m + 1) * W.Ψ'' (m + 3) * W.Ψ'' (m + 4) ^ 2) ^ 2 * W.Ψ₂Sq := by - erw [ΨSq_ofNat, Ψ''_even, if_pos <| even_two_mul _] + (W.preΨ' (m + 2) ^ 2 * W.preΨ' (m + 3) * W.preΨ' (m + 5) - + W.preΨ' (m + 1) * W.preΨ' (m + 3) * W.preΨ' (m + 4) ^ 2) ^ 2 * W.Ψ₂Sq := by + erw [ΨSq_ofNat, preΨ'_even, if_pos <| even_two_mul _] @[simp] lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by - rw [WeierstrassCurve.ΨSq, Ψ'_neg, neg_sq, Int.natAbs_neg, WeierstrassCurve.ΨSq] + simp only [WeierstrassCurve.ΨSq, preΨ_neg, neg_sq, even_neg] /-! ### The bivariate polynomials $\Psi_n$ -/ /-- The bivariate polynomials $\Psi_n$ congruent to the $n$-division polynomials $\psi_n$. -/ protected noncomputable def Ψ (n : ℤ) : R[X][Y] := - C (W.Ψ' n) * if Even n.natAbs then W.ψ₂ else 1 + C (W.preΨ n) * if Even n then W.ψ₂ else 1 @[simp] -lemma Ψ_ofNat (n : ℕ) : W.Ψ n = C (W.Ψ'' n) * if Even n then W.ψ₂ else 1 := by - rw [WeierstrassCurve.Ψ, Ψ'_ofNat, Int.natAbs_cast] +lemma Ψ_ofNat (n : ℕ) : W.Ψ n = C (W.preΨ' n) * if Even n then W.ψ₂ else 1 := by + simp only [WeierstrassCurve.Ψ, preΨ_ofNat, Int.even_coe_nat] @[simp] lemma Ψ_zero : W.Ψ 0 = 0 := by - erw [Ψ_ofNat, C_0, zero_mul] + erw [Ψ_ofNat, preΨ'_zero, C_0, zero_mul] @[simp] lemma Ψ_one : W.Ψ 1 = 1 := by - erw [Ψ_ofNat, C_1, mul_one] + erw [Ψ_ofNat, preΨ'_one, C_1, mul_one] @[simp] lemma Ψ_two : W.Ψ 2 = W.ψ₂ := by - erw [Ψ_ofNat, one_mul, if_pos even_two] + erw [Ψ_ofNat, preΨ'_two, one_mul, if_pos even_two] @[simp] lemma Ψ_three : W.Ψ 3 = C W.Ψ₃ := by - erw [Ψ_ofNat, Ψ''_three, mul_one] + erw [Ψ_ofNat, preΨ'_three, mul_one] @[simp] lemma Ψ_four : W.Ψ 4 = C W.Ψ₄' * W.ψ₂ := by - erw [Ψ_ofNat, Ψ''_four, if_pos <| by decide] + erw [Ψ_ofNat, preΨ'_four, if_pos <| by decide] lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = W.Ψ (m + 4) * W.Ψ (m + 2) ^ 3 - W.Ψ (m + 1) * W.Ψ (m + 3) ^ 3 + W.toAffine.polynomial * (16 * W.toAffine.polynomial - 8 * W.ψ₂ ^ 2) * C - (if Even m then W.Ψ'' (m + 4) * W.Ψ'' (m + 2) ^ 3 - else -W.Ψ'' (m + 1) * W.Ψ'' (m + 3) ^ 3) := by + (if Even m then W.preΨ' (m + 4) * W.preΨ' (m + 2) ^ 3 + else -W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3) := by repeat erw [Ψ_ofNat] - simp_rw [Ψ''_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] + simp_rw [preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] split_ifs <;> C_simp <;> rw [C_Ψ₂Sq_eq] <;> ring1 lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = W.Ψ (m + 2) ^ 2 * W.Ψ (m + 3) * W.Ψ (m + 5) - W.Ψ (m + 1) * W.Ψ (m + 3) * W.Ψ (m + 4) ^ 2 := by repeat erw [Ψ_ofNat] - simp_rw [Ψ''_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] + simp_rw [preΨ'_even, if_pos <| even_two_mul _, Nat.even_add_one, ite_not] split_ifs <;> C_simp <;> ring1 @[simp] lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by - rw [WeierstrassCurve.Ψ, Ψ'_neg, C_neg, neg_mul (α := R[X][Y]), Int.natAbs_neg, WeierstrassCurve.Ψ] + simp only [WeierstrassCurve.Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg] /-! ### The univariate polynomials $\Phi_n$ -/ /-- The univariate polynomials $\Phi_n$ congruent to $\phi_n$. -/ protected noncomputable def Φ (n : ℤ) : R[X] := - X * W.ΨSq n - W.Ψ' (n + 1) * W.Ψ' (n - 1) * if Even n.natAbs then 1 else W.Ψ₂Sq + X * W.ΨSq n - W.preΨ (n + 1) * W.preΨ (n - 1) * if Even n then 1 else W.Ψ₂Sq @[simp] lemma Φ_ofNat (n : ℕ) : W.Φ (n + 1) = - X * W.Ψ'' (n + 1) ^ 2 * (if Even n then 1 else W.Ψ₂Sq) - - W.Ψ'' (n + 2) * W.Ψ'' n * (if Even n then W.Ψ₂Sq else 1) := by - erw [WeierstrassCurve.Φ, ΨSq_ofNat, ← mul_assoc, Ψ'_ofNat, add_sub_cancel_right, Ψ'_ofNat, - Int.natAbs_cast] - simp only [Nat.even_add_one, ite_not] + X * W.preΨ' (n + 1) ^ 2 * (if Even n then 1 else W.Ψ₂Sq) - + W.preΨ' (n + 2) * W.preΨ' n * (if Even n then W.Ψ₂Sq else 1) := by + erw [WeierstrassCurve.Φ, ΨSq_ofNat, ← mul_assoc, preΨ_ofNat, add_sub_cancel_right, preΨ_ofNat] + simp only [Nat.even_add_one, Int.even_add_one, Int.even_coe_nat, ite_not] @[simp] lemma Φ_zero : W.Φ 0 = 1 := by - rw [WeierstrassCurve.Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, Ψ'_one, one_mul, zero_sub, Ψ'_neg, - Ψ'_one, neg_one_mul, neg_neg, if_pos even_zero.natAbs] + rw [WeierstrassCurve.Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, preΨ_one, one_mul, zero_sub, + preΨ_neg, preΨ_one, neg_one_mul, neg_neg, if_pos even_zero] @[simp] lemma Φ_one : W.Φ 1 = X := by - erw [Φ_ofNat, Ψ''_one, one_pow, mul_one, mul_one, mul_zero, zero_mul, sub_zero] + erw [Φ_ofNat, preΨ'_one, one_pow, mul_one, mul_one, preΨ'_zero, mul_zero, zero_mul, sub_zero] @[simp] lemma Φ_two : W.Φ 2 = X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by - erw [Φ_ofNat, Ψ''_two, if_neg Nat.not_even_one, WeierstrassCurve.Ψ₂Sq, Ψ''_three, mul_one, - WeierstrassCurve.Ψ₃] + erw [Φ_ofNat, preΨ'_two, if_neg Nat.not_even_one, WeierstrassCurve.Ψ₂Sq, preΨ'_three, preΨ'_one, + mul_one, WeierstrassCurve.Ψ₃] C_simp ring1 @[simp] lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.Ψ₄' * W.Ψ₂Sq := by - erw [Φ_ofNat, Ψ''_three, mul_one, Ψ''_four, mul_one, if_pos even_two] + erw [Φ_ofNat, preΨ'_three, mul_one, preΨ'_four, preΨ'_two, mul_one, if_pos even_two] @[simp] lemma Φ_four : W.Φ 4 = X * W.Ψ₄' ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.Ψ₄' * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by - erw [Φ_ofNat, Ψ''_four, if_neg <| by decide, show 3 + 2 = 2 * 2 + 1 by rfl, Ψ''_odd, Ψ''_four, - Ψ''_two, if_pos even_zero, Ψ''_three, if_pos even_zero] + erw [Φ_ofNat, preΨ'_four, if_neg <| by decide, show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, + preΨ'_four, preΨ'_two, if_pos even_zero, preΨ'_one, preΨ'_three, if_pos even_zero] ring1 @[simp] lemma Φ_neg (n : ℤ) : W.Φ (-n) = W.Φ n := by - rw [WeierstrassCurve.Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, Ψ'_neg, ← neg_add', Ψ'_neg, - neg_mul_neg, mul_comm <| W.Ψ' _, Int.natAbs_neg, WeierstrassCurve.Φ] + simp only [WeierstrassCurve.Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, preΨ_neg, ← neg_add', + preΨ_neg, neg_mul_neg, mul_comm <| W.preΨ <| n - 1, even_neg] /-! ### The bivariate polynomials $\psi_n$ -/ @@ -418,6 +427,6 @@ lemma φ_four : W.φ 4 = C X * C W.Ψ₄' ^ 2 * W.ψ₂ ^ 2 - C W.Ψ₄' * W.ψ @[simp] lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by rw [WeierstrassCurve.φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, - ← neg_add', ψ_neg, neg_mul_neg, mul_comm <| W.ψ _, WeierstrassCurve.φ] + ← neg_add', ψ_neg, neg_mul_neg (α := R[X][Y]), mul_comm <| W.ψ _, WeierstrassCurve.φ] end WeierstrassCurve From 7190e5e9db9dfab6883a6c4537583aec21be8441 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 8 Jun 2024 18:14:50 +0100 Subject: [PATCH 37/47] Move DivisionPolynomial to DivisionPolynomial.Basic --- Mathlib.lean | 2 +- .../{DivisionPolynomial.lean => DivisionPolynomial/Basic.lean} | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename Mathlib/AlgebraicGeometry/EllipticCurve/{DivisionPolynomial.lean => DivisionPolynomial/Basic.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 0d022a998b2f42..51c4f10fbfe899 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -700,7 +700,7 @@ import Mathlib.Algebra.Tropical.Lattice import Mathlib.Algebra.Vertex.HVertexOperator import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.EllipticCurve.Affine -import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial +import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Basic import Mathlib.AlgebraicGeometry.EllipticCurve.Group import Mathlib.AlgebraicGeometry.EllipticCurve.Jacobian import Mathlib.AlgebraicGeometry.EllipticCurve.Projective diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean similarity index 100% rename from Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial.lean rename to Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean From d97a437af6cc52b314ac72b948c698ab292fb3da Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sat, 8 Jun 2024 23:58:34 +0100 Subject: [PATCH 38/47] Fix names --- .../DivisionPolynomial/Basic.lean | 51 +++++++++---------- 1 file changed, 23 insertions(+), 28 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 2868185a1b5b41..3d381d04a093f0 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -1,5 +1,5 @@ /- -Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. +Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ @@ -9,9 +9,9 @@ import Mathlib.NumberTheory.EllipticDivisibilitySequence /-! # Division polynomials of Weierstrass curves -This file defines certain polynomials associated to division polynomials of Weierstrass curves and -computes their leading terms. These are defined in terms of the auxiliary sequences of normalised -elliptic divisibility sequences defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`. +This file defines certain polynomials associated to division polynomials of Weierstrass curves. +These are defined in terms of the auxiliary sequences of normalised elliptic divisibility sequences +defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`. ## Mathematical background @@ -33,9 +33,8 @@ induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ well-defined as a polynomial, while division by 2 is well-defined when $R$ has characteristic different from 2. In general, it can be shown that 2 always divides the polynomial $\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero universal ring -$\mathcal{R}[X, Y] := \mathbb{Z}[a_i][X, Y] / \langle W(a_i, X, Y)\rangle$ associated to $W$, where -$W(a_i, X, Y)$ is the associated Weierstrass equation. Then $\omega_n$ can be equivalently defined -as its image under the associated universal morphism $\mathcal{R}[X, Y] \to R[X, Y]$. +$\mathbb{Z}[A_i][X, Y]$ of $W$, where $A_i$ are indeterminates. Then $\omega_n$ can be equivalently +defined as its image under the associated universal morphism $\mathbb{Z}[A_i][X, Y] \to R[X, Y]$. Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial $\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences @@ -60,15 +59,10 @@ Furthermore, define the associated sequences $\Psi_n^{[2]}, \Phi_n \in R[X]$ by With these definitions in mind, $\psi_n \in R[X, Y]$ and $\phi_n \in R[X, Y]$ are congruent in $R[W]$ to $\Psi_n \in R[X, Y]$ and $\Phi_n \in R[X]$ respectively, which are defined in terms of -$\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. By induction, their leading terms are - * $\tilde{\Psi}_n = nX^{\tfrac{n^2 - 4}{2}} + \dots$ if $n$ is even, - * $\tilde{\Psi}_n = nX^{\tfrac{n^2 - 1}{2}} + \dots$ if $n$ is odd, - * $\Psi_n^{[2]} = n^2X^{n^2 - 1} + \dots$, and - * $\Phi_n = X^{n^2} + \dots$. +$\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. ## Main definitions - * `WeierstrassCurve.Ψ₂Sq`: the univariate polynomial $\Psi_2^{[2]}$. * `WeierstrassCurve.preΨ`: the univariate polynomials $\tilde{\Psi}_n$. * `WeierstrassCurve.ΨSq`: the univariate polynomials $\Psi_n^{[2]}$. * `WeierstrassCurve.Ψ`: the bivariate polynomials $\Psi_n$. @@ -103,11 +97,11 @@ open Polynomial PolynomialPolynomial local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) -universe u v +universe u namespace WeierstrassCurve -variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] (W : WeierstrassCurve R) (f : R →+* S) +variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) /-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ @@ -136,14 +130,14 @@ protected noncomputable def Ψ₃ : R[X] := /-- The univariate polynomial $\tilde{\Psi}_4$, which is auxiliary to the $4$-division polynomial $\psi_4 = \Psi_4 = \tilde{\Psi}_4\psi_2$. -/ -protected noncomputable def Ψ₄' : R[X] := +protected noncomputable def preΨ₄ : R[X] := 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ protected noncomputable def preΨ' (n : ℕ) : R[X] := - preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n + preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @[simp] lemma preΨ'_zero : W.preΨ' 0 = 0 := @@ -162,7 +156,7 @@ lemma preΨ'_three : W.preΨ' 3 = W.Ψ₃ := preNormEDS'_three .. @[simp] -lemma preΨ'_four : W.preΨ' 4 = W.Ψ₄' := +lemma preΨ'_four : W.preΨ' 4 = W.preΨ₄ := preNormEDS'_four .. lemma preΨ'_odd (m : ℕ) : W.preΨ' (2 * (m + 2) + 1) = @@ -180,7 +174,7 @@ lemma preΨ'_even (m : ℕ) : W.preΨ' (2 * (m + 3)) = /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ protected noncomputable def preΨ (n : ℤ) : R[X] := - preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.Ψ₄' n + preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @[simp] lemma preΨ_ofNat (n : ℕ) : W.preΨ n = W.preΨ' n := @@ -203,7 +197,7 @@ lemma preΨ_three : W.preΨ 3 = W.Ψ₃ := preNormEDS_three .. @[simp] -lemma preΨ_four : W.preΨ 4 = W.Ψ₄' := +lemma preΨ_four : W.preΨ 4 = W.preΨ₄ := preNormEDS_four .. lemma preΨ_odd (m : ℕ) : W.preΨ (2 * (m + 2) + 1) = @@ -247,7 +241,7 @@ lemma ΨSq_three : W.ΨSq 3 = W.Ψ₃ ^ 2 := by erw [ΨSq_ofNat, preΨ'_three, mul_one] @[simp] -lemma ΨSq_four : W.ΨSq 4 = W.Ψ₄' ^ 2 * W.Ψ₂Sq := by +lemma ΨSq_four : W.ΨSq 4 = W.preΨ₄ ^ 2 * W.Ψ₂Sq := by erw [ΨSq_ofNat, preΨ'_four, if_pos <| by decide] lemma ΨSq_odd (m : ℕ) : W.ΨSq (2 * (m + 2) + 1) = @@ -291,7 +285,7 @@ lemma Ψ_three : W.Ψ 3 = C W.Ψ₃ := by erw [Ψ_ofNat, preΨ'_three, mul_one] @[simp] -lemma Ψ_four : W.Ψ 4 = C W.Ψ₄' * W.ψ₂ := by +lemma Ψ_four : W.Ψ 4 = C W.preΨ₄ * W.ψ₂ := by erw [Ψ_ofNat, preΨ'_four, if_pos <| by decide] lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = @@ -343,11 +337,11 @@ lemma Φ_two : W.Φ 2 = X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ ring1 @[simp] -lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.Ψ₄' * W.Ψ₂Sq := by +lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.preΨ₄ * W.Ψ₂Sq := by erw [Φ_ofNat, preΨ'_three, mul_one, preΨ'_four, preΨ'_two, mul_one, if_pos even_two] @[simp] -lemma Φ_four : W.Φ 4 = X * W.Ψ₄' ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.Ψ₄' * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by +lemma Φ_four : W.Φ 4 = X * W.preΨ₄ ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.preΨ₄ * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by erw [Φ_ofNat, preΨ'_four, if_neg <| by decide, show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, preΨ'_four, preΨ'_two, if_pos even_zero, preΨ'_one, preΨ'_three, if_pos even_zero] ring1 @@ -361,7 +355,7 @@ lemma Φ_neg (n : ℤ) : W.Φ (-n) = W.Φ n := by /-- The bivariate $n$-division polynomials $\psi_n$. -/ protected noncomputable def ψ (n : ℤ) : R[X][Y] := - normEDS W.ψ₂ (C W.Ψ₃) (C W.Ψ₄') n + normEDS W.ψ₂ (C W.Ψ₃) (C W.preΨ₄) n @[simp] lemma ψ_zero : W.ψ 0 = 0 := @@ -380,7 +374,7 @@ lemma ψ_three : W.ψ 3 = C W.Ψ₃ := normEDS_three .. @[simp] -lemma ψ_four : W.ψ 4 = C W.Ψ₄' * W.ψ₂ := +lemma ψ_four : W.ψ 4 = C W.preΨ₄ * W.ψ₂ := normEDS_four .. lemma ψ_odd (m : ℕ) : W.ψ (2 * (m + 2) + 1) = @@ -415,11 +409,12 @@ lemma φ_two : W.φ 2 = C X * W.ψ₂ ^ 2 - C W.Ψ₃ := by erw [WeierstrassCurve.φ, ψ_two, ψ_three, ψ_one, mul_one] @[simp] -lemma φ_three : W.φ 3 = C X * C W.Ψ₃ ^ 2 - C W.Ψ₄' * W.ψ₂ ^ 2 := by +lemma φ_three : W.φ 3 = C X * C W.Ψ₃ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 2 := by erw [WeierstrassCurve.φ, ψ_three, ψ_four, mul_assoc, ψ_two, ← sq] @[simp] -lemma φ_four : W.φ 4 = C X * C W.Ψ₄' ^ 2 * W.ψ₂ ^ 2 - C W.Ψ₄' * W.ψ₂ ^ 4 * C W.Ψ₃ + C W.Ψ₃ ^ 4 := by +lemma φ_four : + W.φ 4 = C X * C W.preΨ₄ ^ 2 * W.ψ₂ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 4 * C W.Ψ₃ + C W.Ψ₃ ^ 4 := by erw [WeierstrassCurve.φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, ψ_four, ψ_two, ψ_one, ψ_three] ring1 From 572d3599e60d4b4cdacab61f81b90524650888bd Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 9 Jun 2024 00:22:29 +0100 Subject: [PATCH 39/47] Fix merge --- .../EllipticCurve/DivisionPolynomial/Basic.lean | 4 ++++ Mathlib/NumberTheory/EllipticDivisibilitySequence.lean | 10 +++++----- 2 files changed, 9 insertions(+), 5 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 5db5695fb2bca3..234264ddf05afd 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -464,6 +464,8 @@ end φ section Map +/-! ### Maps across ring homomorphisms -/ + variable (f : R →+* S) lemma map_ψ₂ : (W.map f).ψ₂ = W.ψ₂.map (mapRingHom f) := by @@ -513,6 +515,8 @@ end Map section BaseChange +/-! ### Base changes across algebra homomorphisms -/ + variable [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (f : A →ₐ[S] B) diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index 7a9c74845bd3c2..0b23f4b5990bb8 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -290,11 +290,11 @@ noncomputable def normEDSRec {P : ℕ → Sort u} lemma map_preNormEDS' (n : ℕ) : f (preNormEDS' b c d n) = preNormEDS' (f b) (f c) (f d) n := by induction n using normEDSRec' with - | zero => exact map_zero f - | one => exact map_one f - | two => exact map_one f - | three => rfl - | four => rfl + | zero => rw [preNormEDS'_zero, map_zero, preNormEDS'_zero] + | one => rw [preNormEDS'_one, map_one, preNormEDS'_one] + | two => rw [preNormEDS'_two, map_one, preNormEDS'_two] + | three => rw [preNormEDS'_three, preNormEDS'_three] + | four => rw [preNormEDS'_four, preNormEDS'_four] | _ _ ih => simp only [preNormEDS'_odd, preNormEDS'_even, map_one, map_sub, map_mul, map_pow, apply_ite f] repeat rw [ih _ <| by linarith only] From 83ca5ba28bd392f8eeb716f30a4676fa2661a563 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 9 Jun 2024 11:17:04 +0100 Subject: [PATCH 40/47] Remove protection --- .../DivisionPolynomial/Basic.lean | 90 +++++++++++++------ 1 file changed, 63 insertions(+), 27 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 3d381d04a093f0..2b6ff77d21496c 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -103,18 +103,20 @@ namespace WeierstrassCurve variable {R : Type u} [CommRing R] (W : WeierstrassCurve R) +section Ψ₂Sq + /-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ /-- The $2$-division polynomial $\psi_2 = \Psi_2$. -/ -protected noncomputable def ψ₂ : R[X][Y] := +noncomputable def ψ₂ : R[X][Y] := W.toAffine.polynomialY /-- The univariate polynomial $\Psi_2^{[2]}$ congruent to $\psi_2^2$. -/ -protected noncomputable def Ψ₂Sq : R[X] := +noncomputable def Ψ₂Sq : R[X] := C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by - rw [WeierstrassCurve.Ψ₂Sq, WeierstrassCurve.ψ₂, b₂, b₄, b₆, Affine.polynomialY, Affine.polynomial] + rw [Ψ₂Sq, ψ₂, b₂, b₄, b₆, Affine.polynomialY, Affine.polynomial] C_simp ring1 @@ -122,21 +124,25 @@ lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by lemma Ψ₂Sq_eq : W.Ψ₂Sq = W.twoTorsionPolynomial.toPoly := rfl +end Ψ₂Sq + +section preΨ' + /-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$ -/ /-- The $3$-division polynomial $\psi_3 = \Psi_3$. -/ -protected noncomputable def Ψ₃ : R[X] := +noncomputable def Ψ₃ : R[X] := 3 * X ^ 4 + C W.b₂ * X ^ 3 + 3 * C W.b₄ * X ^ 2 + 3 * C W.b₆ * X + C W.b₈ /-- The univariate polynomial $\tilde{\Psi}_4$, which is auxiliary to the $4$-division polynomial $\psi_4 = \Psi_4 = \tilde{\Psi}_4\psi_2$. -/ -protected noncomputable def preΨ₄ : R[X] := +noncomputable def preΨ₄ : R[X] := 2 * X ^ 6 + C W.b₂ * X ^ 5 + 5 * C W.b₄ * X ^ 4 + 10 * C W.b₆ * X ^ 3 + 10 * C W.b₈ * X ^ 2 + C (W.b₂ * W.b₈ - W.b₄ * W.b₆) * X + C (W.b₄ * W.b₈ - W.b₆ ^ 2) /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{N}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ -protected noncomputable def preΨ' (n : ℕ) : R[X] := +noncomputable def preΨ' (n : ℕ) : R[X] := preNormEDS' (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @[simp] @@ -169,11 +175,15 @@ lemma preΨ'_even (m : ℕ) : W.preΨ' (2 * (m + 3)) = W.preΨ' (m + 1) * W.preΨ' (m + 3) * W.preΨ' (m + 4) ^ 2 := preNormEDS'_even .. +end preΨ' + +section preΨ + /-! ### The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$ -/ /-- The univariate polynomials $\tilde{\Psi}_n$ for $n \in \mathbb{Z}$, which are auxiliary to the bivariate polynomials $\Psi_n$ congruent to the bivariate $n$-division polynomials $\psi_n$. -/ -protected noncomputable def preΨ (n : ℤ) : R[X] := +noncomputable def preΨ (n : ℤ) : R[X] := preNormEDS (W.Ψ₂Sq ^ 2) W.Ψ₃ W.preΨ₄ n @[simp] @@ -214,15 +224,19 @@ lemma preΨ_even (m : ℕ) : W.preΨ (2 * (m + 3)) = lemma preΨ_neg (n : ℤ) : W.preΨ (-n) = -W.preΨ n := preNormEDS_neg .. +end preΨ + +section ΨSq + /-! ### The univariate polynomials $\Psi_n^{[2]}$ -/ /-- The univariate polynomials $\Psi_n^{[2]}$ congruent to $\psi_n^2$. -/ -protected noncomputable def ΨSq (n : ℤ) : R[X] := +noncomputable def ΨSq (n : ℤ) : R[X] := W.preΨ n ^ 2 * if Even n then W.Ψ₂Sq else 1 @[simp] lemma ΨSq_ofNat (n : ℕ) : W.ΨSq n = W.preΨ' n ^ 2 * if Even n then W.Ψ₂Sq else 1 := by - simp only [WeierstrassCurve.ΨSq, preΨ_ofNat, Int.even_coe_nat] + simp only [ΨSq, preΨ_ofNat, Int.even_coe_nat] @[simp] lemma ΨSq_zero : W.ΨSq 0 = 0 := by @@ -256,7 +270,11 @@ lemma ΨSq_even (m : ℕ) : W.ΨSq (2 * (m + 3)) = @[simp] lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by - simp only [WeierstrassCurve.ΨSq, preΨ_neg, neg_sq, even_neg] + simp only [ΨSq, preΨ_neg, neg_sq, even_neg] + +end ΨSq + +section Ψ /-! ### The bivariate polynomials $\Psi_n$ -/ @@ -264,9 +282,11 @@ lemma ΨSq_neg (n : ℤ) : W.ΨSq (-n) = W.ΨSq n := by protected noncomputable def Ψ (n : ℤ) : R[X][Y] := C (W.preΨ n) * if Even n then W.ψ₂ else 1 +open WeierstrassCurve (Ψ) + @[simp] lemma Ψ_ofNat (n : ℕ) : W.Ψ n = C (W.preΨ' n) * if Even n then W.ψ₂ else 1 := by - simp only [WeierstrassCurve.Ψ, preΨ_ofNat, Int.even_coe_nat] + simp only [Ψ, preΨ_ofNat, Int.even_coe_nat] @[simp] lemma Ψ_zero : W.Ψ 0 = 0 := by @@ -305,7 +325,11 @@ lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = @[simp] lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by - simp only [WeierstrassCurve.Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg] + simp only [Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg] + +end Ψ + +section Φ /-! ### The univariate polynomials $\Phi_n$ -/ @@ -313,17 +337,19 @@ lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by protected noncomputable def Φ (n : ℤ) : R[X] := X * W.ΨSq n - W.preΨ (n + 1) * W.preΨ (n - 1) * if Even n then 1 else W.Ψ₂Sq +open WeierstrassCurve (Φ) + @[simp] lemma Φ_ofNat (n : ℕ) : W.Φ (n + 1) = X * W.preΨ' (n + 1) ^ 2 * (if Even n then 1 else W.Ψ₂Sq) - W.preΨ' (n + 2) * W.preΨ' n * (if Even n then W.Ψ₂Sq else 1) := by - erw [WeierstrassCurve.Φ, ΨSq_ofNat, ← mul_assoc, preΨ_ofNat, add_sub_cancel_right, preΨ_ofNat] + erw [Φ, ΨSq_ofNat, ← mul_assoc, preΨ_ofNat, add_sub_cancel_right, preΨ_ofNat] simp only [Nat.even_add_one, Int.even_add_one, Int.even_coe_nat, ite_not] @[simp] lemma Φ_zero : W.Φ 0 = 1 := by - rw [WeierstrassCurve.Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, preΨ_one, one_mul, zero_sub, - preΨ_neg, preΨ_one, neg_one_mul, neg_neg, if_pos even_zero] + rw [Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, preΨ_one, one_mul, zero_sub, preΨ_neg, preΨ_one, + neg_one_mul, neg_neg, if_pos even_zero] @[simp] lemma Φ_one : W.Φ 1 = X := by @@ -331,8 +357,7 @@ lemma Φ_one : W.Φ 1 = X := by @[simp] lemma Φ_two : W.Φ 2 = X ^ 4 - C W.b₄ * X ^ 2 - C (2 * W.b₆) * X - C W.b₈ := by - erw [Φ_ofNat, preΨ'_two, if_neg Nat.not_even_one, WeierstrassCurve.Ψ₂Sq, preΨ'_three, preΨ'_one, - mul_one, WeierstrassCurve.Ψ₃] + erw [Φ_ofNat, preΨ'_two, if_neg Nat.not_even_one, Ψ₂Sq, preΨ'_three, preΨ'_one, mul_one, Ψ₃] C_simp ring1 @@ -348,8 +373,12 @@ lemma Φ_four : W.Φ 4 = X * W.preΨ₄ ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.preΨ₄ @[simp] lemma Φ_neg (n : ℤ) : W.Φ (-n) = W.Φ n := by - simp only [WeierstrassCurve.Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, preΨ_neg, ← neg_add', - preΨ_neg, neg_mul_neg, mul_comm <| W.preΨ <| n - 1, even_neg] + simp only [Φ, ΨSq_neg, neg_add_eq_sub, ← neg_sub n, preΨ_neg, ← neg_add', preΨ_neg, neg_mul_neg, + mul_comm <| W.preΨ <| n - 1, even_neg] + +end Φ + +section ψ /-! ### The bivariate polynomials $\psi_n$ -/ @@ -389,39 +418,46 @@ lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n := normEDS_neg .. +end ψ + +section φ + /-! ### The bivariate polynomials $\phi_n$ -/ /-- The bivariate polynomials $\phi_n$. -/ protected noncomputable def φ (n : ℤ) : R[X][Y] := C X * W.ψ n ^ 2 - W.ψ (n + 1) * W.ψ (n - 1) +open WeierstrassCurve (φ) + @[simp] lemma φ_zero : W.φ 0 = 1 := by - erw [WeierstrassCurve.φ, ψ_zero, zero_pow two_ne_zero, mul_zero, zero_sub, ψ_one, one_mul, + erw [φ, ψ_zero, zero_pow two_ne_zero, mul_zero, zero_sub, ψ_one, one_mul, zero_sub, ψ_neg, neg_neg, ψ_one] @[simp] lemma φ_one : W.φ 1 = C X := by - erw [WeierstrassCurve.φ, ψ_one, one_pow, mul_one, ψ_zero, mul_zero, sub_zero] + erw [φ, ψ_one, one_pow, mul_one, ψ_zero, mul_zero, sub_zero] @[simp] lemma φ_two : W.φ 2 = C X * W.ψ₂ ^ 2 - C W.Ψ₃ := by - erw [WeierstrassCurve.φ, ψ_two, ψ_three, ψ_one, mul_one] + erw [φ, ψ_two, ψ_three, ψ_one, mul_one] @[simp] lemma φ_three : W.φ 3 = C X * C W.Ψ₃ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 2 := by - erw [WeierstrassCurve.φ, ψ_three, ψ_four, mul_assoc, ψ_two, ← sq] + erw [φ, ψ_three, ψ_four, mul_assoc, ψ_two, ← sq] @[simp] lemma φ_four : W.φ 4 = C X * C W.preΨ₄ ^ 2 * W.ψ₂ ^ 2 - C W.preΨ₄ * W.ψ₂ ^ 4 * C W.Ψ₃ + C W.Ψ₃ ^ 4 := by - erw [WeierstrassCurve.φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, ψ_four, ψ_two, ψ_one, - ψ_three] + erw [φ, ψ_four, show (4 + 1 : ℤ) = 2 * 2 + 1 by rfl, ψ_odd, ψ_four, ψ_two, ψ_one, ψ_three] ring1 @[simp] lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by - rw [WeierstrassCurve.φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, - ← neg_add', ψ_neg, neg_mul_neg (α := R[X][Y]), mul_comm <| W.ψ _, WeierstrassCurve.φ] + rw [φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, ← neg_add', ψ_neg, + neg_mul_neg (α := R[X][Y]), mul_comm <| W.ψ _, φ] + +end φ end WeierstrassCurve From 98bbd5f0218d0e64599b18faef84a7cb69874aa1 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 9 Jun 2024 11:26:19 +0100 Subject: [PATCH 41/47] Fix merge --- .../DivisionPolynomial/Basic.lean | 40 ++++++------------- 1 file changed, 13 insertions(+), 27 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 6aa2679bb0a982..b59255c7cd0113 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -111,8 +111,6 @@ variable {R : Type r} {S : Type s} [CommRing R] [CommRing S] (W : WeierstrassCur section Ψ₂Sq -section Ψ₂Sq - /-! ### The univariate polynomial $\Psi_2^{[2]}$ -/ /-- The $2$-division polynomial $\psi_2 = \Psi_2$. -/ @@ -284,10 +282,6 @@ end ΨSq section Ψ -end ΨSq - -section Ψ - /-! ### The bivariate polynomials $\Psi_n$ -/ /-- The bivariate polynomials $\Psi_n$ congruent to the $n$-division polynomials $\psi_n$. -/ @@ -343,10 +337,6 @@ end Ψ section Φ -end Ψ - -section Φ - /-! ### The univariate polynomials $\Phi_n$ -/ /-- The univariate polynomials $\Phi_n$ congruent to $\phi_n$. -/ @@ -396,10 +386,6 @@ end Φ section ψ -end Φ - -section ψ - /-! ### The bivariate polynomials $\psi_n$ -/ /-- The bivariate $n$-division polynomials $\psi_n$. -/ @@ -480,55 +466,55 @@ lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by end φ -end φ - section Map /-! ### Maps across ring homomorphisms -/ +open WeierstrassCurve (Ψ Φ ψ φ) + variable (f : R →+* S) lemma map_ψ₂ : (W.map f).ψ₂ = W.ψ₂.map (mapRingHom f) := by - simp only [WeierstrassCurve.ψ₂, Affine.map_polynomialY] + simp only [ψ₂, Affine.map_polynomialY] lemma map_Ψ₂Sq : (W.map f).Ψ₂Sq = W.Ψ₂Sq.map f := by - simp only [WeierstrassCurve.Ψ₂Sq, map_b₂, map_b₄, map_b₆] + simp only [Ψ₂Sq, map_b₂, map_b₄, map_b₆] map_simp lemma map_Ψ₃ : (W.map f).Ψ₃ = W.Ψ₃.map f := by - simp only [WeierstrassCurve.Ψ₃, map_b₂, map_b₄, map_b₆, map_b₈] + simp only [Ψ₃, map_b₂, map_b₄, map_b₆, map_b₈] map_simp lemma map_preΨ₄ : (W.map f).preΨ₄ = W.preΨ₄.map f := by - simp only [WeierstrassCurve.preΨ₄, map_b₂, map_b₄, map_b₆, map_b₈] + simp only [preΨ₄, map_b₂, map_b₄, map_b₆, map_b₈] map_simp lemma map_preΨ' (n : ℕ) : (W.map f).preΨ' n = (W.preΨ' n).map f := by - simp only [WeierstrassCurve.preΨ', map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS'] + simp only [preΨ', map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS'] map_simp lemma map_preΨ (n : ℤ) : (W.map f).preΨ n = (W.preΨ n).map f := by - simp only [WeierstrassCurve.preΨ, map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS] + simp only [preΨ, map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS] map_simp lemma map_ΨSq (n : ℤ) : (W.map f).ΨSq n = (W.ΨSq n).map f := by - simp only [WeierstrassCurve.ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom] + simp only [ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom] map_simp lemma map_Ψ (n : ℤ) : (W.map f).Ψ n = (W.Ψ n).map (mapRingHom f) := by - simp only [WeierstrassCurve.Ψ, map_preΨ, map_ψ₂, ← coe_mapRingHom] + simp only [Ψ, map_preΨ, map_ψ₂, ← coe_mapRingHom] map_simp lemma map_Φ (n : ℤ) : (W.map f).Φ n = (W.Φ n).map f := by - simp only [WeierstrassCurve.Φ, map_ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom] + simp only [Φ, map_ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom] map_simp lemma map_ψ (n : ℤ) : (W.map f).ψ n = (W.ψ n).map (mapRingHom f) := by - simp only [WeierstrassCurve.ψ, map_ψ₂, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_normEDS] + simp only [ψ, map_ψ₂, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_normEDS] map_simp lemma map_φ (n : ℤ) : (W.map f).φ n = (W.φ n).map (mapRingHom f) := by - simp only [WeierstrassCurve.φ, map_ψ] + simp only [φ, map_ψ] map_simp end Map From a1ef8e9bfe9aeb44a7c2ca83b1752107f3e3c88e Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Sun, 9 Jun 2024 18:20:03 +0100 Subject: [PATCH 42/47] Prove congruences in coordinate ring --- .../DivisionPolynomial/Basic.lean | 28 ++++++++++++++++--- .../EllipticDivisibilitySequence.lean | 6 ++-- 2 files changed, 27 insertions(+), 7 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index b59255c7cd0113..a4593401d2fb72 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.AlgebraicGeometry.EllipticCurve.Affine +import Mathlib.AlgebraicGeometry.EllipticCurve.Group import Mathlib.NumberTheory.EllipticDivisibilitySequence /-! @@ -121,11 +121,17 @@ noncomputable def ψ₂ : R[X][Y] := noncomputable def Ψ₂Sq : R[X] := C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆ -lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by +lemma C_Ψ₂Sq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by rw [Ψ₂Sq, ψ₂, b₂, b₄, b₆, Affine.polynomialY, Affine.polynomial] C_simp ring1 +lemma ψ₂_sq : W.ψ₂ ^ 2 = C W.Ψ₂Sq + 4 * W.toAffine.polynomial := by + rw [C_Ψ₂Sq, sub_add_cancel] + +lemma Affine.CoordinateRing.mk_ψ₂_sq : mk W W.ψ₂ ^ 2 = mk W (C W.Ψ₂Sq) := by + rw [C_Ψ₂Sq, map_sub, map_mul, AdjoinRoot.mk_self, mul_zero, sub_zero, map_pow] + -- TODO: remove `twoTorsionPolynomial` in favour of `Ψ₂Sq` lemma Ψ₂Sq_eq : W.Ψ₂Sq = W.twoTorsionPolynomial.toPoly := rfl @@ -321,7 +327,7 @@ lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) = else -W.preΨ' (m + 1) * W.preΨ' (m + 3) ^ 3) := by repeat erw [Ψ_ofNat] simp_rw [preΨ'_odd, if_neg (m + 2).not_even_two_mul_add_one, Nat.even_add_one, ite_not] - split_ifs <;> C_simp <;> rw [C_Ψ₂Sq_eq] <;> ring1 + split_ifs <;> C_simp <;> rw [C_Ψ₂Sq] <;> ring1 lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = W.Ψ (m + 2) ^ 2 * W.Ψ (m + 3) * W.Ψ (m + 5) - W.Ψ (m + 1) * W.Ψ (m + 3) * W.Ψ (m + 4) ^ 2 := by @@ -333,6 +339,10 @@ lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ = lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by simp only [Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg] +lemma Affine.CoordinateRing.mk_Ψ_sq (n : ℤ) : mk W (W.Ψ n) ^ 2 = mk W (C <| W.ΨSq n) := by + simp only [Ψ, ΨSq, map_one, map_mul, map_pow, one_pow, mul_pow, ite_pow, apply_ite C, + apply_ite <| mk W, mk_ψ₂_sq] + end Ψ section Φ @@ -392,6 +402,8 @@ section ψ protected noncomputable def ψ (n : ℤ) : R[X][Y] := normEDS W.ψ₂ (C W.Ψ₃) (C W.preΨ₄) n +open WeierstrassCurve (Ψ ψ) + @[simp] lemma ψ_zero : W.ψ 0 = 0 := normEDS_zero .. @@ -424,6 +436,9 @@ lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ = lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n := normEDS_neg .. +lemma Affine.CoordinateRing.mk_ψ (n : ℤ) : mk W (W.ψ n) = mk W (W.Ψ n) := by + simp only [ψ, normEDS, Ψ, preΨ, map_mul, map_pow, map_preNormEDS, ← mk_ψ₂_sq, ← pow_mul] + end ψ section φ @@ -434,7 +449,7 @@ section φ protected noncomputable def φ (n : ℤ) : R[X][Y] := C X * W.ψ n ^ 2 - W.ψ (n + 1) * W.ψ (n - 1) -open WeierstrassCurve (φ) +open WeierstrassCurve (Ψ Φ φ) @[simp] lemma φ_zero : W.φ 0 = 1 := by @@ -464,6 +479,11 @@ lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by rw [φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, ← neg_add', ψ_neg, neg_mul_neg (α := R[X][Y]), mul_comm <| W.ψ _, φ] +lemma Affine.CoordinateRing.mk_φ (n : ℤ) : mk W (W.φ n) = mk W (C <| W.Φ n) := by + simp_rw [φ, Φ, map_sub, map_mul, map_pow, mk_ψ, mk_Ψ_sq, Ψ, map_mul, + mul_mul_mul_comm _ <| mk W <| ite .., Int.even_add_one, Int.even_sub_one, ← sq, ite_not, + apply_ite C, apply_ite <| mk W, ite_pow, map_one, one_pow, mk_ψ₂_sq] + end φ section Map diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index 0b23f4b5990bb8..0d73f91ca3bf14 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -213,12 +213,12 @@ with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = This is defined in terms of `preNormEDS` whose even terms differ by a factor of `b`. -/ def normEDS (n : ℤ) : R := - preNormEDS (b ^ 4) c d n * if Even n.natAbs then b else 1 + preNormEDS (b ^ 4) c d n * if Even n then b else 1 @[simp] lemma normEDS_ofNat (n : ℕ) : normEDS b c d n = preNormEDS' (b ^ 4) c d n * if Even n then b else 1 := by - rw [normEDS, preNormEDS_ofNat, Int.natAbs_ofNat] + simp only [normEDS, preNormEDS_ofNat, Int.even_coe_nat] @[simp] lemma normEDS_zero : normEDS b c d 0 = 0 := by @@ -256,7 +256,7 @@ lemma normEDS_even (m : ℕ) : normEDS b c d (2 * (m + 3)) * b = @[simp] lemma normEDS_neg (n : ℤ) : normEDS b c d (-n) = -normEDS b c d n := by - rw [normEDS, preNormEDS_neg, Int.natAbs_neg, neg_mul, normEDS] + simp only [normEDS, preNormEDS_neg, neg_mul, even_neg] /-- Strong recursion principle for a normalised EDS: if we have * `P 0`, `P 1`, `P 2`, `P 3`, and `P 4`, From f75eba2a912435d03213064d5ab70dd439865ccb Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 10 Jun 2024 19:26:15 +0100 Subject: [PATCH 43/47] Improve docstrings --- .../DivisionPolynomial/Basic.lean | 18 ++++++++--------- .../EllipticDivisibilitySequence.lean | 20 +++++++++---------- 2 files changed, 19 insertions(+), 19 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index a4593401d2fb72..1292f4c3dea803 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -10,13 +10,13 @@ import Mathlib.NumberTheory.EllipticDivisibilitySequence # Division polynomials of Weierstrass curves This file defines certain polynomials associated to division polynomials of Weierstrass curves. -These are defined in terms of the auxiliary sequences of normalised elliptic divisibility sequences -defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`. +These are defined in terms of the auxiliary sequences for normalised elliptic divisibility sequences +(EDS) as defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`. ## Mathematical background Let $W$ be a Weierstrass curve over a commutative ring $R$. The sequence of $n$-division polynomials -$\psi_n \in R[X, Y]$ of $W$ is the normalised elliptic divisibility sequence with initial values +$\psi_n \in R[X, Y]$ of $W$ is the normalised EDS with initial values * $\psi_0 := 0$, * $\psi_1 := 1$, * $\psi_2 := 2Y + a_1X + a_3$, @@ -28,8 +28,8 @@ Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by * $\phi_n := X\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, and * $\omega_n := \tfrac{1}{2} \cdot (\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2))$. -Note that $\omega_n$ is always well-defined as a polynomial. As a start, it can be shown by -induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ is always +Note that $\omega_n$ is always well-defined as a polynomial in $R[X, Y]$. As a start, it can be +shown by induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ is always well-defined as a polynomial, while division by 2 is well-defined when $R$ has characteristic different from 2. In general, it can be shown that 2 always divides the polynomial $\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero universal ring @@ -38,16 +38,16 @@ defined as its image under the associated universal morphism $\mathbb{Z}[A_i][X, Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial $\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences -associated to a normalised elliptic divisibility sequence show that $\psi_n / \psi_2$ are congruent -to certain polynomials in $R[X]$. In particular, define $\tilde{\Psi}_n \in R[X]$ as the auxiliary -sequence for a normalised elliptic divisibility sequence with initial values +associated to a normalised EDS show that $\psi_n / \psi_2$ are congruent to certain polynomials in +$R[X]$. In particular, define $\tilde{\Psi}_n \in R[X]$ as the auxiliary sequence for a normalised +EDS with extraneous parameter $\Psi_2^{[2]}$ and initial values * $\tilde{\Psi}_0 := 0$, * $\tilde{\Psi}_1 := 1$, * $\tilde{\Psi}_2 := 1$, * $\tilde{\Psi}_3 := \psi_3$, and * $\tilde{\Psi}_4 := \psi_4 / \psi_2$. -The corresponding normalised elliptic divisibility sequence $\Psi_n \in R[X, Y]$ is then given by +The corresponding normalised EDS $\Psi_n \in R[X, Y]$ is then given by * $\Psi_n := \tilde{\Psi}_n\psi_2$ if $n$ is even, and * $\Psi_n := \tilde{\Psi}_n$ if $n$ is odd. diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index 0d73f91ca3bf14..6bacac27154038 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -42,10 +42,10 @@ Some examples of EDSs include ## Implementation notes The normalised EDS `normEDS b c d n` is defined in terms of the auxiliary sequence -`preNormEDS (b ^ 4) c d n`, which are equal when `n` is odd, and which differ by a factor of `b` -when `n` is even. This coincides with the definition in the references since both agree for -`normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are removed in -`normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. +`preNormEDS (b ^ 4) c d n`, which are equal when `n` is odd, and which differ by a factor of an +extraneous parameter `b` when `n` is even. This coincides with the definition in the references +since both agree for `normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are +removed in `normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. One reason is to avoid the necessity for ring division by `b` in the inductive definition of `normEDS b c d (2 * (m + 3))`. The idea is that, it can be shown that `normEDS b c d (2 * (m + 3))` @@ -102,8 +102,8 @@ lemma IsDivSequence.smul (h : IsDivSequence W) (x : R) : IsDivSequence (x • W) lemma IsEllDivSequence.smul (h : IsEllDivSequence W) (x : R) : IsEllDivSequence (x • W) := ⟨h.left.smul x, h.right.smul x⟩ -/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, -with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ +/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with extraneous parameter `b` and +initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ def preNormEDS' (b c d : R) : ℕ → R | 0 => 0 | 1 => 1 @@ -159,8 +159,8 @@ lemma preNormEDS'_even (m : ℕ) : preNormEDS' b c d (2 * (m + 3)) = simp only [Nat.mul_add_div two_pos] rfl -/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, -with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. +/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, with extraneous parameter `b` and +initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. This extends `preNormEDS'` by defining its values at negative integers. -/ def preNormEDS (n : ℤ) : R := @@ -208,8 +208,8 @@ lemma preNormEDS_even (m : ℕ) : preNormEDS b c d (2 * (m + 3)) = lemma preNormEDS_neg (n : ℤ) : preNormEDS b c d (-n) = -preNormEDS b c d n := by rw [preNormEDS, Int.sign_neg, Int.cast_neg, neg_mul, Int.natAbs_neg, preNormEDS] -/-- The canonical example of a normalised EDS `W : ℤ → R`, -with initial values `W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = d * b`. +/-- The canonical example of a normalised EDS `W : ℤ → R`, with initial values +`W(0) = 0`, `W(1) = 1`, `W(2) = b`, `W(3) = c`, and `W(4) = d * b`. This is defined in terms of `preNormEDS` whose even terms differ by a factor of `b`. -/ def normEDS (n : ℤ) : R := From 66679117531dd20d7707cb846c9fda3b9ee52ccc Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 10 Jun 2024 19:40:01 +0100 Subject: [PATCH 44/47] Add dots --- .../DivisionPolynomial/Basic.lean | 40 ++++++++++--------- 1 file changed, 21 insertions(+), 19 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 1292f4c3dea803..ae1418eba483cd 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -25,16 +25,17 @@ $\psi_n \in R[X, Y]$ of $W$ is the normalised EDS with initial values (2X^6 + b_2X^5 + 5b_4X^4 + 10b_6X^3 + 10b_8X^2 + (b_2b_8 - b_4b_6)X + (b_4b_8 - b_6^2))$. Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by - * $\phi_n := X\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, and - * $\omega_n := \tfrac{1}{2} \cdot (\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2))$. + * $\phi_n := X\psi_n^2 - \psi_{n + 1} \cdot \psi_{n - 1}$, and + * $\omega_n := (\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)) / 2$. Note that $\omega_n$ is always well-defined as a polynomial in $R[X, Y]$. As a start, it can be -shown by induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ is always -well-defined as a polynomial, while division by 2 is well-defined when $R$ has characteristic -different from 2. In general, it can be shown that 2 always divides the polynomial -$\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero universal ring -$\mathbb{Z}[A_i][X, Y]$ of $W$, where $A_i$ are indeterminates. Then $\omega_n$ can be equivalently -defined as its image under the associated universal morphism $\mathbb{Z}[A_i][X, Y] \to R[X, Y]$. +shown by induction that $\psi_n$ always divides $\psi_{2n}$ in $R[X, Y]$, so that +$\psi_{2n} / \psi_n$ is always well-defined as a polynomial, while division by 2 is well-defined +when $R$ has characteristic different from 2. In general, it can be shown that 2 always divides the +polynomial $\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero +universal ring $\mathcal{R}[X, Y] := \mathbb{Z}[A_1, A_2, A_3, A_4, A_6][X, Y]$ of $W$, where the +$A_i$ are indeterminates. Then $\omega_n$ can be equivalently defined as the image of this division +under the associated universal morphism $\mathcal{R}[X, Y] \to R[X, Y]$ mapping $A_i$ to $a_i$. Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the polynomial $\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences @@ -48,14 +49,15 @@ EDS with extraneous parameter $\Psi_2^{[2]}$ and initial values * $\tilde{\Psi}_4 := \psi_4 / \psi_2$. The corresponding normalised EDS $\Psi_n \in R[X, Y]$ is then given by - * $\Psi_n := \tilde{\Psi}_n\psi_2$ if $n$ is even, and + * $\Psi_n := \tilde{\Psi}_n \cdot \psi_2$ if $n$ is even, and * $\Psi_n := \tilde{\Psi}_n$ if $n$ is odd. Furthermore, define the associated sequences $\Psi_n^{[2]}, \Phi_n \in R[X]$ by - * $\Psi_n^{[2]} := \tilde{\Psi}_n^2\Psi_2^{[2]}$ if $n$ is even, + * $\Psi_n^{[2]} := \tilde{\Psi}_n^2 \cdot \Psi_2^{[2]}$ if $n$ is even, * $\Psi_n^{[2]} := \tilde{\Psi}_n^2$ if $n$ is odd, - * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1}\tilde{\Psi}_{n - 1}$ if $n$ is even, and - * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1}\tilde{\Psi}_{n - 1}\Psi_2^{[2]}$ if $n$ is odd. + * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_{n - 1}$ if $n$ is even, and + * $\Phi_n := X\Psi_n^{[2]} - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_{n - 1} \cdot \Psi_2^{[2]}$ + if $n$ is odd. With these definitions in mind, $\psi_n \in R[X, Y]$ and $\phi_n \in R[X, Y]$ are congruent in $R[W]$ to $\Psi_n \in R[X, Y]$ and $\Phi_n \in R[X]$ respectively, which are defined in terms of @@ -73,13 +75,13 @@ $\Psi_2^{[2]} \in R[X]$ and $\tilde{\Psi}_n \in R[X]$. ## Implementation notes -Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate -polynomials $\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is -done partially to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[2]}$ -and $\Phi_n$ as univariate polynomials without needing to work under the coordinate ring, and to -allow the computation of their leading terms without ambiguity. Furthermore, evaluating these -polynomials at a rational point on $W$ recovers their original definition up to linear combinations -of the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring. +Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the bivariate polynomials +$\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is done partially +to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[2]}$ and $\Phi_n$ as +univariate polynomials without needing to work under the coordinate ring, and to allow the +computation of their leading terms without ambiguity. Furthermore, evaluating these polynomials at a +rational point on $W$ recovers their original definition up to linear combinations of the +Weierstrass equation of $W$, hence also avoiding the need to work in the coordinate ring. TODO: implementation notes for the definition of $\omega_n$. From 854c757c60d1c3bf67f934694ed3594c35179390 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 10 Jun 2024 20:26:40 +0100 Subject: [PATCH 45/47] Fix English --- .../EllipticCurve/DivisionPolynomial/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index ae1418eba483cd..a9b0ebc8e1f1cc 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -41,7 +41,7 @@ Now, in the coordinate ring $R[W]$, note that $\psi_2^2$ is congruent to the pol $\Psi_2^{[2]} := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[X]$. As such, in $R[W]$, the recurrences associated to a normalised EDS show that $\psi_n / \psi_2$ are congruent to certain polynomials in $R[X]$. In particular, define $\tilde{\Psi}_n \in R[X]$ as the auxiliary sequence for a normalised -EDS with extraneous parameter $\Psi_2^{[2]}$ and initial values +EDS with extra parameter $(\Psi_2^{[2]})^2$ and initial values * $\tilde{\Psi}_0 := 0$, * $\tilde{\Psi}_1 := 1$, * $\tilde{\Psi}_2 := 1$, From fd29797719c1205bfa89000fe6a956e57a891603 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 10 Jun 2024 20:27:57 +0100 Subject: [PATCH 46/47] Fix English --- .../NumberTheory/EllipticDivisibilitySequence.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index 6bacac27154038..b36c8617d2c1c6 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -43,8 +43,8 @@ Some examples of EDSs include The normalised EDS `normEDS b c d n` is defined in terms of the auxiliary sequence `preNormEDS (b ^ 4) c d n`, which are equal when `n` is odd, and which differ by a factor of an -extraneous parameter `b` when `n` is even. This coincides with the definition in the references -since both agree for `normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are +extra parameter `b` when `n` is even. This coincides with the definition in the references since +both agree for `normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are removed in `normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. One reason is to avoid the necessity for ring division by `b` in the inductive definition of @@ -102,8 +102,8 @@ lemma IsDivSequence.smul (h : IsDivSequence W) (x : R) : IsDivSequence (x • W) lemma IsEllDivSequence.smul (h : IsEllDivSequence W) (x : R) : IsEllDivSequence (x • W) := ⟨h.left.smul x, h.right.smul x⟩ -/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with extraneous parameter `b` and -initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ +/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with extra parameter `b` and initial +values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ def preNormEDS' (b c d : R) : ℕ → R | 0 => 0 | 1 => 1 @@ -159,8 +159,8 @@ lemma preNormEDS'_even (m : ℕ) : preNormEDS' b c d (2 * (m + 3)) = simp only [Nat.mul_add_div two_pos] rfl -/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, with extraneous parameter `b` and -initial values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. +/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, with extra parameter `b` and initial +values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. This extends `preNormEDS'` by defining its values at negative integers. -/ def preNormEDS (n : ℤ) : R := From 541bb1de543679a9729423cacd37296e99e7d143 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 10 Jun 2024 21:15:51 +0100 Subject: [PATCH 47/47] Change order --- .../EllipticDivisibilitySequence.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean index b36c8617d2c1c6..5b9b7bfca33077 100644 --- a/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean +++ b/Mathlib/NumberTheory/EllipticDivisibilitySequence.lean @@ -42,10 +42,10 @@ Some examples of EDSs include ## Implementation notes The normalised EDS `normEDS b c d n` is defined in terms of the auxiliary sequence -`preNormEDS (b ^ 4) c d n`, which are equal when `n` is odd, and which differ by a factor of an -extra parameter `b` when `n` is even. This coincides with the definition in the references since -both agree for `normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are -removed in `normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. +`preNormEDS (b ^ 4) c d n`, which are equal when `n` is odd, and which differ by a factor of `b` +when `n` is even. This coincides with the definition in the references since both agree for +`normEDS b c d 2` and for `normEDS b c d 4`, and the correct factors of `b` are removed in +`normEDS b c d (2 * (m + 2) + 1)` and in `normEDS b c d (2 * (m + 3))`. One reason is to avoid the necessity for ring division by `b` in the inductive definition of `normEDS b c d (2 * (m + 3))`. The idea is that, it can be shown that `normEDS b c d (2 * (m + 3))` @@ -102,8 +102,8 @@ lemma IsDivSequence.smul (h : IsDivSequence W) (x : R) : IsDivSequence (x • W) lemma IsEllDivSequence.smul (h : IsEllDivSequence W) (x : R) : IsEllDivSequence (x • W) := ⟨h.left.smul x, h.right.smul x⟩ -/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with extra parameter `b` and initial -values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. -/ +/-- The auxiliary sequence for a normalised EDS `W : ℕ → R`, with initial values +`W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d` and extra parameter `b`. -/ def preNormEDS' (b c d : R) : ℕ → R | 0 => 0 | 1 => 1 @@ -159,8 +159,8 @@ lemma preNormEDS'_even (m : ℕ) : preNormEDS' b c d (2 * (m + 3)) = simp only [Nat.mul_add_div two_pos] rfl -/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, with extra parameter `b` and initial -values `W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d`. +/-- The auxiliary sequence for a normalised EDS `W : ℤ → R`, with initial values +`W(0) = 0`, `W(1) = 1`, `W(2) = 1`, `W(3) = c`, and `W(4) = d` and extra parameter `b`. This extends `preNormEDS'` by defining its values at negative integers. -/ def preNormEDS (n : ℤ) : R :=