aboutsummaryrefslogtreecommitdiff
path: root/src/Specific/montgomery64_2e413m21_7limbs
diff options
context:
space:
mode:
authorGravatar jadep <jade.philipoom@gmail.com>2017-11-10 13:19:57 -0500
committerGravatar jadep <jade.philipoom@gmail.com>2017-11-12 15:01:59 -0500
commit795d24349b9aca1d9732c7b7fcaa505f24fa4bc6 (patch)
treef517d6c4e83f42e72a303e06567f779c0250fc2e /src/Specific/montgomery64_2e413m21_7limbs
parent7ad53a35ed68777cd21226998a88494e1d97c63e (diff)
new autogenerated files
Diffstat (limited to 'src/Specific/montgomery64_2e413m21_7limbs')
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/CurveParameters.v39
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/Synthesis.v9
-rwxr-xr-xsrc/Specific/montgomery64_2e413m21_7limbs/compiler.sh4
-rwxr-xr-xsrc/Specific/montgomery64_2e413m21_7limbs/compilerxx.sh4
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/feadd.v14
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/feaddDisplay.v4
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/femul.v14
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/femulDisplay.v4
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/fenz.v16
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/fenzDisplay.v4
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/feopp.v14
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/feoppDisplay.v4
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/fesub.v14
-rw-r--r--src/Specific/montgomery64_2e413m21_7limbs/fesubDisplay.v4
-rwxr-xr-xsrc/Specific/montgomery64_2e413m21_7limbs/py_interpreter.sh4
15 files changed, 152 insertions, 0 deletions
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/CurveParameters.v b/src/Specific/montgomery64_2e413m21_7limbs/CurveParameters.v
new file mode 100644
index 000000000..072b07875
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/CurveParameters.v
@@ -0,0 +1,39 @@
+Require Import Crypto.Specific.Framework.RawCurveParameters.
+Require Import Crypto.Util.LetIn.
+
+(***
+Modulus : 2^413 - 21
+Base: 64
+***)
+
+Definition curve : CurveParameters :=
+ {|
+ sz := 7%nat;
+ base := 64;
+ bitwidth := 64;
+ s := 2^413;
+ c := [(1, 21)];
+ carry_chains := None;
+
+ a24 := None;
+ coef_div_modulus := None;
+
+ goldilocks := None;
+ karatsuba := None;
+ montgomery := true;
+ freeze := Some false;
+ ladderstep := false;
+
+ mul_code := None;
+
+ square_code := None;
+
+ upper_bound_of_exponent_loose := None;
+ upper_bound_of_exponent_tight := None;
+ allowable_bit_widths := None;
+ freeze_extra_allowable_bit_widths := None;
+ modinv_fuel := None
+ |}.
+
+Ltac extra_prove_mul_eq _ := idtac.
+Ltac extra_prove_square_eq _ := idtac.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/Synthesis.v b/src/Specific/montgomery64_2e413m21_7limbs/Synthesis.v
new file mode 100644
index 000000000..6e1f63661
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/Synthesis.v
@@ -0,0 +1,9 @@
+Require Import Crypto.Specific.Framework.SynthesisFramework.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.CurveParameters.
+
+Module P <: PrePackage.
+ Definition package : Tag.Context.
+ Proof. make_Synthesis_package curve extra_prove_mul_eq extra_prove_square_eq. Defined.
+End P.
+
+Module Export S := PackageSynthesis P.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/compiler.sh b/src/Specific/montgomery64_2e413m21_7limbs/compiler.sh
new file mode 100755
index 000000000..404f82538
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/compiler.sh
@@ -0,0 +1,4 @@
+#!/bin/sh
+set -eu
+
+clang -fbracket-depth=999999 -march=native -mtune=native -std=gnu11 -O3 -flto -fuse-ld=lld -fomit-frame-pointer -fwrapv -Wno-attributes -fno-strict-aliasing -Da24_hex='0x3039' -Da24_val='12345' -Da_minus_two_over_four_array='{0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x30,0x39}' -Dbitwidth='64' -Dlimb_weight_gaps_array='{64,64,64,64,64,64,64}' -Dmodulus_array='{0x1f,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xeb}' -Dmodulus_bytes_val='52' -Dmodulus_limbs='7' -Dq_mpz='(1_mpz<<413) - 21' "$@"
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/compilerxx.sh b/src/Specific/montgomery64_2e413m21_7limbs/compilerxx.sh
new file mode 100755
index 000000000..934a9dfbe
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/compilerxx.sh
@@ -0,0 +1,4 @@
+#!/bin/sh
+set -eu
+
+clang++ -fbracket-depth=999999 -march=native -mtune=native -std=gnu++11 -O3 -flto -fuse-ld=lld -fomit-frame-pointer -fwrapv -Wno-attributes -fno-strict-aliasing -Da24_hex='0x3039' -Da24_val='12345' -Da_minus_two_over_four_array='{0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x00,0x30,0x39}' -Dbitwidth='64' -Dlimb_weight_gaps_array='{64,64,64,64,64,64,64}' -Dmodulus_array='{0x1f,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xff,0xeb}' -Dmodulus_bytes_val='52' -Dmodulus_limbs='7' -Dq_mpz='(1_mpz<<413) - 21' "$@"
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/feadd.v b/src/Specific/montgomery64_2e413m21_7limbs/feadd.v
new file mode 100644
index 000000000..8cb127d43
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/feadd.v
@@ -0,0 +1,14 @@
+Require Import Crypto.Arithmetic.PrimeFieldTheorems.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.Synthesis.
+
+(* TODO : change this to field once field isomorphism happens *)
+Definition add :
+ { add : feBW_small -> feBW_small -> feBW_small
+ | forall a b, phiM_small (add a b) = F.add (phiM_small a) (phiM_small b) }.
+Proof.
+ Set Ltac Profiling.
+ Time synthesize_add ().
+ Show Ltac Profile.
+Time Defined.
+
+Print Assumptions add.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/feaddDisplay.v b/src/Specific/montgomery64_2e413m21_7limbs/feaddDisplay.v
new file mode 100644
index 000000000..cdd899cb1
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/feaddDisplay.v
@@ -0,0 +1,4 @@
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.feadd.
+Require Import Crypto.Specific.Framework.IntegrationTestDisplayCommon.
+
+Check display add.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/femul.v b/src/Specific/montgomery64_2e413m21_7limbs/femul.v
new file mode 100644
index 000000000..cdad0d00e
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/femul.v
@@ -0,0 +1,14 @@
+Require Import Crypto.Arithmetic.PrimeFieldTheorems.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.Synthesis.
+
+(* TODO : change this to field once field isomorphism happens *)
+Definition mul :
+ { mul : feBW_small -> feBW_small -> feBW_small
+ | forall a b, phiM_small (mul a b) = F.mul (phiM_small a) (phiM_small b) }.
+Proof.
+ Set Ltac Profiling.
+ Time synthesize_mul ().
+ Show Ltac Profile.
+Time Defined.
+
+Print Assumptions mul.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/femulDisplay.v b/src/Specific/montgomery64_2e413m21_7limbs/femulDisplay.v
new file mode 100644
index 000000000..f55c9c649
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/femulDisplay.v
@@ -0,0 +1,4 @@
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.femul.
+Require Import Crypto.Specific.Framework.IntegrationTestDisplayCommon.
+
+Check display mul.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/fenz.v b/src/Specific/montgomery64_2e413m21_7limbs/fenz.v
new file mode 100644
index 000000000..ef30f537d
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/fenz.v
@@ -0,0 +1,16 @@
+Require Import Coq.ZArith.ZArith.
+Require Import Crypto.Arithmetic.PrimeFieldTheorems.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.Synthesis.
+Local Open Scope Z_scope.
+
+(* TODO : change this to field once field isomorphism happens *)
+Definition nonzero :
+ { nonzero : feBW_small -> BoundedWord.BoundedWord 1 adjusted_bitwidth bound1
+ | forall a, (BoundedWord.BoundedWordToZ _ _ _ (nonzero a) =? 0) = (if Decidable.dec (phiM_small a = F.of_Z m 0) then true else false) }.
+Proof.
+ Set Ltac Profiling.
+ Time synthesize_nonzero ().
+ Show Ltac Profile.
+Time Defined.
+
+Print Assumptions nonzero.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/fenzDisplay.v b/src/Specific/montgomery64_2e413m21_7limbs/fenzDisplay.v
new file mode 100644
index 000000000..b94d72a55
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/fenzDisplay.v
@@ -0,0 +1,4 @@
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.fenz.
+Require Import Crypto.Specific.Framework.IntegrationTestDisplayCommon.
+
+Check display nonzero.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/feopp.v b/src/Specific/montgomery64_2e413m21_7limbs/feopp.v
new file mode 100644
index 000000000..45dd48d6f
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/feopp.v
@@ -0,0 +1,14 @@
+Require Import Crypto.Arithmetic.PrimeFieldTheorems.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.Synthesis.
+
+(* TODO : change this to field once field isomorphism happens *)
+Definition opp :
+ { opp : feBW_small -> feBW_small
+ | forall a, phiM_small (opp a) = F.opp (phiM_small a) }.
+Proof.
+ Set Ltac Profiling.
+ Time synthesize_opp ().
+ Show Ltac Profile.
+Time Defined.
+
+Print Assumptions opp.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/feoppDisplay.v b/src/Specific/montgomery64_2e413m21_7limbs/feoppDisplay.v
new file mode 100644
index 000000000..44cf585c9
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/feoppDisplay.v
@@ -0,0 +1,4 @@
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.feopp.
+Require Import Crypto.Specific.Framework.IntegrationTestDisplayCommon.
+
+Check display opp.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/fesub.v b/src/Specific/montgomery64_2e413m21_7limbs/fesub.v
new file mode 100644
index 000000000..4f94daa7d
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/fesub.v
@@ -0,0 +1,14 @@
+Require Import Crypto.Arithmetic.PrimeFieldTheorems.
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.Synthesis.
+
+(* TODO : change this to field once field isomorphism happens *)
+Definition sub :
+ { sub : feBW_small -> feBW_small -> feBW_small
+ | forall a b, phiM_small (sub a b) = F.sub (phiM_small a) (phiM_small b) }.
+Proof.
+ Set Ltac Profiling.
+ Time synthesize_sub ().
+ Show Ltac Profile.
+Time Defined.
+
+Print Assumptions sub.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/fesubDisplay.v b/src/Specific/montgomery64_2e413m21_7limbs/fesubDisplay.v
new file mode 100644
index 000000000..4ffc415f6
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/fesubDisplay.v
@@ -0,0 +1,4 @@
+Require Import Crypto.Specific.montgomery64_2e413m21_7limbs.fesub.
+Require Import Crypto.Specific.Framework.IntegrationTestDisplayCommon.
+
+Check display sub.
diff --git a/src/Specific/montgomery64_2e413m21_7limbs/py_interpreter.sh b/src/Specific/montgomery64_2e413m21_7limbs/py_interpreter.sh
new file mode 100755
index 000000000..8be5e5a3b
--- /dev/null
+++ b/src/Specific/montgomery64_2e413m21_7limbs/py_interpreter.sh
@@ -0,0 +1,4 @@
+#!/bin/sh
+set -eu
+
+/usr/bin/env python3 "$@" -Dq='2**413 - 21' -Dmodulus_bytes='64' -Da24='121665'