Browse Source

Merge b120f9e111 into 38821bbf14

pull/1/merge
Youssef El Housni 2 months ago
committed by GitHub
parent
commit
6263d26260
No known key found for this signature in database GPG Key ID: B5690EEEBB952194
3 changed files with 461 additions and 69 deletions
  1. +166
    -69
      src/groups/curves/short_weierstrass/mod.rs
  2. +119
    -0
      src/groups/curves/short_weierstrass/non_zero_affine.rs
  3. +176
    -0
      src/groups/mod.rs

+ 166
- 69
src/groups/curves/short_weierstrass/mod.rs

@ -350,23 +350,38 @@ where
// We can convert to projective safely because the result is guaranteed to be
// non-zero by the condition on `affine_bits.len()`, and by the fact
// that `accumulator` is non-zero
let result = accumulator.into_projective();
*mul_result += accumulator.into_projective();
// If bits[0] is 0, then we have to subtract `self`; else, we subtract zero.
let subtrahend = bits[0].select(&Self::zero(), &initial_acc_value)?;
*mul_result += result - subtrahend;
let neg = NonZeroAffineVar::new(initial_acc_value.x, initial_acc_value.y.negate()?);
*mul_result = bits[0].select(mul_result, &mul_result.add_mixed(&neg)?)?;
// Now, let's finish off the rest of the bits using our complete formulae
for bit in proj_bits {
for bit in proj_bits.iter().rev().skip(1).rev() {
if bit.is_constant() {
if *bit == &Boolean::TRUE {
*mul_result += &multiple_of_power_of_two.into_projective();
*mul_result = mul_result.add_mixed(&multiple_of_power_of_two)?;
}
} else {
let temp = &*mul_result + &multiple_of_power_of_two.into_projective();
let temp = mul_result.add_mixed(&multiple_of_power_of_two)?;
*mul_result = bit.select(&temp, &mul_result)?;
}
multiple_of_power_of_two.double_in_place()?;
}
// last bit
// we don't need the last doubling of multiple_of_power_of_two
let n = proj_bits.len();
if n >= 1 {
if proj_bits[n - 1].is_constant() {
if proj_bits[n - 1] == &Boolean::TRUE {
*mul_result = mul_result.add_mixed(&multiple_of_power_of_two)?;
}
} else {
let temp = mul_result.add_mixed(&multiple_of_power_of_two)?;
*mul_result = proj_bits[n - 1].select(&temp, &mul_result)?;
}
}
Ok(())
}
}
@ -498,6 +513,78 @@ where
Ok(Self::new(self.x.clone(), self.y.negate()?, self.z.clone()))
}
/// Computes `bits * self`, where `bits` is a little-endian
/// `Boolean` representation of a scalar.
///
/// [Joye07](<https://www.iacr.org/archive/ches2007/47270135/47270135.pdf>), Alg.1.
#[tracing::instrument(target = "r1cs", skip(bits))]
fn scalar_mul_joye_le<'a>(
&self,
bits: impl Iterator<Item = &'a Boolean<<P::BaseField as Field>::BasePrimeField>>,
) -> Result<Self, SynthesisError> {
if self.is_constant() {
if self.value().unwrap().is_zero() {
return Ok(self.clone());
}
}
let self_affine = self.to_affine()?;
let (x, y, infinity) = (self_affine.x, self_affine.y, self_affine.infinity);
// We first handle the non-zero case, and then later will conditionally select
// zero if `self` was zero. However, we also want to make sure that generated
// constraints are satisfiable in both cases.
//
// In particular, using non-sensible values for `x` and `y` in zero-case may
// cause `unchecked` operations to generate constraints that can never
// be satisfied, depending on the curve equation coefficients.
//
// The safest approach is to use coordinates of some point from the curve, thus
// not violating assumptions of `NonZeroAffine`. For instance, generator
// point.
let x = infinity.select(&F::constant(P::GENERATOR.x), &x)?;
let y = infinity.select(&F::constant(P::GENERATOR.y), &y)?;
let non_zero_self = NonZeroAffineVar::new(x, y);
let mut bits = bits.collect::<Vec<_>>();
if bits.len() == 0 {
return Ok(Self::zero());
}
// Remove unnecessary constant zeros in the most-significant positions.
bits = bits
.into_iter()
// We iterate from the MSB down.
.rev()
// Skip leading zeros, if they are constants.
.skip_while(|b| b.is_constant() && (b.value().unwrap() == false))
.collect();
// After collecting we are in big-endian form; we have to reverse to get back to
// little-endian.
bits.reverse();
// second bit
let mut acc = non_zero_self.triple()?;
let mut acc0 = bits[1].select(&acc, &non_zero_self)?;
let mut acc1 = bits[1].select(&non_zero_self, &acc)?;
for bit in bits.iter().skip(2).rev().skip(1).rev() {
acc = acc0.double_and_select_add_unchecked(bit, &acc1)?;
acc0 = bit.select(&acc, &acc0)?;
acc1 = bit.select(&acc1, &acc)?;
}
// last bit
let n = bits.len() - 1;
acc = acc0.double_and_select_add_unchecked(bits[n], &acc1)?;
acc0 = bits[n].select(&acc, &acc0)?;
// first bit
let temp = NonZeroAffineVar::new(non_zero_self.x, non_zero_self.y.negate()?);
acc1 = acc0.add_unchecked(&temp)?;
acc0 = bits[0].select(&acc0, &acc1)?;
let mul_result = acc0.into_projective();
infinity.select(&Self::zero(), &mul_result)
}
/// Computes `bits * self`, where `bits` is a little-endian
/// `Boolean` representation of a scalar.
#[tracing::instrument(target = "r1cs", skip(bits))]
@ -516,12 +603,13 @@ where
// zero if `self` was zero. However, we also want to make sure that generated
// constraints are satisfiable in both cases.
//
// In particular, using non-sensible values for `x` and `y` in zero-case may cause
// `unchecked` operations to generate constraints that can never be satisfied, depending
// on the curve equation coefficients.
// In particular, using non-sensible values for `x` and `y` in zero-case may
// cause `unchecked` operations to generate constraints that can never
// be satisfied, depending on the curve equation coefficients.
//
// The safest approach is to use coordinates of some point from the curve, thus not
// violating assumptions of `NonZeroAffine`. For instance, generator point.
// The safest approach is to use coordinates of some point from the curve, thus
// not violating assumptions of `NonZeroAffine`. For instance, generator
// point.
let x = infinity.select(&F::constant(P::GENERATOR.x), &x)?;
let y = infinity.select(&F::constant(P::GENERATOR.y), &y)?;
let non_zero_self = NonZeroAffineVar::new(x, y);
@ -558,6 +646,73 @@ where
infinity.select(&Self::zero(), &mul_result)
}
/// Computes `bits1 * self + bits2 * p`, where `bits1` and `bits2` are
/// big-endian `Boolean` representation of the scalars.
///
/// `self` and `p` are non-zero and `self` ≠ `-p`.
#[tracing::instrument(target = "r1cs", skip(bits1, bits2))]
fn joint_scalar_mul_be<'a>(
&self,
p: &Self,
bits1: impl Iterator<Item = &'a Boolean<<P::BaseField as Field>::BasePrimeField>>,
bits2: impl Iterator<Item = &'a Boolean<<P::BaseField as Field>::BasePrimeField>>,
) -> Result<Self, SynthesisError> {
// prepare bits decomposition
let mut bits1 = bits1.collect::<Vec<_>>();
if bits1.len() == 0 {
return Ok(Self::zero());
}
// Remove unnecessary constant zeros in the most-significant positions.
bits1 = bits1
.into_iter()
// We iterate from the MSB down.
.rev()
// Skip leading zeros, if they are constants.
.skip_while(|b| b.is_constant() && (b.value().unwrap() == false))
.collect();
let mut bits2 = bits2.collect::<Vec<_>>();
if bits2.len() == 0 {
return Ok(Self::zero());
}
// Remove unnecessary constant zeros in the most-significant positions.
bits2 = bits2
.into_iter()
// We iterate from the MSB down.
.rev()
// Skip leading zeros, if they are constants.
.skip_while(|b| b.is_constant() && (b.value().unwrap() == false))
.collect();
// precompute points
let aff1 = self.to_affine()?;
let nz_aff1 = NonZeroAffineVar::new(aff1.x, aff1.y);
let aff2 = p.to_affine()?;
let nz_aff2 = NonZeroAffineVar::new(aff2.x, aff2.y);
let mut aff1_neg = NonZeroAffineVar::new(nz_aff1.x.clone(), nz_aff1.y.negate()?);
let mut aff2_neg = NonZeroAffineVar::new(nz_aff2.x.clone(), nz_aff2.y.negate()?);
let mut acc = nz_aff1.add_unchecked(&nz_aff2.clone())?;
// double-and-add loop
for (bit1, bit2) in (bits1.iter().rev().skip(1).rev()).zip(bits2.iter().rev().skip(1).rev())
{
let mut b = bit1.select(&nz_aff1, &aff1_neg)?;
acc = acc.double_and_add_unchecked(&b)?;
b = bit2.select(&nz_aff2, &aff2_neg)?;
acc = acc.add_unchecked(&b)?;
}
// last bit
aff1_neg = aff1_neg.add_unchecked(&acc)?;
acc = bits1[bits1.len() - 1].select(&acc, &aff1_neg)?;
aff2_neg = aff2_neg.add_unchecked(&acc)?;
acc = bits2[bits1.len() - 1].select(&acc, &aff2_neg)?;
Ok(acc.into_projective())
}
#[tracing::instrument(target = "r1cs", skip(scalar_bits_with_bases))]
fn precomputed_base_scalar_mul_le<'a, I, B>(
&mut self,
@ -978,61 +1133,3 @@ where
Ok(bytes)
}
}
#[cfg(test)]
mod test_sw_curve {
use crate::{
alloc::AllocVar,
eq::EqGadget,
fields::{fp::FpVar, nonnative::NonNativeFieldVar},
groups::{curves::short_weierstrass::ProjectiveVar, CurveVar},
ToBitsGadget,
};
use ark_ec::{
short_weierstrass::{Projective, SWCurveConfig},
CurveGroup,
};
use ark_ff::PrimeField;
use ark_relations::r1cs::{ConstraintSystem, Result};
use ark_std::UniformRand;
use num_traits::Zero;
fn zero_point_scalar_mul_satisfied<G>() -> Result<bool>
where
G: CurveGroup,
G::BaseField: PrimeField,
G::Config: SWCurveConfig,
{
let mut rng = ark_std::test_rng();
let cs = ConstraintSystem::new_ref();
let point_in = Projective::<G::Config>::zero();
let point_out = Projective::<G::Config>::zero();
let scalar = G::ScalarField::rand(&mut rng);
let point_in =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_witness(cs.clone(), || {
Ok(point_in)
})?;
let point_out =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_input(cs.clone(), || {
Ok(point_out)
})?;
let scalar = NonNativeFieldVar::new_input(cs.clone(), || Ok(scalar))?;
let mul = point_in.scalar_mul_le(scalar.to_bits_le().unwrap().iter())?;
point_out.enforce_equal(&mul)?;
cs.is_satisfied()
}
#[test]
fn test_zero_point_scalar_mul() {
assert!(zero_point_scalar_mul_satisfied::<ark_bls12_381::G1Projective>().unwrap());
assert!(zero_point_scalar_mul_satisfied::<ark_pallas::Projective>().unwrap());
assert!(zero_point_scalar_mul_satisfied::<ark_mnt4_298::G1Projective>().unwrap());
assert!(zero_point_scalar_mul_satisfied::<ark_mnt6_298::G1Projective>().unwrap());
assert!(zero_point_scalar_mul_satisfied::<ark_bn254::G1Projective>().unwrap());
}
}

+ 119
- 0
src/groups/curves/short_weierstrass/non_zero_affine.rs

@ -130,6 +130,79 @@ where
}
}
/// Conditionally computes `(self + other) + self` or `(self + other) +
/// other` depending on the value of `cond`.
///
/// This follows the formulae from [\[ELM03\]](https://arxiv.org/abs/math/0208038).
#[tracing::instrument(target = "r1cs", skip(self))]
pub fn double_and_select_add_unchecked(
&self,
cond: &Boolean<<P::BaseField as Field>::BasePrimeField>,
other: &Self,
) -> Result<Self, SynthesisError> {
if [self].is_constant() || other.is_constant() {
// /!\ TODO: correct constant case /!\
self.double()?.add_unchecked(other)
} else {
// It's okay to use `unchecked` the precondition is that `self != ±other` (i.e.
// same logic as in `add_unchecked`)
let (x1, y1) = (&self.x, &self.y);
let (x2, y2) = (&other.x, &other.y);
// Calculate self + other:
// slope lambda := (y2 - y1)/(x2 - x1);
// x3 = lambda^2 - x1 - x2;
// y3 = lambda * (x1 - x3) - y1
let numerator = y2 - y1;
let denominator = x2 - x1;
let lambda_1 = numerator.mul_by_inverse_unchecked(&denominator)?;
let x3 = lambda_1.square()? - x1 - x2;
let x = &F::conditionally_select(&cond, &x1, &x2)?;
let y = &F::conditionally_select(&cond, &y1, &y2)?;
// Calculate final addition slope:
let lambda_2 =
(lambda_1 + y.double()?.mul_by_inverse_unchecked(&(&x3 - x))?).negate()?;
// Calculate (self + other) + (self or other):
let x4 = lambda_2.square()? - x - x3;
let y4 = lambda_2 * &(x - &x4) - y;
Ok(Self::new(x4, y4))
}
}
/// Triples `self`.
///
/// This follows the formulae from [\[ELM03\]](https://arxiv.org/abs/math/0208038).
#[tracing::instrument(target = "r1cs", skip(self))]
pub fn triple(&self) -> Result<Self, SynthesisError> {
if [self].is_constant() {
self.double()?.add_unchecked(self)
} else {
let (x1, y1) = (&self.x, &self.y);
let x1_sqr = x1.square()?;
// tangent lambda_1 := (3 * x1^2 + a) / (2 * y1);
// x3 = lambda_1^2 - 2x1
// y3 = lambda_1 * (x1 - x3) - y1
let numerator = x1_sqr.double()? + &x1_sqr + P::COEFF_A;
let denominator = y1.double()?;
// It's okay to use `unchecked` here, because the precondition of `double` is
// that self != zero.
let lambda_1 = numerator.mul_by_inverse_unchecked(&denominator)?;
let x3 = lambda_1.square()? - x1.double()?;
// Calculate final addition slope:
let lambda_2 =
(lambda_1 + y1.double()?.mul_by_inverse_unchecked(&(&x3 - x1))?).negate()?;
let x4 = lambda_2.square()? - x1 - x3;
let y4 = lambda_2 * &(x1 - &x4) - y1;
Ok(Self::new(x4, y4))
}
}
/// Doubles `self` in place.
#[tracing::instrument(target = "r1cs", skip(self))]
pub fn double_in_place(&mut self) -> Result<(), SynthesisError> {
@ -390,4 +463,50 @@ mod test_non_zero_affine {
assert!(cs.is_satisfied().unwrap());
}
#[test]
fn correctness_double_and_add_select() {
let cs = ConstraintSystem::<Fq>::new_ref();
let x = FpVar::Var(
AllocatedFp::<Fq>::new_witness(cs.clone(), || Ok(G1Config::GENERATOR.x)).unwrap(),
);
let y = FpVar::Var(
AllocatedFp::<Fq>::new_witness(cs.clone(), || Ok(G1Config::GENERATOR.y)).unwrap(),
);
// The following code tests `double_and_add`.
let sum_a = {
let a = ProjectiveVar::<G1Config, FpVar<Fq>>::new(
x.clone(),
y.clone(),
FpVar::Constant(Fq::one()),
);
let mut cur = a.clone();
cur.double_in_place().unwrap();
for _ in 1..10 {
cur.double_in_place().unwrap();
cur = cur + &a;
}
let sum = cur.value().unwrap().into_affine();
(sum.x, sum.y)
};
let sum_b = {
let a = NonZeroAffineVar::<G1Config, FpVar<Fq>>::new(x, y);
let mut cur = a.double().unwrap();
for _ in 1..10 {
cur = cur.double_and_add_unchecked(&a).unwrap();
}
(cur.x.value().unwrap(), cur.y.value().unwrap())
};
assert!(cs.is_satisfied().unwrap());
assert_eq!(sum_a.0, sum_b.0);
assert_eq!(sum_a.1, sum_b.1);
}
}

+ 176
- 0
src/groups/mod.rs

@ -107,6 +107,44 @@ pub trait CurveVar:
Ok(res)
}
/// Computes `bits * self`, where `bits` is a little-endian
/// `Boolean` representation of a scalar.
///
/// [Joye07](<https://www.iacr.org/archive/ches2007/47270135/47270135.pdf>), Alg.1.
#[tracing::instrument(target = "r1cs", skip(bits))]
fn scalar_mul_joye_le<'a>(
&self,
bits: impl Iterator<Item = &'a Boolean<ConstraintF>>,
) -> Result<Self, SynthesisError> {
// TODO: in the constant case we should call precomputed_scalar_mul_le,
// but rn there's a bug when doing this with TE curves.
// Computes the standard little-endian double-and-add algorithm
// (Algorithm 3.26, Guide to Elliptic Curve Cryptography)
let mut res = Self::zero();
let mut multiple = self.clone();
for bit in bits {
let tmp = res.clone() + &multiple;
res = bit.select(&tmp, &res)?;
multiple.double_in_place()?;
}
Ok(res)
}
/// Computes a `I1 * self + I2 * p` in place, where `I1` and `I2` are
/// `Boolean` *big-endian* representation of the scalars.
#[tracing::instrument(target = "r1cs", skip(bits1, bits2))]
fn joint_scalar_mul_be<'a>(
&self,
p: &Self,
bits1: impl Iterator<Item = &'a Boolean<ConstraintF>>,
bits2: impl Iterator<Item = &'a Boolean<ConstraintF>>,
) -> Result<Self, SynthesisError> {
let res1 = self.scalar_mul_le(bits1)?;
let res2 = p.scalar_mul_le(bits2)?;
Ok(res1 + res2)
}
/// Computes a `I * self` in place, where `I` is a `Boolean` *little-endian*
/// representation of the scalar.
///
@ -161,3 +199,141 @@ pub trait CurveVar:
Ok(result)
}
}
#[cfg(test)]
mod test_sw_arithmetic {
use crate::{
alloc::AllocVar,
eq::EqGadget,
fields::{fp::FpVar, nonnative::NonNativeFieldVar},
groups::{curves::short_weierstrass::ProjectiveVar, CurveVar},
ToBitsGadget,
};
use ark_ec::{
short_weierstrass::{Projective, SWCurveConfig},
CurveGroup,
};
use ark_ff::PrimeField;
use ark_relations::r1cs::{ConstraintSystem, Result};
use ark_std::UniformRand;
fn point_scalar_mul_satisfied<G>() -> Result<bool>
where
G: CurveGroup,
G::BaseField: PrimeField,
G::Config: SWCurveConfig,
{
let mut rng = ark_std::test_rng();
let cs = ConstraintSystem::new_ref();
let point_in = Projective::<G::Config>::rand(&mut rng);
let scalar = G::ScalarField::rand(&mut rng);
let point_out = point_in * scalar;
let point_in =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_witness(cs.clone(), || {
Ok(point_in)
})?;
let point_out =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_input(cs.clone(), || {
Ok(point_out)
})?;
let scalar = NonNativeFieldVar::new_input(cs.clone(), || Ok(scalar))?;
let mul = point_in.scalar_mul_le(scalar.to_bits_le().unwrap().iter())?;
point_out.enforce_equal(&mul)?;
println!("#r1cs for scalar_mul_le: {}", cs.num_constraints());
cs.is_satisfied()
}
fn point_scalar_mul_joye_satisfied<G>() -> Result<bool>
where
G: CurveGroup,
G::BaseField: PrimeField,
G::Config: SWCurveConfig,
{
let mut rng = ark_std::test_rng();
let cs = ConstraintSystem::new_ref();
let point_in = Projective::<G::Config>::rand(&mut rng);
let scalar = G::ScalarField::rand(&mut rng);
let point_out = point_in * scalar;
let point_in =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_witness(cs.clone(), || {
Ok(point_in)
})?;
let point_out =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_input(cs.clone(), || {
Ok(point_out)
})?;
let scalar = NonNativeFieldVar::new_input(cs.clone(), || Ok(scalar))?;
let mul = point_in.scalar_mul_joye_le(scalar.to_bits_le().unwrap().iter())?;
point_out.enforce_equal(&mul)?;
println!("#r1cs for scalar_mul_joye_le: {}", cs.num_constraints());
cs.is_satisfied()
}
fn point_joint_scalar_mul_satisfied<G>() -> Result<bool>
where
G: CurveGroup,
G::BaseField: PrimeField,
G::Config: SWCurveConfig,
{
let mut rng = ark_std::test_rng();
let cs = ConstraintSystem::new_ref();
let point_in1 = Projective::<G::Config>::rand(&mut rng);
let point_in2 = Projective::<G::Config>::rand(&mut rng);
let scalar1 = G::ScalarField::rand(&mut rng);
let scalar2 = G::ScalarField::rand(&mut rng);
let point_out = point_in1 * scalar1 + point_in2 * scalar2;
let point_in1 =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_witness(cs.clone(), || {
Ok(point_in1)
})?;
let point_in2 =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_witness(cs.clone(), || {
Ok(point_in2)
})?;
let point_out =
ProjectiveVar::<G::Config, FpVar<G::BaseField>>::new_input(cs.clone(), || {
Ok(point_out)
})?;
let scalar1 = NonNativeFieldVar::new_input(cs.clone(), || Ok(scalar1))?;
let scalar2 = NonNativeFieldVar::new_input(cs.clone(), || Ok(scalar2))?;
let res = point_in1.joint_scalar_mul_be(
&point_in2,
scalar1.to_bits_le().unwrap().iter(),
scalar2.to_bits_le().unwrap().iter(),
)?;
point_out.enforce_equal(&res)?;
println!("#r1cs for joint_scalar_mul: {}", cs.num_constraints());
cs.is_satisfied()
}
#[test]
fn test_point_scalar_mul() {
assert!(point_scalar_mul_satisfied::<ark_bn254::G1Projective>().unwrap());
}
#[test]
fn test_point_scalar_mul_joye() {
assert!(point_scalar_mul_joye_satisfied::<ark_bn254::G1Projective>().unwrap());
}
#[test]
fn test_point_joint_scalar_mul() {
assert!(point_joint_scalar_mul_satisfied::<ark_bn254::G1Projective>().unwrap());
}
}

Loading…
Cancel
Save