aboutsummaryrefslogtreecommitdiff
path: root/src/SpecificGen/GF25519_64.v
diff options
context:
space:
mode:
Diffstat (limited to 'src/SpecificGen/GF25519_64.v')
-rw-r--r--src/SpecificGen/GF25519_64.v36
1 files changed, 36 insertions, 0 deletions
diff --git a/src/SpecificGen/GF25519_64.v b/src/SpecificGen/GF25519_64.v
index 49b5fb2c4..3645a59f7 100644
--- a/src/SpecificGen/GF25519_64.v
+++ b/src/SpecificGen/GF25519_64.v
@@ -176,6 +176,25 @@ Proof.
etransitivity; apply app_n_correct.
Qed.
+Definition appify9 {T} (op : fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> T) (x0 x1 x2 x3 x4 x5 x6 x7 x8 : fe25519_64) :=
+ app_n x0 (fun x0' =>
+ app_n x1 (fun x1' =>
+ app_n x2 (fun x2' =>
+ app_n x3 (fun x3' =>
+ app_n x4 (fun x4' =>
+ app_n x5 (fun x5' =>
+ app_n x6 (fun x6' =>
+ app_n x7 (fun x7' =>
+ app_n x8 (fun x8' =>
+ op x0' x1' x2' x3' x4' x5' x6' x7' x8'))))))))).
+
+Lemma appify9_correct : forall {T} op x0 x1 x2 x3 x4 x5 x6 x7 x8,
+ @appify9 T op x0 x1 x2 x3 x4 x5 x6 x7 x8 = op x0 x1 x2 x3 x4 x5 x6 x7 x8.
+Proof.
+ intros. cbv [appify9].
+ repeat (etransitivity; [ apply app_n_correct | ]); reflexivity.
+Qed.
+
Definition uncurry_unop_fe25519_64 {T} (op : fe25519_64 -> T)
:= Eval compute in Tuple.uncurry (n:=length_fe25519_64) op.
Definition curry_unop_fe25519_64 {T} op : fe25519_64 -> T
@@ -190,6 +209,23 @@ Definition uncurry_unop_wire_digits {T} (op : wire_digits -> T)
Definition curry_unop_wire_digits {T} op : wire_digits -> T
:= Eval compute in fun f => app_n2 f (Tuple.curry (n:=length wire_widths) op).
+Definition uncurry_9op_fe25519_64 {T} (op : fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> T)
+ := Eval compute in
+ uncurry_unop_fe25519_64 (fun x0 =>
+ uncurry_unop_fe25519_64 (fun x1 =>
+ uncurry_unop_fe25519_64 (fun x2 =>
+ uncurry_unop_fe25519_64 (fun x3 =>
+ uncurry_unop_fe25519_64 (fun x4 =>
+ uncurry_unop_fe25519_64 (fun x5 =>
+ uncurry_unop_fe25519_64 (fun x6 =>
+ uncurry_unop_fe25519_64 (fun x7 =>
+ uncurry_unop_fe25519_64 (fun x8 =>
+ op x0 x1 x2 x3 x4 x5 x6 x7 x8))))))))).
+Definition curry_9op_fe25519_64 {T} op : fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> fe25519_64 -> T
+ := Eval compute in
+ appify9 (fun x0 x1 x2 x3 x4 x5 x6 x7 x8
+ => curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 (curry_unop_fe25519_64 op x0) x1) x2) x3) x4) x5) x6) x7) x8).
+
Definition add_sig (f g : fe25519_64) :
{ fg : fe25519_64 | fg = add_opt f g}.
Proof.