aboutsummaryrefslogtreecommitdiff
path: root/src/CompleteEdwardsCurve
diff options
context:
space:
mode:
authorGravatar Andres Erbsen <andreser@mit.edu>2016-02-15 12:35:37 -0500
committerGravatar Andres Erbsen <andreser@mit.edu>2016-02-15 12:35:37 -0500
commitd2b15c7a598a13fbc6c014db0775001f178d359c (patch)
tree84d5a50bcc864399593ecd37ec9f1084d63dba9c /src/CompleteEdwardsCurve
parent094ccf074fc64cc8256278d26cca46107b9cc813 (diff)
CompleteEdwardsCurve: unifiedAddM1: Closed Under Global Context
Diffstat (limited to 'src/CompleteEdwardsCurve')
-rw-r--r--src/CompleteEdwardsCurve/ExtendedCoordinates.v145
-rw-r--r--src/CompleteEdwardsCurve/ExtendedCoordinatesIssue.v38
2 files changed, 99 insertions, 84 deletions
diff --git a/src/CompleteEdwardsCurve/ExtendedCoordinates.v b/src/CompleteEdwardsCurve/ExtendedCoordinates.v
index d2cd753fe..f2eba81eb 100644
--- a/src/CompleteEdwardsCurve/ExtendedCoordinates.v
+++ b/src/CompleteEdwardsCurve/ExtendedCoordinates.v
@@ -2,20 +2,93 @@ Require Import Crypto.CompleteEdwardsCurve.Pre.
Require Import Crypto.CompleteEdwardsCurve.CompleteEdwardsCurveTheorems.
Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
Require Import Crypto.Tactics.VerdiTactics.
+Local Open Scope F_scope.
-Section ExtendedCoordinates.
- Local Open Scope F_scope.
- Context {prm:TwistedEdwardsParams}.
- Existing Instance prime_q.
+Section ExtendedCoordinatesFieldProofs.
+ (* If [field] worked on (F q) when Definition q := someProj someRecord, we
+ could inline this proof into unifiedAdd_repM1 *)
+ Context p (prime_p : Znumtheory.prime p) (two_lt_p : BinInt.Z.lt 2 p).
+ Existing Instance prime_p.
- Add Field GFfield_Z : (@Ffield_theory q _)
- (morphism (@Fring_morph q),
+ Add Field GFfield_Z : (@Ffield_theory p _)
+ (morphism (@Fring_morph p),
preprocess [Fpreprocess],
postprocess [Fpostprocess],
constants [Fconstant],
- div (@Fmorph_div_theory q),
- power_tac (@Fpower_theory q) [Fexp_tac]).
+ div (@Fmorph_div_theory p),
+ power_tac (@Fpower_theory p) [Fexp_tac]).
+
+ Lemma unifiedAdd_repM1_fieldproof:
+ forall (d XP YP ZP XQ YQ ZQ : F p)
+ (HZP : ZP <> 0)
+ (HZQ : ZQ <> 0)
+ (HoRp : forall x1 y1 x2 y2 : F p,
+ opp 1 * x1 ^ 2 + y1 ^ 2 = 1 + d * x1 ^ 2 * y1 ^ 2 ->
+ opp 1 * x2 ^ 2 + y2 ^ 2 = 1 + d * x2 ^ 2 * y2 ^ 2 ->
+ 1 + d * x1 * x2 * y1 * y2 <> 0)
+ (HoRm : forall x1 y1 x2 y2 : F p,
+ opp 1 * x1 ^ 2 + y1 ^ 2 = 1 + d * x1 ^ 2 * y1 ^ 2 ->
+ opp 1 * x2 ^ 2 + y2 ^ 2 = 1 + d * x2 ^ 2 * y2 ^ 2 ->
+ 1 - d * x1 * x2 * y1 * y2 <> 0)
+ (HoQ: opp 1 * (XQ / ZQ) ^ 2 + (YQ / ZQ) ^ 2 = 1 + d * (XQ / ZQ) ^ 2 * (YQ / ZQ) ^ 2)
+ (HoP : opp 1 * (XP / ZP) ^ 2 + (YP / ZP) ^ 2 = 1 + d * (XP / ZP) ^ 2 * (YP / ZP) ^ 2),
+ (((YP + XP) * (YQ + XQ) - (YP - XP) * (YQ - XQ)) *
+ (ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) /
+ ((ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ))),
+ (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ ((YP + XP) * (YQ + XQ) + (YP - XP) * (YQ - XQ)) /
+ ((ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)))) =
+ ((XP / ZP * (YQ / ZQ) + YP / ZP * (XQ / ZQ)) /
+ (1 + d * (XP / ZP) * (XQ / ZQ) * (YP / ZP) * (YQ / ZQ)),
+ (YP / ZP * (YQ / ZQ) - opp 1 * (XP / ZP) * (XQ / ZQ)) /
+ (1 - d * (XP / ZP) * (XQ / ZQ) * (YP / ZP) * (YQ / ZQ))) /\
+ ((ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) = 0 ->
+ False) /\
+ ((YP + XP) * (YQ + XQ) - (YP - XP) * (YQ - XQ)) *
+ ((YP + XP) * (YQ + XQ) + (YP - XP) * (YQ - XQ)) =
+ ((YP + XP) * (YQ + XQ) - (YP - XP) * (YQ - XQ)) *
+ (ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ ((ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ ((YP + XP) * (YQ + XQ) + (YP - XP) * (YQ - XQ))) /
+ ((ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
+ (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ))).
+ Proof.
+ intros; repeat split; try apply (f_equal2 pair); try field; auto.
+ Ltac tnz := eauto 10 using Fq_mul_nonzero_nonzero, (@char_gt_2 p two_lt_p).
+ (* If we we had reasoning modulo associativity and commutativity,
+ * the following tactic would probably solve all 10 goals here:
+ repeat match goal with [H1: @eq (F p) _ _, H2: @eq (F p) _ _ |- _ ] =>
+ let H := fresh "H" in (
+ pose proof (edwardsAddCompletePlus _ _ _ _ H1 H2) as H;
+ match type of H with ?xs <> 0 => ac_rewrite (eq_refl xs) end
+ ) || (
+ pose proof (edwardsAddCompleteMinus _ _ _ _ H1 H2) as H;
+ match type of H with ?xs <> 0 => ac_rewrite (eq_refl xs) end
+ ); tnz
+ end. *)
+
+ - replace (ZP * ZQ * ZP * ZQ + d * XP * XQ * YP * YQ) with (ZQ*ZQ*ZP*ZP* (1 + d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) + XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 + d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) - XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZQ * ZP * ZQ - d * XP * XQ * YP * YQ) with (ZQ*ZQ*ZP*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) + XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 + d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) - XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZQ * ZP * ZQ - d * XP * XQ * YP * YQ) with (ZQ*ZQ*ZP*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz).
+ repeat apply Fq_mul_nonzero_nonzero.
+ + replace (ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ) ) with (ZToField 2*ZQ*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ + replace (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ) ) with (ZToField 2*ZQ*ZP* (1 + d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) + XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 + d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ - replace (ZP * ZToField 2 * ZQ * (ZP * ZQ) - XP * YP * ZToField 2 * d * (XQ * YQ)) with (ZToField 2*ZQ*ZQ*ZP*ZP* (1 - d * (XQ / ZQ) * (XP / ZP) * (YQ / ZQ) * (YP / ZP))) by (field; tnz); tnz.
+ Qed.
+End ExtendedCoordinatesFieldProofs.
+
+Section ExtendedCoordinates.
+ Context {prm:TwistedEdwardsParams}.
+ Existing Instance prime_q.
(** [extended] represents a point on an elliptic curve using extended projective
* Edwards coordinates with twist a=-1 (see <https://eprint.iacr.org/2008/522.pdf>). *)
@@ -37,8 +110,8 @@ Section ExtendedCoordinates.
Local Hint Unfold twistedToExtended extendedToTwisted rep.
Local Notation "P '~=' rP" := (rep P rP) (at level 70).
- Ltac extended_tac :=
- repeat progress (autounfold; unfold onCurve, unifiedAdd, unifiedAdd', rep in *; intros);
+ Ltac unfoldExtended :=
+ repeat progress (autounfold; unfold onCurve, unifiedAdd, unifiedAdd', rep in *; intros);
repeat match goal with
| [ p : (F q*F q)%type |- _ ] =>
let x := fresh "x" p in
@@ -51,25 +124,29 @@ Section ExtendedCoordinates.
let T := fresh "T" p in
destruct p as [X Y Z T]
| [ H: _ /\ _ |- _ ] => destruct H
- | [ |- _ /\ _ ] => split
| [ H: @eq (F q * F q)%type _ _ |- _ ] => invcs H
| [ H: @eq F q ?x _ |- _ ] => isVar x; rewrite H; clear H
+ end.
+
+ Ltac solveExtended := unfoldExtended;
+ repeat match goal with
+ | [ |- _ /\ _ ] => split
| [ |- @eq (F q * F q)%type _ _] => apply f_equal2
- | _ => progress rewrite ?F_add_0_r, ?F_add_0_l, ?F_sub_0_l, ?F_sub_0_r,
- ?F_mul_0_r, ?F_mul_0_l, ?F_mul_1_l, ?F_mul_1_r, ?F_div_1_r
- | _ => progress eauto using Fq_1_neq_0
+ | _ => progress rewrite ?@F_add_0_r, ?@F_add_0_l, ?@F_sub_0_l, ?@F_sub_0_r,
+ ?@F_mul_0_r, ?@F_mul_0_l, ?@F_mul_1_l, ?@F_mul_1_r, ?@F_div_1_r
+ | _ => solve [eapply @Fq_1_neq_0; eauto with typeclass_instances]
+ | _ => solve [eauto with typeclass_instances]
| [ H: a = _ |- _ ] => rewrite H
- | [ |- @eq (F q) _ _] => fail; field (*FIXME*)
end.
Lemma twistedToExtended_rep : forall P, twistedToExtended P ~= P.
Proof.
- extended_tac.
+ solveExtended.
Qed.
Lemma extendedToTwisted_rep : forall P rP, P ~= rP -> extendedToTwisted P = rP.
Proof.
- extended_tac.
+ solveExtended.
Qed.
Section TwistMinus1.
@@ -97,34 +174,10 @@ Section ExtendedCoordinates.
P ~= rP -> Q ~= rQ -> (unifiedAddM1 P Q) ~= (unifiedAdd' rP rQ).
Proof.
intros P Q rP rQ HoP HoQ HrP HrQ.
- pose proof (@edwardsAddCompletePlus _ _ _ _ two_lt_q nonzero_a square_a nonsquare_d) as HoR; simpl in HoR.
- extended_tac.
- field.
-
- (* If we we had reasoning modulo associativity and commutativity,
- * the following tactic would probably solve all 10 goals here:
- repeat match goal with [H1: @eq GF _ _, H2: @eq GF _ _ |- _ ] =>
- let H := fresh "H" in (
- pose proof (edwardsAddCompletePlus _ _ _ _ H1 H2) as H;
- match type of H with ?xs <> 0 => ac_rewrite (eq_refl xs) end
- ) || (
- pose proof (edwardsAddCompleteMinus _ _ _ _ H1 H2) as H;
- match type of H with ?xs <> 0 => ac_rewrite (eq_refl xs) end
- ); repeat apply mul_nonzero_nonzero; auto 10
- end.
- *)
-
- - replace (Z0 * Z * Z0 * Z + d * X0 * X * Y0 * Y) with (Z*Z*Z0*Z0* (1 + d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) + X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 + d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) - X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * Z * Z0 * Z - d * X0 * X * Y0 * Y) with (Z*Z*Z0*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) + X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 + d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) - X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * Z * Z0 * Z - d * X0 * X * Y0 * Y) with (Z*Z*Z0*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto).
- repeat apply mul_nonzero_nonzero.
- + replace (Z0 * 2 * Z - X0 * Y0 / Z0 * 2 * d * (X * Y / Z)) with (2*Z*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- + replace (Z0 * 2 * Z + X0 * Y0 / Z0 * 2 * d * (X * Y / Z)) with (2*Z*Z0* (1 + d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) + X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 + d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
- - replace (Z0 * inject 2 * Z * (Z0 * Z) - X0 * Y0 * inject 2 * d * (X * Y)) with (2*Z*Z*Z0*Z0* (1 - d * (X / Z) * (X0 / Z0) * (Y / Z) * (Y0 / Z0))) by (field; auto); auto 10.
+ pose proof (@edwardsAddCompletePlus _ _ _ _ two_lt_q nonzero_a square_a nonsquare_d) as HoRp; simpl in HoRp.
+ pose proof (@edwardsAddCompleteMinus _ _ _ _ two_lt_q nonzero_a square_a nonsquare_d) as HoRm; simpl in HoRm.
+ unfoldExtended; rewrite a_eq_minus1 in *.
+ apply unifiedAdd_repM1_fieldproof; auto using prime_q, two_lt_q.
Qed.
+ End TwistMinus1.
End ExtendedCoordinates. \ No newline at end of file
diff --git a/src/CompleteEdwardsCurve/ExtendedCoordinatesIssue.v b/src/CompleteEdwardsCurve/ExtendedCoordinatesIssue.v
deleted file mode 100644
index 8d24f2cd2..000000000
--- a/src/CompleteEdwardsCurve/ExtendedCoordinatesIssue.v
+++ /dev/null
@@ -1,38 +0,0 @@
-Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
-Require Import Crypto.Spec.CompleteEdwardsCurve.
-
-Section ExtendedCoordinates.
- Local Open Scope F_scope.
-
- Context {prm:TwistedEdwardsParams}.
- Context {p:BinInt.Z} {p_eq_q:p = q}.
- Lemma prime_p : Znumtheory.prime p.
- rewrite p_eq_q; exact prime_q.
- Qed.
-
- Add Field Ffield_Z : (@Ffield_theory p prime_p)
- (morphism (@Fring_morph p),
- preprocess [Fpreprocess],
- postprocess [Fpostprocess],
- constants [Fconstant],
- div (@Fmorph_div_theory p),
- power_tac (@Fpower_theory p) [Fexp_tac]).
-
- Lemma biggerFraction : forall XP YP ZP XQ YQ ZQ d : F q,
- ZQ <> 0 ->
- ZP <> 0 ->
- ZP * ZQ * ZP * ZQ + d * XP * XQ * YP * YQ <> 0 ->
- ZP * ZToField 2 * ZQ * (ZP * ZQ) + XP * YP * ZToField 2 * d * (XQ * YQ) <> 0 ->
- ZP * ZToField 2 * ZQ * (ZP * ZQ) - XP * YP * ZToField 2 * d * (XQ * YQ) <> 0 ->
-
- ((YP + XP) * (YQ + XQ) - (YP - XP) * (YQ - XQ)) *
- (ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) /
- ((ZP * ZToField 2 * ZQ - XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ)) *
- (ZP * ZToField 2 * ZQ + XP * YP / ZP * ZToField 2 * d * (XQ * YQ / ZQ))) =
- (XP / ZP * (YQ / ZQ) + YP / ZP * (XQ / ZQ)) / (1 + d * (XP / ZP) * (XQ / ZQ) * (YP / ZP) * (YQ / ZQ)).
- Proof.
- rewrite <-p_eq_q.
- intros.
- abstract (field; assumption).
- Qed.
-End ExtendedCoordinates. \ No newline at end of file