Skip to content

Commit 841cc2f

Browse files
decide point equality (#1848)
1 parent eb5ba09 commit 841cc2f

File tree

5 files changed

+36
-14
lines changed

5 files changed

+36
-14
lines changed

src/Curves/Edwards/AffineProofs.v

+2
Original file line numberDiff line numberDiff line change
@@ -46,6 +46,8 @@ Module E.
4646
Program Definition opp (P:point) : point := (Fopp (fst P), (snd P)).
4747
Next Obligation. match goal with P : point |- _ => destruct P as [ [??]?] end; cbv; fsatz. Qed.
4848

49+
Global Instance Decidable_eq : Decidable.DecidableRel (@E.eq _ Feq Fone Fadd Fmul a d) := _.
50+
4951
Ltac t_step :=
5052
match goal with
5153
| _ => solve [trivial | exact _ ]

src/Curves/EdwardsMontgomery25519.v

+21-13
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,5 @@
11
Require Import Coq.ZArith.ZArith. Local Open Scope Z_scope.
2+
Require Import Crypto.Util.Decidable.
23
Require Import Crypto.Spec.ModularArithmetic. Local Open Scope F_scope.
34
Require Import Crypto.Curves.EdwardsMontgomery. Import M.
45
Require Import Crypto.Curves.Edwards.TwistIsomorphism.
@@ -11,19 +12,26 @@ Import MontgomeryCurve CompleteEdwardsCurve.
1112

1213
Local Definition a' := (M.a + (1 + 1)) / M.b.
1314
Local Definition d' := (M.a - (1 + 1)) / M.b.
14-
Definition r := sqrt (F.inv ((a' / M.b) / E.a)).
15+
Local Definition r := sqrt (F.inv ((a' / M.b) / E.a)).
1516

16-
Local Lemma is_twist : E.a * d' = a' * E.d. Proof. Decidable.vm_decide. Qed.
17-
Local Lemma nonzero_a' : a' <> 0. Proof. Decidable.vm_decide. Qed.
18-
Local Lemma r_correct : E.a = r * r * a'. Proof. Decidable.vm_decide. Qed.
17+
Local Lemma is_twist : E.a * d' = a' * E.d. Proof. vm_decide. Qed.
18+
Local Lemma nonzero_a' : a' <> 0. Proof. vm_decide. Qed.
19+
Local Lemma r_correct : E.a = r * r * a'. Proof. vm_decide. Qed.
1920

20-
Definition Montgomery_of_Edwards (P : Curve25519.E.point) : Curve25519.M.point :=
21-
@of_Edwards _ _ _ _ _ _ _ _ _ _ field _ char_ge_3 M.a M.b M.b_nonzero a' d' eq_refl eq_refl nonzero_a'
22-
(@E.point2_of_point1 _ _ _ _ _ _ _ _ _ _ field _ E.a E.d a' d' is_twist E.nonzero_a nonzero_a' r r_correct P).
23-
24-
Definition Edwards_of_Montgomery (P : Curve25519.M.point) : Curve25519.E.point :=
21+
Module E.
22+
Definition of_Montgomery (P : Curve25519.M.point) : Curve25519.E.point :=
2523
@E.point1_of_point2 _ _ _ _ _ _ _ _ _ _ field _ E.a E.d a' d' is_twist E.nonzero_a nonzero_a' r r_correct
2624
(@to_Edwards _ _ _ _ _ _ _ _ _ _ field _ M.a M.b M.b_nonzero a' d' eq_refl eq_refl nonzero_a' P).
25+
Lemma of_Montgomery_B : E.eq E.B (of_Montgomery M.B). Proof. vm_decide. Qed.
26+
End E.
27+
28+
Module M.
29+
Definition of_Edwards (P : Curve25519.E.point) : Curve25519.M.point :=
30+
@of_Edwards _ _ _ _ _ _ _ _ _ _ field _ char_ge_3 M.a M.b M.b_nonzero a' d' eq_refl eq_refl nonzero_a'
31+
(@E.point2_of_point1 _ _ _ _ _ _ _ _ _ _ field _ E.a E.d a' d' is_twist E.nonzero_a nonzero_a' r r_correct P).
32+
Lemma of_Edwards_B : M.eq M.B (of_Edwards E.B). Proof.
33+
Proof. simple notypeclasses refine (@dec_bool _ _ _). apply Affine.M.Decidable_eq. vm_decide. Qed.
34+
End M.
2735

2836
Local Notation Eopp := ((@AffineProofs.E.opp _ _ _ _ _ _ _ _ _ _ field _ E.a E.d E.nonzero_a)).
2937

@@ -36,16 +44,16 @@ Local Arguments of_Edwards {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _} _ _ { _ _ _ }.
3644
Lemma EdwardsMontgomery25519 : @Group.isomorphic_commutative_groups
3745
Curve25519.E.point E.eq Curve25519.E.add Curve25519.E.zero Eopp Curve25519.M.point
3846
M.eq Curve25519.M.add M.zero Curve25519.M.opp
39-
Montgomery_of_Edwards Edwards_of_Montgomery.
47+
M.of_Edwards E.of_Montgomery.
4048
Proof.
41-
cbv [Montgomery_of_Edwards Edwards_of_Montgomery].
49+
cbv [M.of_Edwards E.of_Montgomery].
4250
epose proof E.twist_isomorphism(a1:=E.a)(a2:=a')(d1:=E.d)(d2:=d')(r:=r) as AB.
4351
epose proof EdwardsMontgomeryIsomorphism(a:=Curve25519.M.a)(b:=Curve25519.M.b)as BC.
4452
destruct AB as [A B ab ba], BC as [_ C bc cb].
4553
pose proof Group.compose_homomorphism(homom:=ab)(homom2:=bc) as ac.
4654
pose proof Group.compose_homomorphism(homom:=cb)(homom2:=ba)(groupH2:=ltac:(eapply A)) as ca.
4755
split; try exact ac; try exact ca; try exact A; try exact C.
4856
Unshelve.
49-
all : try (pose (@PrimeFieldTheorems.F.Decidable_square p prime_p eq_refl); Decidable.vm_decide).
50-
all : try (eapply Hierarchy.char_ge_weaken; try apply ModularArithmeticTheorems.F.char_gt; Decidable.vm_decide).
57+
all : try (pose (@PrimeFieldTheorems.F.Decidable_square p prime_p eq_refl); vm_decide).
58+
all : try (eapply Hierarchy.char_ge_weaken; try apply ModularArithmeticTheorems.F.char_gt; vm_decide).
5159
Qed.

src/Curves/Montgomery/Affine.v

+2
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,8 @@ Module M.
3131
Local Notation add := (M.add(b_nonzero:=b_nonzero)).
3232
Local Notation point := (@M.point F Feq Fadd Fmul a b).
3333

34+
Global Instance Decidable_eq : Decidable.DecidableRel (@M.eq _ Feq Fadd Fmul a b) := _.
35+
3436
Section MontgomeryWeierstrass.
3537
Local Notation "2" := (1+1).
3638
Local Notation "3" := (1+2).

src/Spec/Curve25519.v

+1-1
Original file line numberDiff line numberDiff line change
@@ -97,7 +97,7 @@ Module E.
9797
refine (
9898
exist _ (F.of_Z _ 15112221349535400772501151409588531511454012693041857206046113283949847762202, F.div (F.of_Z _ 4) (F.of_Z _ 5)) _).
9999
Decidable.vm_decide.
100-
Qed.
100+
Defined.
101101
End E.
102102

103103
Require Import Spec.MontgomeryCurve.

src/Util/Decidable.v

+10
Original file line numberDiff line numberDiff line change
@@ -136,12 +136,22 @@ Global Instance dec_match_pair {A B} {P : A -> B -> Prop} {x : A * B}
136136
: Decidable (let '(a, b) := x in P a b) | 1.
137137
Proof. edestruct (_ : _ * _); assumption. Defined.
138138

139+
Global Instance dec_match_sum {A B} {P : A -> Prop} {Q : B -> Prop}
140+
{HP : forall x, Decidable (P x)} {HQ : forall x, Decidable (Q x)}
141+
{x : A + B}
142+
: Decidable (match x with inl a => P a | inr b => Q b end) | 1.
143+
Proof. edestruct (_ : _ + _); eauto. Defined.
144+
139145
Global Instance dec_if_bool {b : bool} {A B} {HA : Decidable A} {HB : Decidable B}
140146
: Decidable (if b then A else B) | 10
141147
:= if b return Decidable (if b then A else B)
142148
then HA
143149
else HB.
144150

151+
Global Instance dec_match_unit {P} {HP : Decidable P} {x : unit}
152+
: Decidable (match x with tt => P end) | 1.
153+
Proof. edestruct x. assumption. Defined.
154+
145155
Lemma not_not P {d:Decidable P} : not (not P) <-> P.
146156
Proof. destruct (dec P); intuition. Qed.
147157

0 commit comments

Comments
 (0)