Negation and addition formulae for nonsingular points in affine coordinates #
Let W
be a Weierstrass curve over a field F
with coefficients aᵢ
. The nonsingular affine
points on W
can be given negation and addition operations defined by a secant-and-tangent process.
Given a nonsingular affine point
P
, its negation-P
is defined to be the unique third nonsingular point of intersection betweenW
and the vertical line throughP
. Explicitly, ifP
is(x, y)
, then-P
is(x, -y - a₁x - a₃)
.Given two nonsingular affine points
P
andQ
, their additionP + Q
is defined to be the negation of the unique third nonsingular point of intersection betweenW
and the lineL
throughP
andQ
. Explicitly, letP
be(x₁, y₁)
and letQ
be(x₂, y₂)
.- If
x₁ = x₂
andy₁ = -y₂ - a₁x₂ - a₃
, thenL
is vertical. - If
x₁ = x₂
andy₁ ≠ -y₂ - a₁x₂ - a₃
, thenL
is the tangent ofW
atP = Q
, and has slopeℓ := (3x₁² + 2a₂x₁ + a₄ - a₁y₁) / (2y₁ + a₁x₁ + a₃)
. - Otherwise
x₁ ≠ x₂
, thenL
is the secant ofW
throughP
andQ
, and has slopeℓ := (y₁ - y₂) / (x₁ - x₂)
.
In the last two cases, the
X
-coordinate ofP + Q
is then the unique third solution of the equation obtained by substituting the lineY = ℓ(X - x₁) + y₁
into the Weierstrass equation, and can be written down explicitly asx := ℓ² + a₁ℓ - a₂ - x₁ - x₂
by inspecting the coefficients ofX²
. TheY
-coordinate ofP + Q
, after applying the final negation that mapsY
to-Y - a₁X - a₃
, is preciselyy := -(ℓ(x - x₁) + y₁) - a₁x - a₃
.- If
This file defines polynomials associated to negation and addition of nonsingular affine points,
including slopes of non-vertical lines. The actual group law on nonsingular points in affine
coordinates will be defined in Mathlib/AlgebraicGeometry/EllipticCurve/Affine/Point.lean
.
Main definitions #
WeierstrassCurve.Affine.negY
: theY
-coordinate of-P
.WeierstrassCurve.Affine.addX
: theX
-coordinate ofP + Q
.WeierstrassCurve.Affine.negAddY
: theY
-coordinate of-(P + Q)
.WeierstrassCurve.Affine.addY
: theY
-coordinate ofP + Q
.
Main statements #
WeierstrassCurve.Affine.equation_neg
: negation preserves the Weierstrass equation.WeierstrassCurve.Affine.nonsingular_neg
: negation preserves the nonsingular condition.WeierstrassCurve.Affine.equation_add
: addition preserves the Weierstrass equation.WeierstrassCurve.Affine.nonsingular_add
: addition preserves the nonsingular condition.
References #
J Silverman, The Arithmetic of Elliptic Curves
Tags #
elliptic curve, affine, negation, doubling, addition, group law
Negation formulae in affine coordinates #
The negation polynomial -Y - a₁X - a₃
associated to the negation of a nonsingular affine point
on a Weierstrass curve.
Equations
- W'.negPolynomial = -Polynomial.X - Polynomial.C (Polynomial.C W'.a₁ * Polynomial.X + Polynomial.C W'.a₃)
Instances For
Alias of WeierstrassCurve.Affine.equation_neg
.
Alias of WeierstrassCurve.Affine.equation_neg
.
Alias of WeierstrassCurve.Affine.nonsingular_neg
.
Alias of WeierstrassCurve.Affine.nonsingular_neg
.
Slope formulae in affine coordinates #
The line polynomial ℓ(X - x) + y
associated to the line Y = ℓ(X - x) + y
that passes through
a nonsingular affine point (x, y)
on a Weierstrass curve W
with a slope of ℓ
.
This does not depend on W
, and has argument order: x
, y
, ℓ
.
Equations
Instances For
The slope of the line through two nonsingular affine points (x₁, y₁)
and (x₂, y₂)
on a
Weierstrass curve W
.
If x₁ ≠ x₂
, then this line is the secant of W
through (x₁, y₁)
and (x₂, y₂)
, and has slope
(y₁ - y₂) / (x₁ - x₂)
. Otherwise, if y₁ ≠ -y₁ - a₁x₁ - a₃
, then this line is the tangent of W
at (x₁, y₁) = (x₂, y₂)
, and has slope (3x₁² + 2a₂x₁ + a₄ - a₁y₁) / (2y₁ + a₁x₁ + a₃)
. Otherwise,
this line is vertical, in which case this returns the value 0
.
This depends on W
, and has argument order: x₁
, x₂
, y₁
, y₂
.
Equations
Instances For
Addition formulae in affine coordinates #
The addition polynomial obtained by substituting the line Y = ℓ(X - x) + y
into the polynomial
W(X, Y)
associated to a Weierstrass curve W
. If such a line intersects W
at another
nonsingular affine point (x', y')
on W
, then the roots of this polynomial are precisely x
,
x'
, and the X
-coordinate of the addition of (x, y)
and (x', y')
.
This depends on W
, and has argument order: x
, y
, ℓ
.
Equations
- W'.addPolynomial x y ℓ = Polynomial.eval (WeierstrassCurve.Affine.linePolynomial x y ℓ) W'.polynomial
Instances For
The X
-coordinate of (x₁, y₁) + (x₂, y₂)
for two nonsingular affine points (x₁, y₁)
and
(x₂, y₂)
on a Weierstrass curve W
, where the line through them has a slope of ℓ
.
This depends on W
, and has argument order: x₁
, x₂
, ℓ
.
Instances For
The Y
-coordinate of -((x₁, y₁) + (x₂, y₂))
for two nonsingular affine points (x₁, y₁)
and
(x₂, y₂)
on a Weierstrass curve W
, where the line through them has a slope of ℓ
.
This depends on W
, and has argument order: x₁
, x₂
, y₁
, ℓ
.
Instances For
The Y
-coordinate of (x₁, y₁) + (x₂, y₂)
for two nonsingular affine points (x₁, y₁)
and
(x₂, y₂)
on a Weierstrass curve W
, where the line through them has a slope of ℓ
.
This depends on W
, and has argument order: x₁
, x₂
, y₁
, ℓ
.
Instances For
The formula x(P₁ + P₂) = x(P₁ - P₂) - ψ(P₁)ψ(P₂) / (x(P₂) - x(P₁))²
,
where ψ(x,y) = 2y + a₁x + a₃
.
The formula y(P₁)(x(P₂) - x(P₃)) + y(P₂)(x(P₃) - x(P₁)) + y(P₃)(x(P₁) - x(P₂)) = 0
,
assuming that P₁ + P₂ + P₃ = O
.
The formula ψ(P₁ + P₂) = (ψ(P₂)(x(P₁) - x(P₃)) - ψ(P₁)(x(P₂) - x(P₃))) / (x(P₂) - x(P₁))
,
where ψ(x,y) = 2y + a₁x + a₃
.