aboutsummaryrefslogtreecommitdiff
path: root/src/Specific/GF25519.v
diff options
context:
space:
mode:
authorGravatar jadep <jade.philipoom@gmail.com>2016-04-21 13:28:53 -0400
committerGravatar jadep <jade.philipoom@gmail.com>2016-04-21 13:28:53 -0400
commit7e3e279a7521a043ed70a0a01214985fc7b92358 (patch)
tree006022b3da14a01c9b366b626273ce0fc6089803 /src/Specific/GF25519.v
parent8720fbfec19d4a4823a5a65771a8122af7e83a09 (diff)
automated most of the code in GF25519
Diffstat (limited to 'src/Specific/GF25519.v')
-rw-r--r--src/Specific/GF25519.v61
1 files changed, 15 insertions, 46 deletions
diff --git a/src/Specific/GF25519.v b/src/Specific/GF25519.v
index 695cfc254..1f6be80cc 100644
--- a/src/Specific/GF25519.v
+++ b/src/Specific/GF25519.v
@@ -15,24 +15,20 @@ Local Open Scope Z.
Definition modulus : Z := 2^255 - 19.
Lemma prime_modulus : prime modulus. Admitted.
-Definition limb_widths : list Z := [26;25;26;25;26;25;26;25;26;25].
-
-Instance params25519 : PseudoMersenneBaseParams modulus := { limb_widths := limb_widths }.
-Proof.
- + abstract (apply fold_right_and_True_forall_In_iff; simpl; repeat (split; [omega |]); auto).
- + abstract (unfold limb_widths; congruence).
- + abstract brute_force_indices limb_widths.
- + abstract apply prime_modulus.
- + abstract brute_force_indices limb_widths.
+Instance params25519 : PseudoMersenneBaseParams modulus.
+ construct_params prime_modulus 10%nat 255.
Defined.
(* END PseudoMersenneBaseParams instance construction. *)
(* Precompute k and c *)
-Definition k_ := 255.
-Lemma k_subst : k_ = k. reflexivity. Qed.
-Definition c_ := 19.
-Lemma c_subst : c_ = c. reflexivity. Qed.
+Definition k_ := Eval compute in k.
+Definition c_ := Eval compute in c.
+
+(* Makes Qed not take forever *)
+Opaque Z.shiftr Pos.iter Z.div2 Pos.div2 Pos.div2_up Pos.succ Z.land
+ Z.of_N Pos.land N.ldiff Pos.pred_N Pos.pred_double Z.opp Z.mul Pos.mul
+ Let_In digits Z.add Pos.add Z.pos_sub.
Local Open Scope nat_scope.
Lemma GF25519Base25Point5_mul_reduce_formula :
@@ -42,29 +38,9 @@ Lemma GF25519Base25Point5_mul_reduce_formula :
-> rep [g0;g1;g2;g3;g4;g5;g6;g7;g8;g9] g
-> rep ls (f*g)%F}.
Proof.
- eexists.
- intros f g Hf Hg.
- pose proof (carry_mul_opt_correct k_ c_ k_subst c_subst [0;9;8;7;6;5;4;3;2;1;0]_ _ _ _ Hf Hg) as Hfg.
- forward Hfg; [abstract (clear; cbv; intros; repeat break_or_hyp; intuition)|].
- specialize (Hfg H); clear H.
- forward Hfg; [exact eq_refl|].
- specialize (Hfg H); clear H.
-
- set (p := params25519) in Hfg at 1.
- change (params25519) with p at 1.
- set (fg := (f * g)%F) in Hfg |- * .
-
- Opaque Z.shiftl Pos.iter Z.div2 Pos.div2 Pos.div2_up Pos.succ Z.land
- Z.of_N Pos.land N.ldiff Pos.pred_N Pos.pred_double Z.opp Z.mul Pos.mul
- Let_In digits Z.add Pos.add Z.pos_sub.
-
- cbv -[fg rep] in Hfg.
-
- change (Z.shiftl 1 26 + -1)%Z with 67108863%Z in Hfg.
- change (Z.shiftl 1 25 + -1)%Z with 33554431%Z in Hfg.
- repeat rewrite ?Z.mul_1_r, ?Z.add_0_l, ?Z.add_assoc, ?Z.mul_assoc in Hfg.
-
- exact Hfg.
+ eexists; intros ? ? Hf Hg.
+ pose proof (carry_mul_opt_correct k_ c_ (eq_refl k_) (eq_refl c_) [0;9;8;7;6;5;4;3;2;1;0]_ _ _ _ Hf Hg) as Hfg.
+ compute_formula.
Time Defined.
Extraction "/tmp/test.ml" GF25519Base25Point5_mul_reduce_formula.
@@ -73,6 +49,7 @@ Extraction "/tmp/test.ml" GF25519Base25Point5_mul_reduce_formula.
* The easiest fix will be to make the proof script above fully automated,
* using [abstract] to contain the proof part. *)
+
Lemma GF25519Base25Point5_add_formula :
forall f0 f1 f2 f3 f4 f5 f6 f7 f8 f9
g0 g1 g2 g3 g4 g5 g6 g7 g8 g9,
@@ -83,13 +60,5 @@ Proof.
eexists.
intros f g Hf Hg.
pose proof (add_opt_rep _ _ _ _ Hf Hg) as Hfg.
-
- set (p := params25519) in Hfg at 1.
- change (params25519) with p at 1.
- set (fg := (f + g)%F) in Hfg |- * .
-
- cbv -[fg rep] in Hfg.
- change (Z.shiftl 1 26 + -1)%Z with 67108863%Z in Hfg.
- change (Z.shiftl 1 25 + -1)%Z with 33554431%Z in Hfg.
- exact Hfg.
-Defined.
+ compute_formula.
+Defined. \ No newline at end of file