Skip to content

Commit d3af796

Browse files
Merge pull request #1953 from andres-erbsen/weier2
Update Curves/Weierstrass
2 parents 9158342 + 6f13372 commit d3af796

File tree

4 files changed

+469
-215
lines changed

4 files changed

+469
-215
lines changed

src/Curves/Weierstrass/Affine.v

+59-7
Original file line numberDiff line numberDiff line change
@@ -1,20 +1,72 @@
11
Require Import Crypto.Spec.WeierstrassCurve.
22
Require Import Crypto.Algebra.Field.
33
Require Import Crypto.Util.Decidable Crypto.Util.Tactics.DestructHead Crypto.Util.Tactics.BreakMatch.
4+
Require Import Crypto.Util.Tactics.SetoidSubst.
5+
Import RelationClasses Morphisms.
46

57
Module W.
68
Section W.
79
Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv} {a b:F}
810
{field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
911
{Feq_dec:DecidableRel Feq}.
12+
Local Infix "+" := Fadd. Local Infix "-" := Fsub.
13+
Local Infix "*" := Fmul. Local Infix "/" := Fdiv.
14+
Local Notation "x ^ 2" := (x*x) (at level 30).
15+
Local Notation point := (@W.point F Feq Fadd Fmul a b).
1016

11-
Program Definition opp (P:@W.point F Feq Fadd Fmul a b) : @W.point F Feq Fadd Fmul a b
12-
:= match W.coordinates P return F*F+_ with
13-
| inl (x1, y1) => inl (x1, Fopp y1)
14-
| inr tt => inr tt
15-
end.
16-
Next Obligation.
17-
cbv [W.coordinates]; break_match; trivial; fsatz.
17+
Definition opp (P : point) : point. refine (exist _ (
18+
match W.coordinates P with
19+
| inl (x1, y1) => inl (x1, Fopp y1)
20+
| inr tt => inr tt
21+
end) _).
22+
Proof. abstract (cbv [W.coordinates]; break_match; trivial; fsatz). Defined.
23+
24+
Global Instance Equivalence_eq : Equivalence (@W.eq _ Feq Fadd Fmul a b).
25+
Proof.
26+
cbv [W.eq W.coordinates]; split; repeat intros [ [ []|[] ] ?]; intuition try solve
27+
[contradiction | apply reflexivity | apply symmetry; trivial | eapply transitivity; eauto 1].
28+
Qed.
29+
30+
Global Instance Proper_opp : Proper (W.eq ==> W.eq) opp.
31+
Proof.
32+
repeat (intros [ [[]|[] ]?] || intro); cbv [W.coordinates opp W.eq] in *;
33+
repeat (try destruct_head' @and; try case dec as []; try contradiction; try split); trivial.
34+
setoid_subst_rel Feq; reflexivity.
35+
Qed.
36+
37+
(* Weierstraß Elliptic Curves and Side-Channel Attacks
38+
by Eric Brier and Marc Joye, 2002 *)
39+
Definition add' (P1 P2 : point) : point. refine (exist _
40+
match W.coordinates P1, W.coordinates P2 with
41+
| inl (x1, y1), inl (x2, y2) =>
42+
if dec (Feq y1 (Fopp y2)) then
43+
if dec (Feq x1 x2) then inr tt
44+
else let k := (y2-y1)/(x2-x1) in
45+
let x3 := k^2-x1-x2 in
46+
let y3 := k*(x1-x3)-y1 in
47+
inl (x3, y3)
48+
else let k := ((x1^2 + x1*x2 + x2^2 + a)/(y1+y2)) in
49+
let x3 := k^2-x1-x2 in
50+
let y3 := k*(x1-x3)-y1 in
51+
inl (x3, y3)
52+
| inr tt, inr tt => inr tt
53+
| inr tt, _ => W.coordinates P2
54+
| _, inr tt => W.coordinates P1
55+
end _).
56+
Proof. abstract (cbv [W.coordinates]; break_match; trivial; fsatz). Defined.
57+
58+
Lemma add'_correct char_ge_3 : forall P Q : point, W.eq (W.add' P Q) (W.add(char_ge_3:=char_ge_3) P Q).
59+
Proof. intros [ [[]|]?] [ [[]|]?]; cbv [W.coordinates W.add W.add' W.eq]; break_match; try split; try fsatz. Qed.
60+
61+
Global Instance Proper_add' : Proper (W.eq ==> W.eq ==> W.eq) add'.
62+
Proof.
63+
repeat (intros [ [[]|[] ]?] || intro); cbv [W.coordinates W.add' W.eq] in *;
64+
repeat (try destruct_head' @and; try case dec as []; try contradiction; try split); trivial.
65+
Time par : fsatz. (* setoid_subst_rel is slower *)
1866
Qed.
67+
68+
Global Instance Proper_add {char_ge_3} :
69+
Proper (W.eq ==> W.eq ==> W.eq) (@W.add _ Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv _ _ char_ge_3 a b).
70+
Proof. repeat intro. rewrite <-2add'_correct. apply Proper_add'; trivial. Qed.
1971
End W.
2072
End W.

src/Curves/Weierstrass/Jacobian/CoZ.v

+54-33
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,7 @@ Module Jacobian.
2424
Local Infix "+" := Fadd. Local Infix "-" := Fsub.
2525
Local Infix "*" := Fmul. Local Infix "/" := Fdiv.
2626
Local Notation "x ^ 2" := (x*x). Local Notation "x ^ 3" := (x^2*x).
27-
Local Notation "2" := (1+1). Local Notation "4" := (2+2).
27+
Local Notation "2" := (1+1). Local Notation "4" := (1+1+1+1).
2828
Local Notation "8" := (4+4). Local Notation "27" := (4*4 + 4+4 +1+1+1).
2929
Local Notation Wpoint := (@W.point F Feq Fadd Fmul a b).
3030
Context {char_ge_12:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 12%positive}
@@ -67,8 +67,8 @@ Module Jacobian.
6767
| _ => progress destruct_head'_sum
6868
| _ => progress destruct_head'_bool
6969
| _ => progress destruct_head'_or
70-
| H: context[@dec ?P ?pf] |- _ => destruct (@dec P pf)
71-
| |- context[@dec ?P ?pf] => destruct (@dec P pf)
70+
| H: context[@dec ?P ?pf] |- _ => destruct (@dec P pf) || destruct (@dec P pf) at 1
71+
| |- context[@dec ?P ?pf] => destruct (@dec P pf) || destruct (@dec P pf) at 1
7272
| |- ?P => lazymatch type of P with Prop => split end
7373
end.
7474
Ltac prept := repeat prept_step.
@@ -405,36 +405,57 @@ Module Jacobian.
405405
end;
406406
fsatz ].
407407

408-
Hint Unfold Jacobian.double negb andb : points_as_coordinates.
409-
Hint Unfold W.eq W.add Jacobian.to_affine Jacobian.of_affine Jacobian.add Jacobian.add_impl Jacobian.opp : points_as_coordinates.
410-
411-
(* These could go into Jacobian.v *)
412-
Global Instance Proper_opp : Proper (eq ==> eq) opp. Proof. faster_t_noclear. Qed.
408+
Local Hint Unfold W.add W.add' W.zero W.opp : points_as_coordinates.
409+
Local Hint Unfold Jacobian.double Jacobian.double_impl negb andb : points_as_coordinates.
410+
Local Hint Unfold Jacobian.to_affine Jacobian.to_affine_impl Jacobian.of_affine Jacobian.of_affine_impl Jacobian.add Jacobian.add_impl Jacobian.opp : points_as_coordinates.
413411

414412
Lemma of_affine_add P Q
415413
: eq (of_affine (W.add P Q)) (add (of_affine P) (of_affine Q)).
416-
Proof. t. Qed.
417-
418-
Lemma add_opp (P : point) :
419-
z_of (add P (opp P)) = 0.
420-
Proof. faster_t_noclear. Qed.
414+
Proof. rewrite Jacobian.eq_iff, Jacobian.to_affine_add, 3Jacobian.to_affine_of_affine; reflexivity. Qed.
421415

422416
Lemma add_comm (P Q : point) :
423417
eq (add P Q) (add Q P).
424-
Proof. faster_t_noclear. Qed.
418+
Proof.
419+
pose proof W.commutative_group(discriminant_nonzero:=discriminant_nonzero) _.
420+
rewrite Jacobian.eq_iff, 2Jacobian.to_affine_add.
421+
rewrite Hierarchy.commutative. reflexivity.
422+
Qed.
425423

426424
Lemma add_zero_l (P Q : point) (H : z_of P = 0) :
427425
eq (add P Q) Q.
428-
Proof. faster_t. Qed.
426+
Proof.
427+
pose proof W.commutative_group(discriminant_nonzero:=discriminant_nonzero) _.
428+
rewrite Jacobian.eq_iff, Jacobian.to_affine_add.
429+
rewrite (proj1 (Jacobian.iszero_iff P)), Hierarchy.left_identity; [reflexivity|].
430+
case P as [ [ []?] ?]; cbv [Jacobian.iszero z_of proj1_sig snd] in *; trivial.
431+
Qed.
429432

430433
Lemma add_zero_r (P Q : point) (H : z_of Q = 0) :
431434
eq (add P Q) P.
432-
Proof. faster_t. Qed.
435+
Proof. rewrite add_comm, add_zero_l; trivial; reflexivity. Qed.
433436

434437
Lemma add_double (P Q : point) :
435438
eq P Q ->
436439
eq (add P Q) (double P).
437-
Proof. faster_t_noclear. Qed.
440+
Proof.
441+
rewrite 2Jacobian.eq_iff, Jacobian.to_affine_add, Jacobian.to_affine_double.
442+
intros ->; reflexivity.
443+
Qed.
444+
445+
Lemma add_opp_same_r (P : point) :
446+
eq (add P (opp P)) (of_affine W.zero).
447+
Proof.
448+
pose proof W.commutative_group(discriminant_nonzero:=discriminant_nonzero) _.
449+
rewrite Jacobian.eq_iff, Jacobian.to_affine_add, Jacobian.to_affine_of_affine, Jacobian.to_affine_opp.
450+
rewrite Hierarchy.right_inverse. reflexivity.
451+
Qed.
452+
453+
Lemma z_of_eq_zero (P : point) : eq P (of_affine W.zero) -> z_of P = 0.
454+
Proof. prept. match goal with H : 0 <> 0 |- _ => case (H ltac:(reflexivity)) end. Qed.
455+
456+
Lemma z_of_add_opp_same_r (P : point) :
457+
z_of (add P (opp P)) = 0.
458+
Proof. apply z_of_eq_zero, add_opp_same_r. Qed.
438459

439460
(* This uses assumptions not present in Jacobian.v,
440461
namely char_ge_12 and discriminant_nonzero. *)
@@ -465,7 +486,7 @@ Module Jacobian.
465486

466487
Lemma opp_co_z (P : point) :
467488
co_z P (opp P).
468-
Proof. unfold co_z; faster_t. Qed.
489+
Proof. unfold co_z. prept. Qed.
469490

470491
Program Definition make_co_z (P Q : point) (HQaff : z_of Q = 1) : point*point :=
471492
match proj1_sig P, proj1_sig Q return (F*F*F)*(F*F*F) with
@@ -479,8 +500,8 @@ Module Jacobian.
479500
let t2 := t3 * t2 in
480501
(P, (t1, t2, t3))
481502
end.
482-
Next Obligation. Proof. faster_t. Qed.
483-
Next Obligation. Proof. faster_t. Qed.
503+
Next Obligation. Proof. prept. Qed.
504+
Next Obligation. Proof. prept. par: faster_t. Qed.
484505

485506
Hint Unfold is_point co_z make_co_z : points_as_coordinates.
486507

@@ -520,17 +541,17 @@ Module Jacobian.
520541
let t5 := t5 - t2 in
521542
((t4, t5, t3), (t1, t2, t3))
522543
end.
523-
Next Obligation. Proof. faster_t_noclear. Qed.
524-
Next Obligation. Proof. faster_t. Qed.
544+
Next Obligation. Proof. prept. all : faster_t_noclear. Qed.
545+
Next Obligation. Proof. prept. all : faster_t. Qed.
525546

526-
Hint Unfold zaddu : points_as_coordinates.
547+
Hint Unfold zaddu Jacobian.add_nz_nz : points_as_coordinates.
527548

528549
(* ZADDU(P, Q) = (P + Q, P) if P <> Q, Q <> -P *)
529550
Lemma zaddu_correct (P Q : point) (H : co_z P Q)
530551
(Hneq : x_of P <> x_of Q):
531552
let '(R1, R2) := zaddu P Q H in
532553
eq (add P Q) R1 /\ eq P R2 /\ co_z R1 R2.
533-
Proof. faster_t_noclear. Qed.
554+
Proof. prept. par : faster_t_noclear. Qed.
534555

535556
Lemma zaddu_correct_alt (P Q : point) (H : co_z P Q) :
536557
let '(R1, R2) := zaddu P Q H in
@@ -546,7 +567,7 @@ Module Jacobian.
546567
Lemma zaddu_correct0 (P : point) :
547568
let '(R1, R2) := zaddu P (opp P) (opp_co_z P) in
548569
z_of R1 = 0 /\ co_z R1 R2.
549-
Proof. faster_t_noclear. Qed.
570+
Proof. prept. all : faster_t_noclear. Qed.
550571

551572
(* Scalar Multiplication on Weierstraß Elliptic Curves from Co-Z Arithmetic *)
552573
(* Goundar, Joye, Miyaji, Rivain, Vanelli *)
@@ -587,16 +608,16 @@ Module Jacobian.
587608
let t2 := t2 + t6 in
588609
((t1, t2, t3), (t4, t5, t3))
589610
end.
590-
Next Obligation. Proof. faster_t_noclear. Qed.
591-
Next Obligation. Proof. faster_t_noclear. Qed.
611+
Next Obligation. Proof. prept. all : faster_t_noclear. Qed.
612+
Next Obligation. Proof. prept. all : faster_t_noclear. Qed.
592613

593614
Hint Unfold zaddc : points_as_coordinates.
594615
(* ZADDC(P, Q) = (P + Q, P - Q) if P <> Q, Q <> -P *)
595616
Lemma zaddc_correct (P Q : point) (H : co_z P Q)
596617
(Hneq : x_of P <> x_of Q):
597618
let '(R1, R2) := zaddc P Q H in
598619
eq (add P Q) R1 /\ eq (add P (opp Q)) R2 /\ co_z R1 R2.
599-
Proof. faster_t_noclear. Qed.
620+
Proof. prept. par : faster_t_noclear. Qed.
600621

601622
Lemma zaddc_correct_alt (P Q : point) (H : co_z P Q) :
602623
let '(R1, R2) := zaddc P Q H in
@@ -735,7 +756,7 @@ Module Jacobian.
735756
rewrite add_assoc, add_comm. reflexivity.
736757
- rewrite <- A2, <- B1, <- B2.
737758
rewrite (add_comm P Q).
738-
rewrite add_assoc. rewrite add_zero_r; [reflexivity|apply add_opp].
759+
rewrite add_assoc. rewrite add_zero_r; [reflexivity|apply z_of_add_opp_same_r].
739760
Qed.
740761

741762
Lemma zdau_naive_correct_alt (P Q : point) (H : co_z P Q)
@@ -756,7 +777,7 @@ Module Jacobian.
756777
rewrite add_assoc, add_comm. reflexivity.
757778
- rewrite <- A2, <- B1, <- B2.
758779
rewrite (add_comm P Q).
759-
rewrite add_assoc. rewrite add_zero_r; [reflexivity|apply add_opp].
780+
rewrite add_assoc. rewrite add_zero_r; [reflexivity|apply z_of_add_opp_same_r].
760781
Qed.
761782

762783
(* Scalar Multiplication on Weierstraß Elliptic Curves from Co-Z Arithmetic *)
@@ -816,16 +837,16 @@ Module Jacobian.
816837
let t5 := t7 - t5 in
817838
((t1, t2, t3), (t4, t5, t3))
818839
end.
819-
Next Obligation. Proof. faster_t_noclear. Qed.
820-
Next Obligation. Proof. faster_t_noclear. Qed.
840+
Next Obligation. Proof. prept. par:setoid_subst_rel Feq; faster_t_noclear. Qed.
841+
Next Obligation. Proof. prept. par:faster_t_noclear. Qed.
821842

822843
Hint Unfold zdau : points_as_coordinates.
823844

824845
Lemma zdau_naive_eq_zdau (P Q : point) (H : co_z P Q) :
825846
let '(R1, R2) := zdau_naive P Q H in
826847
let '(S1, S2) := zdau P Q H in
827848
eq R1 S1 /\ eq R2 S2.
828-
Proof. faster_t. all: try fsatz. Qed.
849+
Proof. prept. par : faster_t; try fsatz. Qed.
829850

830851
(* Direct proof is intensive, which is why we need the naive implementation *)
831852
Lemma zdau_correct (P Q : point) (H : co_z P Q)

0 commit comments

Comments
 (0)