aboutsummaryrefslogtreecommitdiff
path: root/src/Specific/GF25519.v
diff options
context:
space:
mode:
Diffstat (limited to 'src/Specific/GF25519.v')
-rw-r--r--src/Specific/GF25519.v70
1 files changed, 25 insertions, 45 deletions
diff --git a/src/Specific/GF25519.v b/src/Specific/GF25519.v
index 2f1da1846..7ece3a5da 100644
--- a/src/Specific/GF25519.v
+++ b/src/Specific/GF25519.v
@@ -1,10 +1,11 @@
Require Import Crypto.BaseSystem.
Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
-Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParams.
Require Import Crypto.ModularArithmetic.PseudoMersenneBaseParamProofs.
Require Import Crypto.ModularArithmetic.ModularBaseSystem.
Require Import Crypto.ModularArithmetic.ModularBaseSystemProofs.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemOpt.
+Require Import Crypto.ModularArithmetic.ModularBaseSystemField.
Require Import Coq.Lists.List Crypto.Util.ListUtil.
Require Import Crypto.Tactics.VerdiTactics.
Require Import Crypto.Util.ZUtil.
@@ -27,7 +28,7 @@ Defined.
Definition fe25519 := Eval compute in (tuple Z (length limb_widths)).
-Definition mul2modulus :=
+Definition mul2modulus : fe25519 :=
Eval compute in (from_list_default 0%Z (length limb_widths) (construct_mul2modulus params25519)).
Instance subCoeff : SubtractionCoefficient modulus params25519.
@@ -51,7 +52,7 @@ Definition c_subst : c = c_ := eq_refl c_.
Local Opaque Z.shiftr Z.shiftl Z.land Z.mul Z.add Z.sub Let_In.
-Definition app_10 (f : fe25519) (P : fe25519 -> fe25519) : fe25519.
+Definition app_5 (f : fe25519) (P : fe25519 -> fe25519) : fe25519.
Proof.
cbv [fe25519] in *.
set (f0 := f).
@@ -60,7 +61,7 @@ Proof.
apply f0.
Defined.
-Lemma app_10_correct : forall f P, app_10 f P = P f.
+Definition app_5_correct f (P : fe25519 -> fe25519) : app_5 f P = P f.
Proof.
intros.
cbv [fe25519] in *.
@@ -69,16 +70,16 @@ Proof.
Qed.
Definition appify2 (op : fe25519 -> fe25519 -> fe25519) (f g : fe25519):=
- app_10 f (fun f0 => (app_10 g (fun g0 => op f0 g0))).
+ app_5 f (fun f0 => (app_5 g (fun g0 => op f0 g0))).
Lemma appify2_correct : forall op f g, appify2 op f g = op f g.
Proof.
intros. cbv [appify2].
- etransitivity; apply app_10_correct.
+ etransitivity; apply app_5_correct.
Qed.
Definition add_sig (f g : fe25519) :
- { fg : fe25519 | fg = ModularBaseSystem.add f g}.
+ { fg : fe25519 | fg = add_opt f g}.
Proof.
eexists.
rewrite <-appify2_correct.
@@ -91,12 +92,12 @@ Definition add (f g : fe25519) : fe25519 :=
proj1_sig (add_sig f g).
Definition add_correct (f g : fe25519)
- : add f g = ModularBaseSystem.add f g :=
+ : add f g = add_opt f g :=
Eval cbv beta iota delta [proj1_sig add_sig] in
- proj2_sig (add_sig f g). (* Coq 8.4 : 10s *)
+ proj2_sig (add_sig f g).
Definition sub_sig (f g : fe25519) :
- { fg : fe25519 | fg = ModularBaseSystem.sub mul2modulus f g}.
+ { fg : fe25519 | fg = sub_opt f g}.
Proof.
eexists.
rewrite <-appify2_correct.
@@ -109,19 +110,18 @@ Definition sub (f g : fe25519) : fe25519 :=
proj1_sig (sub_sig f g).
Definition sub_correct (f g : fe25519)
- : sub f g = ModularBaseSystem.sub mul2modulus f g :=
+ : sub f g = sub_opt f g :=
Eval cbv beta iota delta [proj1_sig sub_sig] in
proj2_sig (sub_sig f g).
(* For multiplication, we add another layer of definition so that we can
rewrite under the [let] binders. *)
Definition mul_simpl_sig (f g : fe25519) :
- { fg : fe25519 | fg = ModularBaseSystem.mul f g}.
+ { fg : fe25519 | fg = carry_mul_opt k_ c_ f g}.
Proof.
cbv [fe25519] in *.
repeat match goal with p : (_ * Z)%type |- _ => destruct p end.
eexists.
- rewrite <-mul_opt_correct with (k_ := k_) (c_ := c_) by auto.
cbv.
autorewrite with zsimplify.
reflexivity.
@@ -132,12 +132,12 @@ Definition mul_simpl (f g : fe25519) : fe25519 :=
proj1_sig (mul_simpl_sig f g).
Definition mul_simpl_correct (f g : fe25519)
- : mul_simpl f g = ModularBaseSystem.mul f g :=
+ : mul_simpl f g = carry_mul_opt k_ c_ f g :=
Eval cbv beta iota delta [proj1_sig mul_simpl_sig] in
proj2_sig (mul_simpl_sig f g).
Definition mul_sig (f g : fe25519) :
- { fg : fe25519 | fg = ModularBaseSystem.mul f g}.
+ { fg : fe25519 | fg = carry_mul_opt k_ c_ f g}.
Proof.
eexists.
rewrite <-mul_simpl_correct.
@@ -151,45 +151,25 @@ Definition mul (f g : fe25519) : fe25519 :=
proj1_sig (mul_sig f g).
Definition mul_correct (f g : fe25519)
- : mul f g = ModularBaseSystem.mul f g :=
+ : mul f g = carry_mul_opt k_ c_ f g :=
Eval cbv beta iota delta [proj1_sig add_sig] in
proj2_sig (mul_sig f g).
-Definition decode := Eval compute in ModularBaseSystem.decode.
-
Import Morphisms.
Lemma field25519 : @field fe25519 eq zero one opp add sub mul inv div.
Proof.
pose proof (Equivalence_Reflexive : Reflexive eq).
- eapply (Field.isomorphism_to_subfield_field (phi := decode)
- (fieldR := PrimeFieldTheorems.field_modulo (prime_q := prime_modulus))).
+ eapply (Field.equivalent_operations_field (fieldR := modular_base_system_field k_ c_ k_subst c_subst)).
Grab Existential Variables.
- + intros; change decode with ModularBaseSystem.decode; apply encode_rep.
- + intros; change decode with ModularBaseSystem.decode; apply encode_rep.
- + intros; change decode with ModularBaseSystem.decode; apply encode_rep.
- + intros; change decode with ModularBaseSystem.decode; apply encode_rep.
- + intros; change decode with ModularBaseSystem.decode.
- rewrite mul_correct; apply mul_rep; reflexivity.
- + intros; change decode with ModularBaseSystem.decode.
- rewrite sub_correct; apply sub_rep; try reflexivity.
- rewrite <- coeff_mod. reflexivity.
- + intros; change decode with ModularBaseSystem.decode.
- rewrite add_correct; apply add_rep; reflexivity.
- + intros; change decode with ModularBaseSystem.decode; apply encode_rep.
- + cbv [eq zero one]. change decode with ModularBaseSystem.decode.
- rewrite !encode_rep. intro A.
- eapply (PrimeFieldTheorems.Fq_1_neq_0 (prime_q := prime_modulus)). congruence.
- + trivial.
- + repeat intro. cbv [div]. congruence.
- + repeat intro. cbv [inv]. congruence.
- + repeat intro. cbv [eq] in *. rewrite !mul_correct, !mul_rep by reflexivity; congruence.
- + repeat intro. cbv [eq] in *. rewrite !sub_correct. rewrite !sub_rep by
- (rewrite <-?coeff_mod; reflexivity); congruence.
- + repeat intro. cbv [eq] in *. rewrite !add_correct, !add_rep by reflexivity; congruence.
- + repeat intro. cbv [opp]. congruence.
- + cbv [eq]. auto using ModularArithmeticTheorems.F_eq_dec.
- + apply (eq_Equivalence (prm := params25519)).
+ + reflexivity.
+ + reflexivity.
+ + reflexivity.
+ + intros; rewrite mul_correct. reflexivity.
+ + intros; rewrite sub_correct; reflexivity.
+ + intros; rewrite add_correct; reflexivity.
+ + reflexivity.
+ + reflexivity.
Qed.