aboutsummaryrefslogtreecommitdiff
path: root/src
diff options
context:
space:
mode:
authorGravatar Andres Erbsen <andreser@mit.edu>2017-02-06 18:25:59 -0500
committerGravatar Andres Erbsen <andreser@mit.edu>2017-03-02 13:37:14 -0500
commit55de48a40d269bd900ce0af04449cb33ebb8577d (patch)
tree4857924d17d8d767d7144b2f6a7462d5d639c307 /src
parentd641525010b941607605b7e998b53ba6d99783e5 (diff)
use field_nsatz in CompleteEdwardsCurve.Pre
Diffstat (limited to 'src')
-rw-r--r--src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v4
-rw-r--r--src/CompleteEdwardsCurve/Pre.v129
-rw-r--r--src/Spec/CompleteEdwardsCurve.v18
3 files changed, 36 insertions, 115 deletions
diff --git a/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v b/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
index e90aad0d8..c8986cee9 100644
--- a/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
+++ b/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
@@ -22,7 +22,7 @@ Module E.
Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
Local Notation "x ^ 2" := (x*x).
Local Notation point := (@point F Feq Fone Fadd Fmul a d).
- Local Notation onCurve := (@onCurve F Feq Fone Fadd Fmul a d).
+ Local Notation onCurve x y := (a*x^2 + y^2 = 1 + d*x^2*y^2).
Add Field _edwards_curve_theorems_field : (field_theory_for_stdlib_tactic (H:=field)).
@@ -119,7 +119,7 @@ Module E.
destruct (dec (y=0)); [apply nonzero_a | apply nonsquare_d with (sqrt_a/y)]; super_nsatz.
Qed.
- Lemma solve_correct : forall x y, onCurve (x, y) <-> (x^2 = solve_for_x2 y).
+ Lemma solve_correct : forall x y, onCurve x y <-> (x^2 = solve_for_x2 y).
Proof.
unfold solve_for_x2; simpl; split; intros;
(common_denominator_all; [nsatz | auto using a_d_y2_nonzero]).
diff --git a/src/CompleteEdwardsCurve/Pre.v b/src/CompleteEdwardsCurve/Pre.v
index fe98f5dcc..aa10fbd5f 100644
--- a/src/CompleteEdwardsCurve/Pre.v
+++ b/src/CompleteEdwardsCurve/Pre.v
@@ -1,11 +1,12 @@
Require Import Coq.Classes.Morphisms. Require Coq.Setoids.Setoid.
Require Import Crypto.Algebra Crypto.Algebra.
-Require Import Crypto.Util.Notations Crypto.Util.Decidable.
+Require Import Crypto.Util.Notations Crypto.Util.Decidable Crypto.Util.Tactics.
-Generalizable All Variables.
Section Pre.
Context {F eq zero one opp add sub mul inv div}
- `{field F eq zero one opp add sub mul inv div} {eq_dec:DecidableRel eq}.
+ {field:@Algebra.field F eq zero one opp add sub mul inv div}
+ {eq_dec:DecidableRel eq}.
+ Add Field EdwardsCurveField : (Field.field_theory_for_stdlib_tactic (T:=F)).
Local Infix "=" := eq. Local Notation "a <> b" := (not (a = b)).
Local Infix "=" := eq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
@@ -14,102 +15,26 @@ Section Pre.
Local Infix "-" := sub. Local Infix "/" := div.
Local Notation "x ^ 2" := (x*x).
- Add Field EdwardsCurveField : (Field.field_theory_for_stdlib_tactic (T:=F)).
-
- Context {a:F} {a_nonzero : a<>0} {a_square : exists sqrt_a, sqrt_a^2 = a}.
- Context {d:F} {d_nonsquare : forall sqrt_d, sqrt_d^2 <> d}.
- Context {char_gt_2 : 1+1 <> 0}.
-
- (* the canonical definitions are in Spec *)
- Definition onCurve (P:F*F) := let (x, y) := P in a*x^2 + y^2 = 1 + d*x^2*y^2.
- Definition unifiedAdd' (P1' P2':F*F) : F*F :=
- let (x1, y1) := P1' in
- let (x2, y2) := P2' in
- pair (((x1*y2 + y1*x2)/(1 + d*x1*x2*y1*y2))) (((y1*y2 - a*x1*x2)/(1 - d*x1*x2*y1*y2))).
-
- Ltac use_sqrt_a := destruct a_square as [sqrt_a a_square']; rewrite <-a_square' in *.
-
- Lemma onCurve_subst : forall x1 x2 y1 y2, and (eq x1 y1) (eq x2 y2) -> onCurve (x1, x2) ->
- onCurve (y1, y2).
- Proof.
- unfold onCurve.
- intros ? ? ? ? [eq_1 eq_2] ?.
- rewrite eq_1, eq_2 in *.
- assumption.
- Qed.
-
- Lemma edwardsAddComplete' x1 y1 x2 y2 :
- onCurve (pair x1 y1) ->
- onCurve (pair x2 y2) ->
- (d*x1*x2*y1*y2)^2 <> 1.
- Proof.
- unfold onCurve, not; use_sqrt_a; intros.
- destruct (dec ((sqrt_a*x2 + y2) = 0)); destruct (dec ((sqrt_a*x2 - y2) = 0));
- lazymatch goal with
- | [H: not (eq (?f (sqrt_a * x2) y2) 0) |- _ ]
- => apply d_nonsquare with (sqrt_d:= (f (sqrt_a * x1) (d * x1 * x2 * y1 * y2 * y1))
- /(f (sqrt_a * x2) y2 * x1 * y1 ))
- | _ => apply a_nonzero
- end; super_nsatz.
- Qed.
-
- Lemma edwardsAddCompletePlus x1 y1 x2 y2 :
- onCurve (x1, y1) -> onCurve (x2, y2) -> (1 + d*x1*x2*y1*y2) <> 0.
- Proof. intros H1 H2 ?. apply (edwardsAddComplete' _ _ _ _ H1 H2); super_nsatz. Qed.
-
- Lemma edwardsAddCompleteMinus x1 y1 x2 y2 :
- onCurve (x1, y1) -> onCurve (x2, y2) -> (1 - d*x1*x2*y1*y2) <> 0.
- Proof. intros H1 H2 ?. apply (edwardsAddComplete' _ _ _ _ H1 H2); super_nsatz. Qed.
-
- Lemma zeroOnCurve : onCurve (0, 1). Proof. simpl. super_nsatz. Qed.
-
- Lemma unifiedAdd'_onCurve : forall P1 P2,
- onCurve P1 -> onCurve P2 -> onCurve (unifiedAdd' P1 P2).
- Proof.
- unfold onCurve, unifiedAdd'; intros [x1 y1] [x2 y2] ? ?.
- common_denominator; [ | auto using edwardsAddCompleteMinus, edwardsAddCompletePlus..].
- nsatz.
- Qed.
-End Pre.
-
-Import Group Ring Field.
-
-(* TODO: move -- this does not need to be defined before [point] *)
-Section RespectsFieldHomomorphism.
- Context {F EQ ZERO ONE OPP ADD MUL SUB INV DIV} `{@field F EQ ZERO ONE OPP ADD SUB MUL INV DIV}.
- Context {K eq zero one opp add mul sub inv div} `{@field K eq zero one opp add sub mul inv div}.
- Local Infix "=" := eq. Local Infix "=" := eq : type_scope.
- Context {phi:F->K} `{@is_homomorphism F EQ ONE ADD MUL K eq one add mul phi}.
- Context {A D:F} {a d:K} {a_ok:phi A=a} {d_ok:phi D=d}.
-
- Let phip := fun (P':F*F) => let (x, y) := P' in (phi x, phi y).
-
- Let eqp := fun (P1' P2':K*K) =>
- let (x1, y1) := P1' in
- let (x2, y2) := P2' in
- and (eq x1 x2) (eq y1 y2).
-
- Create HintDb field_homomorphism discriminated.
- Hint Rewrite
- homomorphism_one
- homomorphism_add
- homomorphism_sub
- homomorphism_mul
- homomorphism_div
- a_ok
- d_ok
- : field_homomorphism.
-
- Lemma morphism_unifiedAdd' : forall P Q:F*F,
- (~ EQ (ADD ONE (MUL (MUL (MUL (MUL D (fst P)) (fst Q)) (snd P)) (snd Q))) ZERO) ->
- (~ EQ (SUB ONE (MUL (MUL (MUL (MUL D (fst P)) (fst Q)) (snd P)) (snd Q))) ZERO) ->
- eqp
- (phip (unifiedAdd'(F:=F)(one:=ONE)(add:=ADD)(sub:=SUB)(mul:=MUL)(div:=DIV)(a:=A)(d:=D) P Q))
- (unifiedAdd'(F:=K)(one:=one)(add:=add)(sub:=sub)(mul:=mul)(div:=div)(a:=a)(d:=d) (phip P) (phip Q)).
- Proof.
- intros [x1 y1] [x2 y2] Hnonzero1 Hnonzero2;
- cbv [unifiedAdd' phip eqp];
- apply conj;
- (rewrite_strat topdown hints field_homomorphism); try assumption; reflexivity.
- Qed.
-End RespectsFieldHomomorphism.
+ Context (a:F) (a_nonzero : a<>0) (a_square : exists sqrt_a, sqrt_a^2 = a).
+ Context (d:F) (d_nonsquare : forall sqrt_d, sqrt_d^2 <> d).
+ Context (char_gt_2 : 1+1 <> 0).
+
+ Local Notation onCurve x y := (a*x^2 + y^2 = 1 + d*x^2*y^2).
+ Lemma zeroOnCurve : onCurve 0 1. nsatz. Qed.
+
+ Section Addition.
+ Context (x1 y1:F) (P1onCurve: onCurve x1 y1).
+ Context (x2 y2:F) (P2onCurve: onCurve x2 y2).
+ Lemma denominator_nonzero : (d*x1*x2*y1*y2)^2 <> 1.
+ Proof.
+ destruct a_square as [sqrt_a], (dec(sqrt_a*x2+y2 = 0)), (dec(sqrt_a*x2-y2 = 0));
+ try match goal with [H: ?f (sqrt_a * x2) y2 <> 0 |- _ ]
+ => specialize (d_nonsquare ((f (sqrt_a * x1) (d * x1 * x2 * y1 * y2 * y1))
+ /(f (sqrt_a * x2) y2 * x1 * y1 )))
+ end; field_nsatz.
+ Qed.
+
+ Lemma add_onCurve : onCurve ((x1*y2 + y1*x2)/(1 + d*x1*x2*y1*y2)) ((y1*y2 - a*x1*x2)/(1 - d*x1*x2*y1*y2)).
+ Proof. pose proof denominator_nonzero; field_nsatz. Qed.
+ End Addition.
+End Pre. \ No newline at end of file
diff --git a/src/Spec/CompleteEdwardsCurve.v b/src/Spec/CompleteEdwardsCurve.v
index 584013015..e081ded1e 100644
--- a/src/Spec/CompleteEdwardsCurve.v
+++ b/src/Spec/CompleteEdwardsCurve.v
@@ -24,23 +24,19 @@ Module E.
square_a : exists sqrt_a, sqrt_a^2 = a;
nonsquare_d : forall x, x^2 <> d
}.
- Context `{twisted_edwards_params}.
+ Context `{twisted_edwards_params}. (* TODO: name me *)
Definition point := { P | let '(x,y) := P in a*x^2 + y^2 = 1 + d*x^2*y^2 }.
Definition coordinates (P:point) : (F*F) := proj1_sig P.
- (** The following points are indeed on the curve -- see [CompleteEdwardsCurve.Pre] for proof *)
- Local Obligation Tactic := intros;
- apply (Pre.zeroOnCurve(a_nonzero:=nonzero_a)(char_gt_2:=char_gt_2)) ||
- apply (Pre.unifiedAdd'_onCurve (char_gt_2:=char_gt_2) (d_nonsquare:=nonsquare_d)
- (a_nonzero:=nonzero_a) (a_square:=square_a) _ _ (proj2_sig _) (proj2_sig _)).
-
Program Definition zero : point := (0, 1).
+ Next Obligation. auto using Pre.zeroOnCurve. Defined.
- Program Definition add (P1 P2:point) : point := exist _ (
- let (x1, y1) := coordinates P1 in
- let (x2, y2) := coordinates P2 in
- (((x1*y2 + y1*x2)/(1 + d*x1*x2*y1*y2)) , ((y1*y2 - a*x1*x2)/(1 - d*x1*x2*y1*y2)))) _.
+ Program Definition add (P1 P2:point) : point :=
+ let x1y1 := coordinates P1 in let x1 := fst x1y1 in let y1 := snd x1y1 in
+ let x2y2 := coordinates P2 in let x2 := fst x2y2 in let y2 := snd x2y2 in
+ (((x1*y2 + y1*x2)/(1 + d*x1*x2*y1*y2)) , ((y1*y2 - a*x1*x2)/(1 - d*x1*x2*y1*y2))).
+ Next Obligation. destruct P1 as [[??]?], P2 as [[??]?], H; auto using Pre.add_onCurve. Defined.
Fixpoint mul (n:nat) (P : point) : point :=
match n with