Rational Points and Group Structure over Finite Fields

Contents

Enumeration of Points

Points(J) : JacHyp -> SetIndx
RationalPoints(J) : JacHyp -> SetIndx
Given a Jacobian J of a hyperelliptic curve defined over a finite field, determine all rational points on the Jacobian J.

Counting Points on the Jacobian

Several algorithms are used to compute the order of a Jacobian depending of its size, its genus and its type. In particular, in genus 2 it includes all the techniques described in [GH00]. The best current algorithms are the ones based on p-adic liftings. These are the defaults when they apply (see below) and the characteristic is not too large. In odd characteristic, Kedlaya's algorithm is used as described in [Ked01]. For characteristic 2, we have Mestre's canonical lift method as adapted by Lercier and Lubicz. Details can be found in [LL]. We also now have an implementation of Vercauteren's characteristic 2 version of Kedlaya (described in [Ver02]). As an added bonus, the p-adic methods actually give the Euler factor (see below). A verbose flag can be set to see which strategy is chosen and the progress of the computation.

The latest p-adic methods, faster than Kedlaya in odd characteristic, are those of Alan Lauder based on deformations over parametrised families. Magma incorporates an implementation for hyperelliptic families by Hendrik Hubrechts (see [Hub06] for details). This has a slightly different interface to the other methods (which all run through the same top-level intrinsics) as it can be used to count points on several members of the family at once.

SetVerbose("JacHypCnt", v) : MonStgElt, RngIntElt ->
Set the verbose printing level for the point-counting on the Jacobians. Currently the legal values for v are true, false, 0, 1, 2, 3 or 4 (false is the same as 0, and true is the same as 1).
# J : JacHyp -> RngIntElt
Order(J) : JacHyp -> RngIntElt
Given the Jacobian J of a hyperelliptic curve defined over a finite field, determine the order of the group of rational points.

There are 4 optional parameters which concern every genus.

     NaiveAlg: BoolElt                   Default: false
     ShanksLimit: RngIntElt              Default: 1012
     CartierManinLimit: RngIntElt        Default: 5 x 105
     UseSubexpAlg: BoolElt               Default: true
The parameter NaiveAlg can be set to one if one want to use the naive algorithm for point counting, consisting in counting points on the curve over the first g extensions, where g is the genus. Setting this parameter to true annihilates the effects of the others. By default, this strategy is used for small groups and Jacobians in characteristic 2 for which no group law is available.

In odd characteristic, Kedlaya's algorithm can be applied to curves of any genus and it is the default when the group is not too small or the characteristic too large.

In contrast, Mestre's characteristic 2 method has several conditions on its applicability. When these are satisfied and the genus is not too large or the group order too small, it is the default.

Firstly, the Jacobian must be ordinary. This is equivalent to the defining polynomial h(x) having degree g or g + 1 and having distinct roots (ie non-zero discriminant).
Next, there must be a group law available for a final elimination: the algorithm produces several possibilities for the characteristic polynomial of Frobenius. If no group law is available, we attempt to replace the curve with an isomorphic one for which one is. Note that a group law exists if there is a single point at infinity (degree of h(x) = g).

{Warning:} Even if these conditions are satisfied, the last phase of the algorithm may still fail for g ≥4 in unusual cases. However, this is very rare. Finally, it should be noted that the main algorithm works over the smallest extension of the ground field K over which h(x) splits. It is therefore much faster if h(x) has all its roots in K.

When the genus is > 4 or the canonical lift method doesn't apply, Vercauteren's version of Kedlaya is now used in its place. It is also used for genus 4 when an extension of the base field is required for Mestre: ie, h(x) doesn't split.

     UseGenus2: BoolElt                  Default: false
For genus 2 the methods mentioned at the start are applied if the parameter UseGenus2 is set to true. Otherwise they are now only used for large characteristic.

If none of the above apply, the remaining methods are the subexponential algorithm, which uses the generic machinery of the function fields, and general group order algorithms of Shanks and Pollard. The subexponential algorithm will be the first choice if no group law is available or the genus is large and the base field small. It can be disabled by setting UseSubexpAlg to false.

The parameter ShanksLimit determines when we switch from a memory-consuming Shanks' algorithm to a slower Pollard's method. The value of the parameter CartierManinLimit is the maximal characteristic for which we may use the Cartier-Manin trick (see function EulerFactorModChar below).

The next 2 parameters are used only for the genus 2 methods in odd characteristic.

     UseSchoof: BoolElt                  Default: true
     UseHalving: BoolElt                 Default: true
When the base field is large enough, it is better to firstly compute the cardinality of the Jacobian modulo some odd primes and some power of two. These two parameters allow the user to disable one or both of these methods.

Example CrvHyp_Jac_Point_Counting (H134E17)

In the following sequence of examples the different point counting methods are illustrated. The first example uses direct counting.

> SetVerbose("JacHypCnt",true);
> P<x> := PolynomialRing(GF(31));
> f := x^8 + 3*x^7 + 2*x^6 + 23*x^5 + 5*x^4 + 21*x^3 + 29*x^2 + 12*x + 9;
> J := Jacobian(HyperellipticCurve(f));
> time #J;
Using naive counting algorithm.
20014
Time: 0.020

Example CrvHyp_kedlaya (H134E18)

For the second example, we apply Kedlaya's algorithm to a genus curve 2 over GF(320).
> K := FiniteField(3,20);
> P<x> := PolynomialRing(K);
> f :=  x^5 + x + K.1;
> J := Jacobian(HyperellipticCurve(f));
> #J;
Using Kedlaya's algorithm.
Applying Kedlaya's algorithm
Total time: 2.310
12158175772023384771

Example CrvHyp_kedlaya2 (H134E19)

For the third example, we apply Vercauteren's algorithm to a non-ordinary genus curve 2 over GF(225)..
> K := FiniteField(2,25);
> P<x> := PolynomialRing(K);
> f :=  x^5 + x^3 + x^2 + K.1;
> J := Jacobian(HyperellipticCurve([f,K!1]));
> time #J;
Using Kedlaya/Vercauteren's algorithm.
1125899940397057
Time: 0.680

Example CrvHyp_mestre (H134E20)

For the third example, we apply Mestre's algorithm to a genus 3 ordinary curve over GF(225).
> K := FiniteField(2,25);
> P<x> := PolynomialRing(K);
> h := x*(x+1)*(x+K.1);
> f :=  x^7 + x^5 + x + K.1;
> C := HyperellipticCurve([f,h]);
> J := Jacobian(C);
> #J;
Using Mestre's method.
Total time: 0.440
37780017712685037895040
> SetVerbose("JacHypCnt",false);

As the full Euler factor of J has been computed and stored the number of points on C, the Zeta function for C and the Euler factor for C are now immediate.

> time #C;
33555396
Time: 0.000

Example CrvHyp_shanks-pollard (H134E21)

We compare the Shanks and Pollard methods for large prime fields.

> // Comparison between Shanks and Pollard:
> P<x> := PolynomialRing(GF(1000003));
> f := x^7 + 123456*x^6 + 123*x^5 + 456*x^4 + 98*x^3 + 76*x^2 + 54*x + 32;
> J := Jacobian(HyperellipticCurve(f));
> curr_mem := GetMemoryUsage();
> ResetMaximumMemoryUsage();
> time Order(J : ShanksLimit := 10^15);
1001800207033014252
Time: 19.140
> GetMaximumMemoryUsage() - curr_mem;
133583360

The computation took about 100 MB of central memory.

> delete J`Order;   // clear the result which has been stored
> curr_mem := GetMemoryUsage();
> ResetMaximumMemoryUsage();
> time Order(J : ShanksLimit := 0);
1001800207033014252
Time: 95.670
> GetMaximumMemoryUsage() - curr_mem;
0

Now it takes almost no memory, but it is slower (and runtime may vary a lot).

Example CrvHyp_shanks-pollard (H134E22)

In the case of genus 2 curves, the parameters UseSchoof and UseHalving do help.

> // Using Schoof and Halving true is generally best in genus 2
> P<x> := PolynomialRing(GF(100000007));
> f := x^5 + 456*x^4 + 98*x^3 + 76*x^2 + 54*x + 32;
> J := Jacobian(HyperellipticCurve(f));
> time Order(J);
10001648178050390
Time: 7.350
> delete J`Order;
> time Order(J: UseSchoof := false, UseHalving := false);
10001648178050390
Time: 21.080

But if the Jacobian is known in advance to be highly non--cyclic, it may be slightly better to switch them off. The Jacobian below is the direct product of two copies of the same supersingular elliptic curve.

> // ... but not always for highly non-cyclic Jacobians
> P<x> := PolynomialRing(GF(500083));
> f := x^5 + 250039*x^4 + 222262*x^3 + 416734*x^2 + 166695*x + 222259;
> J := Jacobian(HyperellipticCurve(f));
> time Order(J);
250084007056
Time: 1.920
> delete J`Order;
> time Order(J : UseSchoof:=false, UseHalving := false);
250084007056
Time: 1.870
FactoredOrder(J) : JacHyp -> [ <RngIntElt, RngIntElt> ]
Given the Jacobian J of a hyperelliptic curve defined over a finite field, the function returns the factorization of the order of the group of rational points.
EulerFactor(J) : JacHyp -> RngUPolElt
Given the Jacobian J of a hyperelliptic curve defined over a finite field, the function returns the Euler factor J, i.e., the reciprocal of the characteristic polynomial of Frobenius acting on H1(J). (see also ZetaFunction(C) which is essentially the same function and has the same behaviour).
EulerFactorModChar(J) : JacHyp -> RngUPolElt
Given the Jacobian J of a hyperelliptic curve defined over a finite field, the function returns the Euler factor J modulo the characteristic of the base field. This function should not be used in high characteristic (say p should be ≤106).
EulerFactor(J, K) : JacHyp, FldFin -> RngUPolElt
Given a Jacobian J of a hyperelliptic curve defined over the rationals and a finite field K at which J has good reduction, the function returns the Euler factor of the base extension of J to K.

Deformation Point Counting

JacobianOrdersByDeformation(Q, Y) : RngMPolElt, SeqEnum -> SeqEnum
EulerFactorsByDeformation(Q, Y) : RngMPolElt, SeqEnum -> SeqEnum
ZetaFunctionsByDeformation(Q, Y) : RngMPolElt, SeqEnum -> SeqEnum
These functions compute the orders of Jacobians (resp. Euler factors, resp. Zeta functions) of one or more hyperelliptic curves in a 1-parameter family over a finite field using deformation methods.

The 2-variable polynomial Q(x, z) should be in variables x and z over a finite field of odd characteristic k. Currently Q must be monic of odd degree as a polynomial in x. z is the parameter and the family of curves is given by y2=Q(x, z).

The sequence Y should contain elements in a finite field extension K of k. The function then computes and returns the sequence of orders, Euler factors or Zeta functions associated to the hyperelliptic curves over K obtained by specialising the z parameter of Q(x, z) to the elements of Y.

Over a field of size pn, for n not too large, it is efficient to compute the results for several curves in the family at once, as roughly half of the computation for an individual curve is taken up in computing a Frobenius matrix that is then specialised to the parameter value. For a sequence of input parameter values, this part of the computation is only performed once at the start. However, as n becomes larger, the time taken for the specialisation at a particular value starts to dominate.

The field k over which the family is defined should be of small degree over the prime field. This is because Kedlaya's algorithm is applied over k to a particular member of the family with the parameter specialised to a k-value to initialise the deformation.

Similarly, the algorithm is much faster when Q(x, z) is of small degree in the parameter z. In practice, the degree of Q(x, z) as a polynomial in z, probably shouldn't exceed 3 or 4 (linear is obviously best).

There are two further conditions on Q and Y which we will try to remove in the near future. The first is that no Y-value is zero. The second is that the generic discriminant of hat(Q) (resultant of hat(Q) and ∂hat(Q)/∂x w.r.t. x) must have leading coefficient (as a polynomial in z) a unit in W(k). Here, W(k) is the ring of integers of the unramified extension of Qp with residue class field k and hat(Q) is a lift of Q to a 2-variable polynomial over W(k). By translating and scaling the parameter, the user may be able to effect these two conditions if they aren't initially satisfied.

As with the other point-counting functions, the verbose flag JacHypCnt can be used to output information during processing.

Example CrvHyp_def_hyp_pt_cnt_ex (H134E23)

The following is a small example where the Euler factors for 4 random members of a linear family of elliptic curves are computed with a single call. The base field is F9 and the parameter values are taken in F_(340) which is the field over which the counting occurs.
> Fp := FiniteField(3^2);
> R<X,Y> := PolynomialRing(Fp,2);
> Q  := X^3+X^2-2*X+Fp.1+Y*(X^2-X+1);
> Fq := ext<Fp | 20>;
> values := [Random(Fq) : i in [1..4]];
> EFs := EulerFactorsByDeformation(Q,values);
> EFs;
[
    12157665459056928801*$.1^2 + 852459373*$.1 + 1,
    12157665459056928801*$.1^2 + 1088717005*$.1 + 1,
    12157665459056928801*$.1^2 + 6911064226*$.1 + 1,
    12157665459056928801*$.1^2 - 607949990*$.1 + 1
]

Abelian Group Structure

Sylow(J, p) : JacHyp, RngIntElt -> GrpAb, Map, Eseq
Given the Jacobian J of a hyperelliptic curve defined over a finite field and a prime p, this function returns the Sylow p-subgroup of the group of rational points of J, as an abstract abelian group A. The injection from A to J is also returned as well as the generators of the p-Sylow subgroup.

AbelianGroup(J) : JacHyp -> GrpAb, Map
    UseGenerators: Bool                 Default: false
    Generators: SetEnum                 Default: 
Given the Jacobian J of a hyperelliptic curve defined over a finite field K, this function returns the group of rational points of J as an abstract abelian group A. The isomorphism from A to J(K) is returned as a second value.

If UseGenerators is set then the group structure computation is achieved by extracting relations from the user-supplied set of generators in Generators.

HasAdditionAlgorithm(J) : JacHyp -> Bool
Returns true if and only if the Jacobian J has an addition algorithm. This is of interest when trying to construct the (abelian) group structure of J: When no user-supplied generators are given, such an algorithm must be present.

V2.28, 13 July 2023