@@ -3,20 +3,20 @@ Copyright (c) 2024 David Kurniadi Angdinata. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: David Kurniadi Angdinata
5
5
-/
6
- import Mathlib.AlgebraicGeometry.EllipticCurve.Affine
6
+ import Mathlib.AlgebraicGeometry.EllipticCurve.Group
7
7
import Mathlib.NumberTheory.EllipticDivisibilitySequence
8
8
9
9
/-!
10
10
# Division polynomials of Weierstrass curves
11
11
12
12
This file defines certain polynomials associated to division polynomials of Weierstrass curves.
13
- These are defined in terms of the auxiliary sequences of normalised elliptic divisibility sequences
14
- defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`.
13
+ These are defined in terms of the auxiliary sequences for normalised elliptic divisibility sequences
14
+ (EDS) as defined in `Mathlib.NumberTheory.EllipticDivisibilitySequence`.
15
15
16
16
## Mathematical background
17
17
18
18
Let $W$ be a Weierstrass curve over a commutative ring $R$. The sequence of $n$-division polynomials
19
- $\psi_n \in R[X, Y]$ of $W$ is the normalised elliptic divisibility sequence with initial values
19
+ $\psi_n \in R[X, Y]$ of $W$ is the normalised EDS with initial values
20
20
* $\psi_0 := 0$,
21
21
* $\psi_1 := 1$,
22
22
* $\psi_2 := 2Y + a_1X + a_3$,
@@ -25,37 +25,39 @@ $\psi_n \in R[X, Y]$ of $W$ is the normalised elliptic divisibility sequence wit
25
25
(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))$.
26
26
27
27
Furthermore, define the associated sequences $\phi_n, \omega_n \in R[X, Y]$ by
28
- * $\phi_n := X\psi_n^2 - \psi_{n + 1}\psi_{n - 1}$, and
29
- * $\omega_n := \tfrac{1}{2} \cdot (\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2))$.
30
-
31
- Note that $\omega_n$ is always well-defined as a polynomial. As a start, it can be shown by
32
- induction that $\psi_n$ always divides $\psi_{2n}$, so that $\psi_{2n} / \psi_n$ is always
33
- well-defined as a polynomial, while division by 2 is well-defined when $R$ has characteristic
34
- different from 2. In general, it can be shown that 2 always divides the polynomial
35
- $\psi_{2n} / \psi_n - \psi_n(a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero universal ring
36
- $\mathbb{Z}[ A_i ] [X, Y ]$ of $W$, where $A_i$ are indeterminates. Then $\omega_n$ can be equivalently
37
- defined as its image under the associated universal morphism $\mathbb{Z}[ A_i ] [X, Y ] \to R[X, Y]$.
28
+ * $\phi_n := X\psi_n^2 - \psi_{n + 1} \cdot \psi_{n - 1}$, and
29
+ * $\omega_n := (\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)) / 2$.
30
+
31
+ Note that $\omega_n$ is always well-defined as a polynomial in $R[X, Y]$. As a start, it can be
32
+ shown by induction that $\psi_n$ always divides $\psi_{2n}$ in $R[X, Y]$, so that
33
+ $\psi_{2n} / \psi_n$ is always well-defined as a polynomial, while division by 2 is well-defined
34
+ when $R$ has characteristic different from 2. In general, it can be shown that 2 always divides the
35
+ polynomial $\psi_{2n} / \psi_n - \psi_n \cdot (a_1\phi_n + a_3\psi_n^2)$ in the characteristic zero
36
+ universal ring $\mathcal{R}[X, Y] := \mathbb{Z}[ A_1, A_2, A_3, A_4, A_6 ] [X, Y ]$ of $W$, where the
37
+ $A_i$ are indeterminates. Then $\omega_n$ can be equivalently defined as the image of this division
38
+ under the associated universal morphism $\mathcal{R}[X, Y] \to R[X, Y]$ mapping $A_i$ to $a_i$.
38
39
39
40
Now, in the coordinate ring $R[ W ] $, note that $\psi_2^2$ is congruent to the polynomial
40
41
$\Psi_2^{[ 2 ] } := 4X^3 + b_2X^2 + 2b_4X + b_6 \in R[ X ] $. As such, in $R[ W ] $, the recurrences
41
- associated to a normalised elliptic divisibility sequence show that $\psi_n / \psi_2$ are congruent
42
- to certain polynomials in $R[ X ] $. In particular, define $\tilde{\Psi}_n \in R[ X ] $ as the auxiliary
43
- sequence for a normalised elliptic divisibility sequence with initial values
42
+ associated to a normalised EDS show that $\psi_n / \psi_2$ are congruent to certain polynomials in
43
+ $R[ X ] $. In particular, define $\tilde{\Psi}_n \in R[ X ] $ as the auxiliary sequence for a normalised
44
+ EDS with extra parameter $(\Psi_2^{ [ 2 ] })^2$ and initial values
44
45
* $\tilde{\Psi}_0 := 0$,
45
46
* $\tilde{\Psi}_1 := 1$,
46
47
* $\tilde{\Psi}_2 := 1$,
47
48
* $\tilde{\Psi}_3 := \psi_3$, and
48
49
* $\tilde{\Psi}_4 := \psi_4 / \psi_2$.
49
50
50
- The corresponding normalised elliptic divisibility sequence $\Psi_n \in R[X, Y]$ is then given by
51
- * $\Psi_n := \tilde{\Psi}_n\psi_2$ if $n$ is even, and
51
+ The corresponding normalised EDS $\Psi_n \in R[X, Y]$ is then given by
52
+ * $\Psi_n := \tilde{\Psi}_n \cdot \psi_2$ if $n$ is even, and
52
53
* $\Psi_n := \tilde{\Psi}_n$ if $n$ is odd.
53
54
54
55
Furthermore, define the associated sequences $\Psi_n^{[ 2 ] }, \Phi_n \in R[ X ] $ by
55
- * $\Psi_n^{[ 2 ] } := \tilde{\Psi}_n^2\Psi_2^{[ 2 ] }$ if $n$ is even,
56
+ * $\Psi_n^{[ 2 ] } := \tilde{\Psi}_n^2 \cdot \Psi_2^{[ 2 ] }$ if $n$ is even,
56
57
* $\Psi_n^{[ 2 ] } := \tilde{\Psi}_n^2$ if $n$ is odd,
57
- * $\Phi_n := X\Psi_n^{[ 2 ] } - \tilde{\Psi}_{n + 1}\tilde{\Psi}_ {n - 1}$ if $n$ is even, and
58
- * $\Phi_n := X\Psi_n^{[ 2 ] } - \tilde{\Psi}_{n + 1}\tilde{\Psi}_ {n - 1}\Psi_2^{[ 2 ] }$ if $n$ is odd.
58
+ * $\Phi_n := X\Psi_n^{[ 2 ] } - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_ {n - 1}$ if $n$ is even, and
59
+ * $\Phi_n := X\Psi_n^{[ 2 ] } - \tilde{\Psi}_{n + 1} \cdot \tilde{\Psi}_ {n - 1} \cdot \Psi_2^{[ 2 ] }$
60
+ if $n$ is odd.
59
61
60
62
With these definitions in mind, $\psi_n \in R[X, Y]$ and $\phi_n \in R[X, Y]$ are congruent in
61
63
$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]$.
73
75
74
76
## Implementation notes
75
77
76
- Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the $n$-division bivariate
77
- polynomials $\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is
78
- done partially to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[ 2 ] }$
79
- and $\Phi_n$ as univariate polynomials without needing to work under the coordinate ring, and to
80
- allow the computation of their leading terms without ambiguity. Furthermore, evaluating these
81
- polynomials at a rational point on $W$ recovers their original definition up to linear combinations
82
- of the Weierstrass equation of $W$, hence also avoiding the need to work under the coordinate ring.
78
+ Analogously to `Mathlib.NumberTheory.EllipticDivisibilitySequence`, the bivariate polynomials
79
+ $\Psi_n$ are defined in terms of the univariate polynomials $\tilde{\Psi}_n$. This is done partially
80
+ to avoid ring division, but more crucially to allow the definition of $\Psi_n^{[ 2 ] }$ and $\Phi_n$ as
81
+ univariate polynomials without needing to work under the coordinate ring, and to allow the
82
+ computation of their leading terms without ambiguity. Furthermore, evaluating these polynomials at a
83
+ rational point on $W$ recovers their original definition up to linear combinations of the
84
+ Weierstrass equation of $W$, hence also avoiding the need to work in the coordinate ring.
83
85
84
86
TODO: implementation notes for the definition of $\omega_n$.
85
87
@@ -97,11 +99,17 @@ open Polynomial PolynomialPolynomial
97
99
local macro "C_simp" : tactic =>
98
100
`(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow])
99
101
100
- universe u
102
+ local macro "map_simp" : tactic =>
103
+ `(tactic| simp only [map_ofNat, map_neg, map_add, map_sub, map_mul, map_pow, map_div₀,
104
+ Polynomial.map_ofNat, Polynomial.map_one, map_C, map_X, Polynomial.map_neg, Polynomial.map_add,
105
+ Polynomial.map_sub, Polynomial.map_mul, Polynomial.map_pow, Polynomial.map_div, coe_mapRingHom,
106
+ apply_ite <| mapRingHom _, WeierstrassCurve.map])
107
+
108
+ universe r s u v
101
109
102
110
namespace WeierstrassCurve
103
111
104
- variable {R : Type u} [CommRing R] (W : WeierstrassCurve R)
112
+ variable {R : Type r} {S : Type s} [CommRing R] [CommRing S ] (W : WeierstrassCurve R)
105
113
106
114
section Ψ₂Sq
107
115
@@ -115,11 +123,17 @@ noncomputable def ψ₂ : R[X][Y] :=
115
123
noncomputable def Ψ₂Sq : R[X] :=
116
124
C 4 * X ^ 3 + C W.b₂ * X ^ 2 + C (2 * W.b₄) * X + C W.b₆
117
125
118
- lemma C_Ψ₂Sq_eq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by
126
+ lemma C_Ψ₂Sq : C W.Ψ₂Sq = W.ψ₂ ^ 2 - 4 * W.toAffine.polynomial := by
119
127
rw [Ψ₂Sq, ψ₂, b₂, b₄, b₆, Affine.polynomialY, Affine.polynomial]
120
128
C_simp
121
129
ring1
122
130
131
+ lemma ψ₂_sq : W.ψ₂ ^ 2 = C W.Ψ₂Sq + 4 * W.toAffine.polynomial := by
132
+ rw [C_Ψ₂Sq, sub_add_cancel]
133
+
134
+ lemma Affine.CoordinateRing.mk_ψ₂_sq : mk W W.ψ₂ ^ 2 = mk W (C W.Ψ₂Sq) := by
135
+ rw [C_Ψ₂Sq, map_sub, map_mul, AdjoinRoot.mk_self, mul_zero, sub_zero, map_pow]
136
+
123
137
-- TODO: remove `twoTorsionPolynomial` in favour of `Ψ₂Sq`
124
138
lemma Ψ₂Sq_eq : W.Ψ₂Sq = W.twoTorsionPolynomial.toPoly :=
125
139
rfl
@@ -315,7 +329,7 @@ lemma Ψ_odd (m : ℕ) : W.Ψ (2 * (m + 2) + 1) =
315
329
else -W.preΨ' (m + 1 ) * W.preΨ' (m + 3 ) ^ 3 ) := by
316
330
repeat erw [Ψ_ofNat]
317
331
simp_rw [preΨ'_odd, if_neg (m + 2 ).not_even_two_mul_add_one, Nat.even_add_one, ite_not]
318
- split_ifs <;> C_simp <;> rw [C_Ψ₂Sq_eq ] <;> ring1
332
+ split_ifs <;> C_simp <;> rw [C_Ψ₂Sq ] <;> ring1
319
333
320
334
lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3 )) * W.ψ₂ =
321
335
W.Ψ (m + 2 ) ^ 2 * W.Ψ (m + 3 ) * W.Ψ (m + 5 ) - W.Ψ (m + 1 ) * W.Ψ (m + 3 ) * W.Ψ (m + 4 ) ^ 2 := by
@@ -327,6 +341,10 @@ lemma Ψ_even (m : ℕ) : W.Ψ (2 * (m + 3)) * W.ψ₂ =
327
341
lemma Ψ_neg (n : ℤ) : W.Ψ (-n) = -W.Ψ n := by
328
342
simp only [Ψ, preΨ_neg, C_neg, neg_mul (α := R[X][Y]), even_neg]
329
343
344
+ lemma Affine.CoordinateRing.mk_Ψ_sq (n : ℤ) : mk W (W.Ψ n) ^ 2 = mk W (C <| W.ΨSq n) := by
345
+ simp only [Ψ, ΨSq, map_one, map_mul, map_pow, one_pow, mul_pow, ite_pow, apply_ite C,
346
+ apply_ite <| mk W, mk_ψ₂_sq]
347
+
330
348
end Ψ
331
349
332
350
section Φ
@@ -386,6 +404,8 @@ section ψ
386
404
protected noncomputable def ψ (n : ℤ) : R[X][Y] :=
387
405
normEDS W.ψ₂ (C W.Ψ₃) (C W.preΨ₄) n
388
406
407
+ open WeierstrassCurve (Ψ ψ)
408
+
389
409
@[simp]
390
410
lemma ψ_zero : W.ψ 0 = 0 :=
391
411
normEDS_zero ..
@@ -418,6 +438,9 @@ lemma ψ_even (m : ℕ) : W.ψ (2 * (m + 3)) * W.ψ₂ =
418
438
lemma ψ_neg (n : ℤ) : W.ψ (-n) = -W.ψ n :=
419
439
normEDS_neg ..
420
440
441
+ lemma Affine.CoordinateRing.mk_ψ (n : ℤ) : mk W (W.ψ n) = mk W (W.Ψ n) := by
442
+ simp only [ψ, normEDS, Ψ, preΨ, map_mul, map_pow, map_preNormEDS, ← mk_ψ₂_sq, ← pow_mul]
443
+
421
444
end ψ
422
445
423
446
section φ
@@ -428,7 +451,7 @@ section φ
428
451
protected noncomputable def φ (n : ℤ) : R[X][Y] :=
429
452
C X * W.ψ n ^ 2 - W.ψ (n + 1 ) * W.ψ (n - 1 )
430
453
431
- open WeierstrassCurve (φ)
454
+ open WeierstrassCurve (Ψ Φ φ)
432
455
433
456
@[simp]
434
457
lemma φ_zero : W.φ 0 = 1 := by
@@ -458,6 +481,106 @@ lemma φ_neg (n : ℤ) : W.φ (-n) = W.φ n := by
458
481
rw [φ, ψ_neg, neg_sq (R := R[X][Y]), neg_add_eq_sub, ← neg_sub n, ψ_neg, ← neg_add', ψ_neg,
459
482
neg_mul_neg (α := R[X][Y]), mul_comm <| W.ψ _, φ]
460
483
484
+ lemma Affine.CoordinateRing.mk_φ (n : ℤ) : mk W (W.φ n) = mk W (C <| W.Φ n) := by
485
+ simp_rw [φ, Φ, map_sub, map_mul, map_pow, mk_ψ, mk_Ψ_sq, Ψ, map_mul,
486
+ mul_mul_mul_comm _ <| mk W <| ite .., Int.even_add_one, Int.even_sub_one, ← sq, ite_not,
487
+ apply_ite C, apply_ite <| mk W, ite_pow, map_one, one_pow, mk_ψ₂_sq]
488
+
461
489
end φ
462
490
491
+ section Map
492
+
493
+ /-! ### Maps across ring homomorphisms -/
494
+
495
+ open WeierstrassCurve (Ψ Φ ψ φ)
496
+
497
+ variable (f : R →+* S)
498
+
499
+ lemma map_ψ₂ : (W.map f).ψ₂ = W.ψ₂.map (mapRingHom f) := by
500
+ simp only [ψ₂, Affine.map_polynomialY]
501
+
502
+ lemma map_Ψ₂Sq : (W.map f).Ψ₂Sq = W.Ψ₂Sq.map f := by
503
+ simp only [Ψ₂Sq, map_b₂, map_b₄, map_b₆]
504
+ map_simp
505
+
506
+ lemma map_Ψ₃ : (W.map f).Ψ₃ = W.Ψ₃.map f := by
507
+ simp only [Ψ₃, map_b₂, map_b₄, map_b₆, map_b₈]
508
+ map_simp
509
+
510
+ lemma map_preΨ₄ : (W.map f).preΨ₄ = W.preΨ₄.map f := by
511
+ simp only [preΨ₄, map_b₂, map_b₄, map_b₆, map_b₈]
512
+ map_simp
513
+
514
+ lemma map_preΨ' (n : ℕ) : (W.map f).preΨ' n = (W.preΨ' n).map f := by
515
+ simp only [preΨ', map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS']
516
+ map_simp
517
+
518
+ lemma map_preΨ (n : ℤ) : (W.map f).preΨ n = (W.preΨ n).map f := by
519
+ simp only [preΨ, map_Ψ₂Sq, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_preNormEDS]
520
+ map_simp
521
+
522
+ lemma map_ΨSq (n : ℤ) : (W.map f).ΨSq n = (W.ΨSq n).map f := by
523
+ simp only [ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom]
524
+ map_simp
525
+
526
+ lemma map_Ψ (n : ℤ) : (W.map f).Ψ n = (W.Ψ n).map (mapRingHom f) := by
527
+ simp only [Ψ, map_preΨ, map_ψ₂, ← coe_mapRingHom]
528
+ map_simp
529
+
530
+ lemma map_Φ (n : ℤ) : (W.map f).Φ n = (W.Φ n).map f := by
531
+ simp only [Φ, map_ΨSq, map_preΨ, map_Ψ₂Sq, ← coe_mapRingHom]
532
+ map_simp
533
+
534
+ lemma map_ψ (n : ℤ) : (W.map f).ψ n = (W.ψ n).map (mapRingHom f) := by
535
+ simp only [ψ, map_ψ₂, map_Ψ₃, map_preΨ₄, ← coe_mapRingHom, map_normEDS]
536
+ map_simp
537
+
538
+ lemma map_φ (n : ℤ) : (W.map f).φ n = (W.φ n).map (mapRingHom f) := by
539
+ simp only [φ, map_ψ]
540
+ map_simp
541
+
542
+ end Map
543
+
544
+ section BaseChange
545
+
546
+ /-! ### Base changes across algebra homomorphisms -/
547
+
548
+ variable [Algebra R S] {A : Type u} [CommRing A] [Algebra R A] [Algebra S A] [IsScalarTower R S A]
549
+ {B : Type v} [CommRing B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] (f : A →ₐ[S] B)
550
+
551
+ lemma baseChange_ψ₂ : (W.baseChange B).ψ₂ = (W.baseChange A).ψ₂.map (mapRingHom f) := by
552
+ rw [← map_ψ₂, map_baseChange]
553
+
554
+ lemma baseChange_Ψ₂Sq : (W.baseChange B).Ψ₂Sq = (W.baseChange A).Ψ₂Sq.map f := by
555
+ rw [← map_Ψ₂Sq, map_baseChange]
556
+
557
+ lemma baseChange_Ψ₃ : (W.baseChange B).Ψ₃ = (W.baseChange A).Ψ₃.map f := by
558
+ rw [← map_Ψ₃, map_baseChange]
559
+
560
+ lemma baseChange_preΨ₄ : (W.baseChange B).preΨ₄ = (W.baseChange A).preΨ₄.map f := by
561
+ rw [← map_preΨ₄, map_baseChange]
562
+
563
+ lemma baseChange_preΨ' (n : ℕ) : (W.baseChange B).preΨ' n = ((W.baseChange A).preΨ' n).map f := by
564
+ rw [← map_preΨ', map_baseChange]
565
+
566
+ lemma baseChange_preΨ (n : ℤ) : (W.baseChange B).preΨ n = ((W.baseChange A).preΨ n).map f := by
567
+ rw [← map_preΨ, map_baseChange]
568
+
569
+ lemma baseChange_ΨSq (n : ℤ) : (W.baseChange B).ΨSq n = ((W.baseChange A).ΨSq n).map f := by
570
+ rw [← map_ΨSq, map_baseChange]
571
+
572
+ lemma baseChange_Ψ (n : ℤ) : (W.baseChange B).Ψ n = ((W.baseChange A).Ψ n).map (mapRingHom f) := by
573
+ rw [← map_Ψ, map_baseChange]
574
+
575
+ lemma baseChange_Φ (n : ℤ) : (W.baseChange B).Φ n = ((W.baseChange A).Φ n).map f := by
576
+ rw [← map_Φ, map_baseChange]
577
+
578
+ lemma baseChange_ψ (n : ℤ) : (W.baseChange B).ψ n = ((W.baseChange A).ψ n).map (mapRingHom f) := by
579
+ rw [← map_ψ, map_baseChange]
580
+
581
+ lemma baseChange_φ (n : ℤ) : (W.baseChange B).φ n = ((W.baseChange A).φ n).map (mapRingHom f) := by
582
+ rw [← map_φ, map_baseChange]
583
+
584
+ end BaseChange
585
+
463
586
end WeierstrassCurve
0 commit comments