aboutsummaryrefslogtreecommitdiff
path: root/src/Curves/Montgomery
diff options
context:
space:
mode:
authorGravatar Andres Erbsen <andreser@mit.edu>2017-04-06 22:53:07 -0400
committerGravatar Andres Erbsen <andreser@mit.edu>2017-04-06 22:53:07 -0400
commitc9fc5a3cdf1f5ea2d104c150c30d1b1a6ac64239 (patch)
treedb7187f6984acff324ca468e7b33d9285806a1eb /src/Curves/Montgomery
parent21198245dab432d3c0ba2bb8a02254e7d0594382 (diff)
rename-everything
Diffstat (limited to 'src/Curves/Montgomery')
-rw-r--r--src/Curves/Montgomery/Affine.v67
-rw-r--r--src/Curves/Montgomery/AffineProofs.v83
-rw-r--r--src/Curves/Montgomery/XZ.v57
-rw-r--r--src/Curves/Montgomery/XZProofs.v58
4 files changed, 265 insertions, 0 deletions
diff --git a/src/Curves/Montgomery/Affine.v b/src/Curves/Montgomery/Affine.v
new file mode 100644
index 000000000..721908a6a
--- /dev/null
+++ b/src/Curves/Montgomery/Affine.v
@@ -0,0 +1,67 @@
+Require Import Crypto.Algebra.Field.
+Require Import Crypto.Util.GlobalSettings.
+Require Import Crypto.Util.Sum Crypto.Util.Prod.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Spec.MontgomeryCurve Crypto.Spec.WeierstrassCurve.
+
+Module M.
+ Section MontgomeryCurve.
+ Import BinNat.
+ Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {Feq_dec:Decidable.DecidableRel Feq}
+ {char_ge_3:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul (BinNat.N.succ_pos (BinNat.N.two))}.
+ Local Infix "=" := Feq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
+ Local Infix "+" := Fadd. Local Infix "*" := Fmul.
+ Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
+ Local Notation "- x" := (Fopp x).
+ Local Notation "x ^ 2" := (x*x) (at level 30).
+ Local Notation "x ^ 3" := (x*x^2) (at level 30).
+ Local Notation "0" := Fzero. Local Notation "1" := Fone.
+ Local Notation "'∞'" := unit : type_scope.
+ Local Notation "'∞'" := (inr tt) : core_scope.
+ Local Notation "( x , y )" := (inl (pair x y)).
+ Local Open Scope core_scope.
+
+ Context {a b: F} {b_nonzero:b <> 0}.
+
+ Program Definition opp (P:@M.point F Feq Fadd Fmul a b) : @M.point F Feq Fadd Fmul a b :=
+ match P return F*F+∞ with
+ | (x, y) => (x, -y)
+ | ∞ => ∞
+ end.
+ Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.
+
+ Local Notation add := (M.add(b_nonzero:=b_nonzero)).
+ Local Notation point := (@M.point F Feq Fadd Fmul a b).
+
+ Section MontgomeryWeierstrass.
+ Local Notation "2" := (1+1).
+ Local Notation "3" := (1+2).
+ Local Notation "4" := (1+3).
+ Local Notation "16" := (4*4).
+ Local Notation "9" := (3*3).
+ Local Notation "27" := (3*9).
+ Context {char_ge_28:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 28}.
+
+
+ Local Notation WeierstrassA := ((3-a^2)/(3*b^2)).
+ Local Notation WeierstrassB := ((2*a^3-9*a)/(27*b^3)).
+ Local Notation Wpoint := (@W.point F Feq Fadd Fmul WeierstrassA WeierstrassB).
+ Local Notation Wadd := (@W.add F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv field Feq_dec char_ge_3 WeierstrassA WeierstrassB).
+ Program Definition to_Weierstrass (P:@point) : Wpoint :=
+ match M.coordinates P return F*F+∞ with
+ | (x, y) => ((x + a/3)/b, y/b)
+ | _ => ∞
+ end.
+ Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.
+
+ Program Definition of_Weierstrass (P:Wpoint) : point :=
+ match W.coordinates P return F*F+∞ with
+ | (x,y) => (b*x-a/3, b*y)
+ | _ => ∞
+ end.
+ Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.
+ End MontgomeryWeierstrass.
+ End MontgomeryCurve.
+End M. \ No newline at end of file
diff --git a/src/Curves/Montgomery/AffineProofs.v b/src/Curves/Montgomery/AffineProofs.v
new file mode 100644
index 000000000..a83109a55
--- /dev/null
+++ b/src/Curves/Montgomery/AffineProofs.v
@@ -0,0 +1,83 @@
+Require Import Crypto.Algebra.Field.
+Require Import Crypto.Util.GlobalSettings.
+Require Import Crypto.Util.Sum Crypto.Util.Prod.
+Require Import Crypto.Util.Tactics.BreakMatch.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Spec.MontgomeryCurve Crypto.Curves.Montgomery.Affine.
+Require Import Crypto.Spec.WeierstrassCurve Crypto.Curves.Weierstrass.Affine.
+Require Import Crypto.Curves.Weierstrass.AffineProofs.
+
+Module M.
+ Section MontgomeryCurve.
+ Import BinNat.
+ Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {Feq_dec:Decidable.DecidableRel Feq}
+ {char_ge_28:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 28}.
+ Let char_ge_12 : @Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 12.
+ Proof. eapply Algebra.Hierarchy.char_ge_weaken; eauto. vm_decide. Qed.
+ Let char_ge_3 : @Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 3.
+ Proof. eapply Algebra.Hierarchy.char_ge_weaken; eauto; vm_decide. Qed.
+
+ Local Infix "=" := Feq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
+ Local Infix "+" := Fadd. Local Infix "*" := Fmul.
+ Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
+ Local Notation "- x" := (Fopp x).
+ Local Notation "x ^ 2" := (x*x) (at level 30).
+ Local Notation "x ^ 3" := (x*x^2) (at level 30).
+ Local Notation "0" := Fzero. Local Notation "1" := Fone.
+ Local Notation "2" := (1+1). Local Notation "3" := (1+2).
+ Local Notation "9" := (3*3). Local Notation "27" := (3*9).
+ Local Notation "'∞'" := unit : type_scope.
+ Local Notation "'∞'" := (inr tt) : core_scope.
+ Local Notation "( x , y )" := (inl (pair x y)).
+ Local Open Scope core_scope.
+
+ Context {a b: F} {b_nonzero:b <> 0}.
+
+ Local Notation WeierstrassA := ((3-a^2)/(3*b^2)).
+ Local Notation WeierstrassB := ((2*a^3-9*a)/(27*b^3)).
+ Local Notation Wpoint := (@W.point F Feq Fadd Fmul WeierstrassA WeierstrassB).
+ Local Notation Wadd := (@W.add F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv field Feq_dec char_ge_3 WeierstrassA WeierstrassB).
+ Local Notation Wopp := (@W.opp F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv WeierstrassA WeierstrassB field Feq_dec).
+
+ Ltac t :=
+ repeat
+ match goal with
+ | _ => solve [ trivial ]
+ | _ => progress intros
+ | _ => progress subst
+ | _ => progress Tactics.DestructHead.destruct_head' @M.point
+ | _ => progress Tactics.DestructHead.destruct_head' @prod
+ | _ => progress Tactics.DestructHead.destruct_head' @sum
+ | _ => progress Tactics.DestructHead.destruct_head' @and
+ | _ => progress Sum.inversion_sum
+ | _ => progress Prod.inversion_prod
+ | _ => progress Tactics.BreakMatch.break_match_hyps
+ | _ => progress Tactics.BreakMatch.break_match
+ | _ => progress cbv [M.coordinates M.add M.zero M.eq M.opp proj1_sig
+ W.coordinates W.add W.zero W.eq W.opp
+ M.of_Weierstrass M.to_Weierstrass] in *
+ | |- _ /\ _ => split | |- _ <-> _ => split
+ end.
+
+ Program Definition _MW (discr_nonzero:id _) : _ /\ _ /\ _ :=
+ @Group.group_from_redundant_representation
+ Wpoint W.eq Wadd W.zero Wopp
+ (Algebra.Hierarchy.abelian_group_group (W.commutative_group(char_ge_12:=char_ge_12)(discriminant_nonzero:=discr_nonzero)))
+ (@M.point F Feq Fadd Fmul a b) M.eq (M.add(char_ge_3:=char_ge_3)(b_nonzero:=b_nonzero)) M.zero (M.opp(b_nonzero:=b_nonzero))
+ (M.of_Weierstrass(b_nonzero:=b_nonzero))
+ (M.to_Weierstrass(b_nonzero:=b_nonzero))
+ _ _ _ _ _
+ .
+ Next Obligation. Proof. t; fsatz. Qed.
+ Next Obligation. Proof. t; fsatz. Qed.
+ Next Obligation. Proof. t; fsatz. Qed.
+ Next Obligation. Proof. t; fsatz. Qed.
+ Next Obligation. Proof. t; fsatz. Qed.
+
+ Global Instance group discr_nonzero : Algebra.Hierarchy.group := proj1 (_MW discr_nonzero).
+ Global Instance homomorphism_of_Weierstrass discr_nonzero : Monoid.is_homomorphism(phi:=M.of_Weierstrass) := proj1 (proj2 (_MW discr_nonzero)).
+ Global Instance homomorphism_to_Weierstrass discr_nonzero : Monoid.is_homomorphism(phi:=M.to_Weierstrass) := proj2 (proj2 (_MW discr_nonzero)).
+ End MontgomeryCurve.
+End M.
diff --git a/src/Curves/Montgomery/XZ.v b/src/Curves/Montgomery/XZ.v
new file mode 100644
index 000000000..60020827c
--- /dev/null
+++ b/src/Curves/Montgomery/XZ.v
@@ -0,0 +1,57 @@
+Require Import Crypto.Algebra.Field.
+Require Import Crypto.Util.GlobalSettings Crypto.Util.Notations.
+Require Import Crypto.Util.Sum Crypto.Util.Prod Crypto.Util.LetIn.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Spec.MontgomeryCurve Crypto.Curves.Montgomery.Affine.
+
+Module M.
+ Section MontgomeryCurve.
+ Import BinNat.
+ Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {Feq_dec:Decidable.DecidableRel Feq}
+ {char_ge_5:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 5}.
+ Local Infix "=" := Feq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
+ Local Infix "+" := Fadd. Local Infix "*" := Fmul.
+ Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
+ Local Notation "x ^ 2" := (x*x).
+ Local Notation "0" := Fzero. Local Notation "1" := Fone.
+ Local Notation "'∞'" := (inr tt) : core_scope.
+ Local Notation "( x , y )" := (inl (pair x y)).
+
+ Context {a b: F} {b_nonzero:b <> 0}.
+ Local Notation add := (M.add(b_nonzero:=b_nonzero)).
+ Local Notation opp := (M.opp(b_nonzero:=b_nonzero)).
+ Local Notation point := (@M.point F Feq Fadd Fmul a b).
+
+ Program Definition to_xz (P:point) : F*F :=
+ match M.coordinates P with
+ | (x, y) => pair x 1
+ | ∞ => pair 1 0
+ end.
+
+ Let char_ge_3:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul (BinNat.N.succ_pos (BinNat.N.two)).
+ Proof. eapply Algebra.Hierarchy.char_ge_weaken; eauto; vm_decide. Qed.
+
+ (* From Curve25519 paper by djb, appendix B. Credited to Montgomery *)
+ Context {a24:F} {a24_correct:(1+1+1+1)*a24 = a-(1+1)}.
+ Definition xzladderstep (x1:F) (Q Q':F*F) : ((F*F)*(F*F)) :=
+ match Q, Q' with
+ pair x z, pair x' z' =>
+ dlet A := x+z in
+ dlet B := x-z in
+ dlet AA := A^2 in
+ dlet BB := B^2 in
+ dlet x2 := AA*BB in
+ dlet E := AA-BB in
+ dlet z2 := E*(AA + a24*E) in
+ dlet C := x'+z' in
+ dlet D := x'-z' in
+ dlet CB := C*B in
+ dlet DA := D*A in
+ dlet x3 := (DA+CB)^2 in
+ dlet z3 := x1*(DA-CB)^2 in
+ (pair (pair x2 z2) (pair x3 z3))
+ end.
+ End MontgomeryCurve.
+End M.
diff --git a/src/Curves/Montgomery/XZProofs.v b/src/Curves/Montgomery/XZProofs.v
new file mode 100644
index 000000000..d24d1398c
--- /dev/null
+++ b/src/Curves/Montgomery/XZProofs.v
@@ -0,0 +1,58 @@
+Require Import Crypto.Algebra.Field.
+Require Import Crypto.Util.Sum Crypto.Util.Prod Crypto.Util.LetIn.
+Require Import Crypto.Util.Decidable.
+Require Import Crypto.Spec.MontgomeryCurve Crypto.Curves.Montgomery.Affine.
+Require Import Crypto.Curves.Montgomery.XZ BinPos.
+
+Module M.
+ Section MontgomeryCurve.
+ Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
+ {Feq_dec:Decidable.DecidableRel Feq}
+ {char_ge_5:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 5}.
+ Local Infix "=" := Feq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
+ Local Infix "+" := Fadd. Local Infix "*" := Fmul.
+ Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
+ Local Notation "0" := Fzero. Local Notation "1" := Fone.
+ Local Notation "'∞'" := (inr tt) : core_scope.
+ Local Notation "( x , y )" := (inl (pair x y)).
+
+ Let char_ge_3:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul (BinNat.N.succ_pos (BinNat.N.two)).
+ Proof. eapply Algebra.Hierarchy.char_ge_weaken; eauto; vm_decide. Qed.
+
+ Context {a b: F} {b_nonzero:b <> 0}.
+ Context {a24:F} {a24_correct:(1+1+1+1)*a24 = a-(1+1)}.
+ Local Notation add := (M.add(a:=a)(b_nonzero:=b_nonzero)(char_ge_3:=char_ge_3)).
+ Local Notation opp := (M.opp(a:=a)(b_nonzero:=b_nonzero)).
+ Local Notation point := (@M.point F Feq Fadd Fmul a b).
+ Local Notation xzladderstep := (M.xzladderstep(a24:=a24)(Fadd:=Fadd)(Fsub:=Fsub)(Fmul:=Fmul)).
+
+ Ltac t :=
+ repeat
+ match goal with
+ | _ => solve [ contradiction | trivial ]
+ | _ => progress intros
+ | _ => progress subst
+ | _ => progress Tactics.DestructHead.destruct_head' @M.point
+ | _ => progress Tactics.DestructHead.destruct_head' @prod
+ | _ => progress Tactics.DestructHead.destruct_head' @sum
+ | _ => progress Tactics.DestructHead.destruct_head' @and
+ | _ => progress Sum.inversion_sum
+ | _ => progress Prod.inversion_prod
+ | _ => progress Tactics.BreakMatch.break_match_hyps
+ | _ => progress Tactics.BreakMatch.break_match
+ | _ => progress cbv [fst snd M.coordinates M.add M.zero M.eq M.opp proj1_sig M.xzladderstep M.to_xz Let_In] in *
+ | |- _ /\ _ => split
+ end.
+
+ Lemma xzladderstep_correct
+ (Q Q':point) x z x' z' x1 x2 z2 x3 z3
+ (Hl:Logic.eq (pair(pair x2 z2)(pair x3 z3)) (xzladderstep x1 (pair x z) (pair x' z')))
+ (H:match M.coordinates Q with∞=>z=0/\x<>0|(xQ,y)=>xQ=x/z/\z<>0 (* TODO *) /\ y <> 0 (* TODO: prove this from non-squareness of a^2 - 4 *) end)
+ (H':match M.coordinates Q' with∞=>z'=0/\x'<>0|(xQ',_)=>xQ'=x'/z'/\z'<>0 end)
+ (H1:match M.coordinates (add Q (opp Q')) with∞=>False|(x,y)=>x=x1/\x<>0 end):
+ match M.coordinates (add Q Q) with∞=>z2=0/\x2<>0|(xQQ,_)=>xQQ=x2/z2/\z2<>0 end /\
+ match M.coordinates (add Q Q') with∞=>z3=0/\x3<>0|(xQQ',_)=>xQQ'=x3/z3/\z3<>0 end.
+ Proof using a24_correct char_ge_5. t; abstract fsatz. Qed.
+ End MontgomeryCurve.
+End M.