Refactor bit iteration infrastructure:

* `to_bits` -> `to_bits_le`
* `BitIterator` -> `BitIteratorLE` + `BitIteratorBE`
* `found_one`/`seen_one` -> `BitIteratorBE::without_leading_zeros`
This commit is contained in:
Pratyush Mishra
2020-08-28 11:39:38 -07:00
parent bce788419f
commit 6cca9327be
31 changed files with 398 additions and 455 deletions

View File

@@ -1,4 +1,4 @@
use algebra::{BitIterator, Field};
use algebra::{BitIteratorBE, Field};
use crate::{prelude::*, Assignment, Vec};
use core::borrow::Borrow;
@@ -373,14 +373,15 @@ impl<F: Field> Boolean<F> {
}
}
/// Asserts that this bit_gadget representation is "in
/// the field" when interpreted in big endian.
pub fn enforce_in_field(bits: &[Self]) -> Result<(), SynthesisError> {
// b = char() - 1
/// Enforces that `bits`, when interpreted as a integer, is less than `F::characteristic()`,
/// That is, interpret bits as a little-endian integer, and enforce that this integer
/// is "in the field F".
pub fn enforce_in_field_le(bits: &[Self]) -> Result<(), SynthesisError> {
// `bits` < F::characteristic() <==> `bits` <= F::characteristic() -1
let mut b = F::characteristic().to_vec();
assert_eq!(b[0] % 2, 1);
b[0] -= 1;
let run = Self::enforce_smaller_or_equal_than(bits, b)?;
b[0] -= 1; // This works, because the LSB is one, so there's no borrows.
let run = Self::enforce_smaller_or_equal_than_le(bits, b)?;
// We should always end in a "run" of zeros, because
// the characteristic is an odd prime. So, this should
@@ -390,57 +391,35 @@ impl<F: Field> Boolean<F> {
Ok(())
}
/// Asserts that this bit_gadget representation is smaller
/// or equal than the provided element
pub fn enforce_smaller_or_equal_than(
/// Enforces that `bits` is less than or equal to `element`,
/// when both are interpreted as (little-endian) integers.
pub fn enforce_smaller_or_equal_than_le<'a>(
bits: &[Self],
element: impl AsRef<[u64]>,
) -> Result<Vec<Self>, SynthesisError> {
let mut bits_iter = bits.iter();
let b: &[u64] = element.as_ref();
let mut bits_iter = bits.iter().rev(); // Iterate in big-endian
// Runs of ones in r
let mut last_run = Boolean::constant(true);
let mut current_run = vec![];
let mut found_one = false;
let mut element_num_bits = 0;
for _ in BitIteratorBE::without_leading_zeros(b) {
element_num_bits += 1;
}
let char_num_bits = {
let mut leading_zeros = 0;
let mut total_bits = 0;
let mut found_one = false;
for b in BitIterator::new(b.clone()) {
total_bits += 1;
if !b && !found_one {
leading_zeros += 1
}
if b {
found_one = true;
}
}
total_bits - leading_zeros
};
if bits.len() > char_num_bits {
let num_extra_bits = bits.len() - char_num_bits;
if bits.len() > element_num_bits {
let mut or_result = Boolean::constant(false);
for should_be_zero in &bits[0..num_extra_bits] {
for should_be_zero in &bits[element_num_bits..] {
or_result = or_result.or(should_be_zero)?;
let _ = bits_iter.next().unwrap();
}
or_result.enforce_equal(&Boolean::constant(false))?;
}
for b in BitIterator::new(b) {
// Skip over unset bits at the beginning
found_one |= b;
if !found_one {
continue;
}
let a = bits_iter.next().unwrap();
for (b, a) in BitIteratorBE::without_leading_zeros(b).zip(bits_iter.by_ref()) {
if b {
// This is part of a run of ones.
current_run.push(a.clone());
@@ -586,9 +565,8 @@ impl<F: Field> EqGadget<F> for Boolean<F> {
impl<F: Field> ToBytesGadget<F> for Boolean<F> {
/// Outputs `1u8` if `self` is true, and `0u8` otherwise.
fn to_bytes(&self) -> Result<Vec<UInt8<F>>, SynthesisError> {
let mut bits = vec![Boolean::constant(false); 7];
bits.push(self.clone());
bits.reverse();
let mut bits = vec![self.clone()];
bits.extend(vec![Boolean::constant(false); 7]);
let value = self.value().map(|val| val as u8).ok();
let byte = UInt8 { bits, value };
Ok(vec![byte])
@@ -655,7 +633,9 @@ impl<F: Field> CondSelectGadget<F> for Boolean<F> {
mod test {
use super::{AllocatedBit, Boolean};
use crate::prelude::*;
use algebra::{bls12_381::Fr, BitIterator, Field, One, PrimeField, UniformRand, Zero};
use algebra::{
bls12_381::Fr, BitIteratorBE, BitIteratorLE, Field, One, PrimeField, UniformRand, Zero,
};
use r1cs_core::{ConstraintSystem, Namespace, SynthesisError};
use rand::SeedableRng;
use rand_xorshift::XorShiftRng;
@@ -1331,17 +1311,58 @@ mod test {
Ok(())
}
#[test]
fn test_smaller_than_or_equal_to() -> Result<(), SynthesisError> {
let mut rng = XorShiftRng::seed_from_u64(1231275789u64);
for _ in 0..1000 {
let mut r = Fr::rand(&mut rng);
let mut s = Fr::rand(&mut rng);
if r > s {
core::mem::swap(&mut r, &mut s)
}
let cs = ConstraintSystem::<Fr>::new_ref();
let native_bits: Vec<_> = BitIteratorLE::new(r.into_repr()).collect();
let bits = Vec::new_witness(cs.clone(), || Ok(native_bits))?;
Boolean::enforce_smaller_or_equal_than_le(&bits, s.into_repr())?;
assert!(cs.is_satisfied().unwrap());
}
for _ in 0..1000 {
let r = Fr::rand(&mut rng);
if r == -Fr::one() {
continue;
}
let s = r + Fr::one();
let s2 = r.double();
let cs = ConstraintSystem::<Fr>::new_ref();
let native_bits: Vec<_> = BitIteratorLE::new(r.into_repr()).collect();
let bits = Vec::new_witness(cs.clone(), || Ok(native_bits))?;
Boolean::enforce_smaller_or_equal_than_le(&bits, s.into_repr())?;
if r < s2 {
Boolean::enforce_smaller_or_equal_than_le(&bits, s2.into_repr())?;
}
assert!(cs.is_satisfied().unwrap());
}
Ok(())
}
#[test]
fn test_enforce_in_field() -> Result<(), SynthesisError> {
{
let cs = ConstraintSystem::<Fr>::new_ref();
let mut bits = vec![];
for b in BitIterator::new(Fr::characteristic()).skip(1) {
for b in BitIteratorBE::new(Fr::characteristic()).skip(1) {
bits.push(Boolean::new_witness(cs.clone(), || Ok(b))?);
}
bits.reverse();
Boolean::enforce_in_field(&bits)?;
Boolean::enforce_in_field_le(&bits)?;
assert!(!cs.is_satisfied().unwrap());
}
@@ -1353,11 +1374,12 @@ mod test {
let cs = ConstraintSystem::<Fr>::new_ref();
let mut bits = vec![];
for b in BitIterator::new(r.into_repr()).skip(1) {
for b in BitIteratorBE::new(r.into_repr()).skip(1) {
bits.push(Boolean::new_witness(cs.clone(), || Ok(b))?);
}
bits.reverse();
Boolean::enforce_in_field(&bits)?;
Boolean::enforce_in_field_le(&bits)?;
assert!(cs.is_satisfied().unwrap());
}

View File

@@ -15,66 +15,63 @@ make_uint!(UInt32, 32, u32, uint32);
make_uint!(UInt64, 64, u64, uint64);
pub trait ToBitsGadget<F: Field> {
/// Outputs the canonical bit-wise representation of `self`.
/// Outputs the canonical little-endian bit-wise representation of `self`.
///
/// This is the correct default for 99% of use cases.
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError>;
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError>;
/// Outputs a possibly non-unique bit-wise representation of `self`.
///
/// If you're not absolutely certain that your usecase can get away with a
/// non-canonical representation, please use `self.to_bits(cs)` instead.
fn to_non_unique_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
self.to_bits()
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
self.to_bits_le()
}
}
impl<F: Field> ToBitsGadget<F> for Boolean<F> {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
Ok(vec![self.clone()])
}
}
impl<F: Field> ToBitsGadget<F> for [Boolean<F>] {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
/// Outputs `self`.
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
Ok(self.to_vec())
}
}
impl<F: Field> ToBitsGadget<F> for Vec<Boolean<F>> {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
Ok(self.clone())
}
}
impl<F: Field> ToBitsGadget<F> for UInt8<F> {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
Ok(self.into_bits_le())
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
Ok(self.bits.to_vec())
}
}
impl<F: Field> ToBitsGadget<F> for [UInt8<F>] {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let mut result = Vec::with_capacity(&self.len() * 8);
for byte in self {
result.extend_from_slice(&byte.into_bits_le());
}
Ok(result)
/// Interprets `self` as an integer, and outputs the little-endian
/// bit-wise decomposition of that integer.
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let bits = self.iter().flat_map(|b| &b.bits).cloned().collect();
Ok(bits)
}
}
impl<F: Field> ToBitsGadget<F> for Vec<UInt8<F>> {
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let mut result = Vec::with_capacity(&self.len() * 8);
for byte in self {
result.extend_from_slice(&byte.into_bits_le());
}
Ok(result)
impl<F: Field, T> ToBitsGadget<F> for Vec<T>
where
[T]: ToBitsGadget<F>,
{
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
self.as_slice().to_bits_le().map(|v| v.to_vec())
}
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
self.as_slice().to_non_unique_bits_le().map(|v| v.to_vec())
}
}
pub trait ToBytesGadget<F: Field> {
/// Outputs a canonical byte-wise representation of `self`.
/// Outputs a canonical, little-endian, byte decomposition of `self`.
///
/// This is the correct default for 99% of use cases.
fn to_bytes(&self) -> Result<Vec<UInt8<F>>, SynthesisError>;

View File

@@ -98,11 +98,7 @@ impl<F: Field> UInt8<F> {
let mut allocated_bits = Vec::new();
for field_element in field_elements.into_iter() {
let fe = AllocatedFp::new_input(cs.clone(), || Ok(field_element))?;
let mut fe_bits = fe.to_bits()?;
// FpGadget::to_bits outputs a big-endian binary representation of
// fe_gadget's value, so we have to reverse it to get the little-endian
// form.
fe_bits.reverse();
let fe_bits = fe.to_bits_le()?;
// Remove the most significant bit, because we know it should be zero
// because `values.to_field_elements()` only
@@ -113,19 +109,12 @@ impl<F: Field> UInt8<F> {
}
// Chunk up slices of 8 bit into bytes.
Ok(allocated_bits[0..8 * values_len]
Ok(allocated_bits[0..(8 * values_len)]
.chunks(8)
.map(Self::from_bits_le)
.collect())
}
/// Turns this `UInt8` into its little-endian byte order representation.
/// LSB-first means that we can easily get the corresponding field element
/// via double and add.
pub fn into_bits_le(&self) -> Vec<Boolean<F>> {
self.bits.to_vec()
}
/// Converts a little-endian byte order representation of bits into a
/// `UInt8`.
pub fn from_bits_le(bits: &[Boolean<F>]) -> Self {
@@ -134,25 +123,10 @@ impl<F: Field> UInt8<F> {
let bits = bits.to_vec();
let mut value = Some(0u8);
for b in bits.iter().rev() {
value.as_mut().map(|v| *v <<= 1);
match *b {
Boolean::Constant(b) => {
value.as_mut().map(|v| *v |= u8::from(b));
}
Boolean::Is(ref b) => match b.value() {
Ok(b) => {
value.as_mut().map(|v| *v |= u8::from(b));
}
Err(_) => value = None,
},
Boolean::Not(ref b) => match b.value() {
Ok(b) => {
value.as_mut().map(|v| *v |= u8::from(!b));
}
Err(_) => value = None,
},
for (i, b) in bits.iter().enumerate() {
value = match b.value().ok() {
Some(b) => value.map(|v| v + (u8::from(b) << i)),
None => None,
}
}
@@ -241,7 +215,7 @@ mod test {
let cs = ConstraintSystem::<Fr>::new_ref();
let byte_val = 0b01110001;
let byte = UInt8::new_witness(cs.ns("alloc value"), || Ok(byte_val)).unwrap();
let bits = byte.into_bits_le();
let bits = byte.to_bits_le()?;
for (i, bit) in bits.iter().enumerate() {
assert_eq!(bit.value()?, (byte_val >> i) & 1 == 1)
}
@@ -253,10 +227,17 @@ mod test {
let cs = ConstraintSystem::<Fr>::new_ref();
let byte_vals = (64u8..128u8).collect::<Vec<_>>();
let bytes = UInt8::new_input_vec(cs.ns("alloc value"), &byte_vals).unwrap();
dbg!(bytes.value())?;
for (native, variable) in byte_vals.into_iter().zip(bytes) {
let bits = variable.into_bits_le();
let bits = variable.to_bits_le()?;
for (i, bit) in bits.iter().enumerate() {
assert_eq!(bit.value()?, (native >> i) & 1 == 1)
assert_eq!(
bit.value()?,
(native >> i) & 1 == 1,
"native value {}: bit {:?}",
native,
i
)
}
}
Ok(())
@@ -280,7 +261,7 @@ mod test {
}
}
let expected_to_be_same = val.into_bits_le();
let expected_to_be_same = val.to_bits_le()?;
for x in v.iter().zip(expected_to_be_same.iter()) {
match x {

View File

@@ -376,19 +376,19 @@ where
for<'a> &'a BF: FieldOpsBounds<'a, P::BaseField, BF>,
P: CubicExtVarParams<BF>,
{
fn to_bits(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_bits()?;
let mut c1 = self.c1.to_bits()?;
let mut c2 = self.c2.to_bits()?;
fn to_bits_le(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_bits_le()?;
let mut c1 = self.c1.to_bits_le()?;
let mut c2 = self.c2.to_bits_le()?;
c0.append(&mut c1);
c0.append(&mut c2);
Ok(c0)
}
fn to_non_unique_bits(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_non_unique_bits()?;
let mut c1 = self.c1.to_non_unique_bits()?;
let mut c2 = self.c2.to_non_unique_bits()?;
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_non_unique_bits_le()?;
let mut c1 = self.c1.to_non_unique_bits_le()?;
let mut c2 = self.c2.to_non_unique_bits_le()?;
c0.append(&mut c1);
c0.append(&mut c2);
Ok(c0)

View File

@@ -95,8 +95,10 @@ impl<F: PrimeField> FpVar<F> {
pub fn enforce_smaller_or_equal_than_mod_minus_one_div_two(
&self,
) -> Result<(), SynthesisError> {
let _ = Boolean::enforce_smaller_or_equal_than(
&self.to_bits()?,
// It's okay to use `to_non_unique_bits` bits here because we're enforcing
// self <= (p-1)/2, which implies self < p.
let _ = Boolean::enforce_smaller_or_equal_than_le(
&self.to_non_unique_bits_le()?,
F::modulus_minus_one_div_two(),
)?;
Ok(())
@@ -114,7 +116,12 @@ impl<F: PrimeField> FpVar<F> {
/// assumes `self` and `other` are `<= (p-1)/2` and does not generate constraints
/// to verify that.
fn is_smaller_than_unchecked(&self, other: &FpVar<F>) -> Result<Boolean<F>, SynthesisError> {
Ok((self - other).double()?.to_bits()?.last().unwrap().clone())
Ok((self - other)
.double()?
.to_bits_le()?
.first()
.unwrap()
.clone())
}
/// Helper function to enforce `self < other`. This function verifies `self` and `other`
@@ -186,6 +193,7 @@ mod test {
}
assert!(cs.is_satisfied().unwrap());
}
println!("Finished with satisfaction tests");
for _i in 0..10 {
let cs = ConstraintSystem::<Fr>::new_ref();

View File

@@ -1,10 +1,10 @@
use algebra::{bytes::ToBytes, FpParameters, PrimeField};
use algebra::{BigInteger, FpParameters, PrimeField};
use r1cs_core::{lc, ConstraintSystemRef, LinearCombination, Namespace, SynthesisError, Variable};
use core::borrow::Borrow;
use crate::fields::{FieldOpsBounds, FieldVar};
use crate::{boolean::AllocatedBit, prelude::*, Assignment, Vec};
use crate::{prelude::*, Assignment, Vec};
pub mod cmp;
@@ -338,48 +338,41 @@ impl<F: PrimeField> AllocatedFp<F> {
impl<F: PrimeField> ToBitsGadget<F> for AllocatedFp<F> {
/// Outputs the unique bit-wise decomposition of `self` in *big-endian*
/// form.
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let bits = self.to_non_unique_bits()?;
Boolean::enforce_in_field(&bits)?;
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let bits = self.to_non_unique_bits_le()?;
Boolean::enforce_in_field_le(&bits)?;
Ok(bits)
}
fn to_non_unique_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
let cs = self.cs.clone();
let num_bits = F::Params::MODULUS_BITS;
use algebra::BitIterator;
let bit_values = match self.value {
Some(value) => {
let mut field_char = BitIterator::new(F::characteristic());
let mut tmp = Vec::with_capacity(num_bits as usize);
let mut found_one = false;
for b in BitIterator::new(value.into_repr()) {
// Skip leading bits
found_one |= field_char.next().unwrap();
if !found_one {
continue;
}
tmp.push(Some(b));
}
assert_eq!(tmp.len(), num_bits as usize);
tmp
}
None => vec![None; num_bits as usize],
use algebra::BitIteratorBE;
let mut bits = if let Some(value) = self.value {
let field_char = BitIteratorBE::new(F::characteristic());
let bits: Vec<_> = BitIteratorBE::new(value.into_repr())
.zip(field_char)
.skip_while(|(_, c)| !c)
.map(|(b, _)| Some(b))
.collect();
assert_eq!(bits.len(), F::Params::MODULUS_BITS as usize);
bits
} else {
vec![None; F::Params::MODULUS_BITS as usize]
};
let mut bits = vec![];
for b in bit_values {
bits.push(AllocatedBit::new_witness(cs.clone(), || b.get())?);
}
// Convert to little-endian
bits.reverse();
let bits: Vec<_> = bits
.into_iter()
.map(|b| Boolean::new_witness(cs.clone(), || b.get()))
.collect::<Result<_, _>>()?;
let mut lc = LinearCombination::zero();
let mut coeff = F::one();
for bit in bits.iter().rev() {
lc += (coeff, bit.variable());
for bit in bits.iter() {
lc = &lc + bit.lc() * coeff;
coeff.double_in_place();
}
@@ -388,7 +381,7 @@ impl<F: PrimeField> ToBitsGadget<F> for AllocatedFp<F> {
cs.enforce_constraint(lc!(), lc!(), lc)?;
Ok(bits.into_iter().map(Boolean::from).collect())
Ok(bits)
}
}
@@ -396,52 +389,26 @@ impl<F: PrimeField> ToBytesGadget<F> for AllocatedFp<F> {
/// Outputs the unique byte decomposition of `self` in *little-endian*
/// form.
fn to_bytes(&self) -> Result<Vec<UInt8<F>>, SynthesisError> {
let bytes = self.to_non_unique_bytes()?;
Boolean::enforce_in_field(
&bytes.iter()
.flat_map(|b| b.into_bits_le())
// This reverse maps the bits into big-endian form, as required by `enforce_in_field`.
.rev()
.collect::<Vec<_>>(),
)?;
let num_bits = F::BigInt::NUM_LIMBS * 64;
let mut bits = self.to_bits_le()?;
let remainder = core::iter::repeat(Boolean::constant(false)).take(num_bits - bits.len());
bits.extend(remainder);
let bytes = bits
.chunks(8)
.map(|chunk| UInt8::from_bits_le(chunk))
.collect();
Ok(bytes)
}
fn to_non_unique_bytes(&self) -> Result<Vec<UInt8<F>>, SynthesisError> {
let cs = self.cs.clone();
let byte_values = match self.value {
Some(value) => to_bytes![&value.into_repr()]
.unwrap()
.into_iter()
.map(Some)
.collect::<Vec<_>>(),
None => {
let default = F::default();
let default_len = to_bytes![&default].unwrap().len();
vec![None; default_len]
}
};
let bytes = UInt8::new_witness_vec(cs.clone(), &byte_values)?;
let mut lc = LinearCombination::zero();
let mut coeff = F::one();
for bit in bytes.iter().flat_map(|b| b.bits.clone()) {
match bit {
Boolean::Is(bit) => {
lc += (coeff, bit.variable());
coeff.double_in_place();
}
Boolean::Constant(_) | Boolean::Not(_) => unreachable!(),
}
}
lc = lc - &self.variable;
cs.enforce_constraint(lc!(), lc!(), lc)?;
let num_bits = F::BigInt::NUM_LIMBS * 64;
let mut bits = self.to_non_unique_bits_le()?;
let remainder = core::iter::repeat(Boolean::constant(false)).take(num_bits - bits.len());
bits.extend(remainder);
let bytes = bits
.chunks(8)
.map(|chunk| UInt8::from_bits_le(chunk))
.collect();
Ok(bytes)
}
}
@@ -806,19 +773,20 @@ impl<F: PrimeField> EqGadget<F> for FpVar<F> {
}
impl<F: PrimeField> ToBitsGadget<F> for FpVar<F> {
/// Outputs the unique bit-wise decomposition of `self` in *big-endian*
/// form.
fn to_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
fn to_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
match self {
Self::Constant(c) => UInt8::constant_vec(&to_bytes![c].unwrap()).to_bits(),
Self::Var(v) => v.to_bits(),
Self::Constant(_) => self.to_non_unique_bits_le(),
Self::Var(v) => v.to_bits_le(),
}
}
fn to_non_unique_bits(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<F>>, SynthesisError> {
use algebra::BitIteratorLE;
match self {
Self::Constant(c) => UInt8::constant_vec(&to_bytes![c].unwrap()).to_non_unique_bits(),
Self::Var(v) => v.to_non_unique_bits(),
Self::Constant(c) => Ok(BitIteratorLE::without_trailing_zeros(&c.into_repr())
.map(Boolean::constant)
.collect::<Vec<_>>()),
Self::Var(v) => v.to_non_unique_bits_le(),
}
}
}

View File

@@ -1,4 +1,4 @@
use algebra::{prelude::*, BitIterator};
use algebra::{prelude::*, BitIteratorBE};
use core::{
fmt::Debug,
ops::{Add, AddAssign, Mul, MulAssign, Sub, SubAssign},
@@ -132,37 +132,28 @@ pub trait FieldVar<F: Field, ConstraintF: Field>:
Ok(self)
}
/// Accepts as input a list of bits which, when interpreted in little-endian
/// form, are a scalar.
//
// TODO: check that the input really should be in little-endian or not...
fn pow(&self, bits: &[Boolean<ConstraintF>]) -> Result<Self, SynthesisError> {
/// Comptues `self^bits`, where `bits` is a *little-endian* bit-wise decomposition
/// of the exponent.
fn pow_le(&self, bits: &[Boolean<ConstraintF>]) -> Result<Self, SynthesisError> {
let mut res = Self::one();
for bit in bits.iter() {
res.square_in_place()?;
let tmp = res.clone() * self;
let mut power = self.clone();
for bit in bits {
let tmp = res.clone() * &power;
res = bit.select(&tmp, &res)?;
power.square_in_place()?;
}
Ok(res)
}
/// Computes `self^S`, where S is interpreted as an integer.
fn pow_by_constant<S: AsRef<[u64]>>(&self, exp: S) -> Result<Self, SynthesisError> {
let mut res = self.clone();
let mut found_one = false;
for bit in BitIterator::new(exp) {
if found_one {
res = res.square()?;
}
if bit {
if found_one {
res *= self;
}
found_one = true;
let mut res = Self::one();
for i in BitIteratorBE::without_leading_zeros(exp) {
res.square_in_place()?;
if i {
res *= self;
}
}
Ok(res)
}
}
@@ -173,7 +164,7 @@ pub(crate) mod tests {
use rand_xorshift::XorShiftRng;
use crate::{fields::*, Vec};
use algebra::{test_rng, BitIterator, Field, UniformRand};
use algebra::{test_rng, BitIteratorLE, Field, UniformRand};
use r1cs_core::{ConstraintSystem, SynthesisError};
#[allow(dead_code)]
@@ -303,13 +294,14 @@ pub(crate) mod tests {
assert_eq!(a_b_inv.value()?, a_native * b_native.inverse().unwrap());
// a * a * a = a^3
let bits = BitIterator::new([0x3])
let bits = BitIteratorLE::without_trailing_zeros([3u64])
.map(Boolean::constant)
.collect::<Vec<_>>();
assert_eq!(a_native.pow([0x3]), a.pow(&bits)?.value()?);
assert_eq!(a_native.pow([0x3]), a.pow_le(&bits)?.value()?);
// a * a * a = a^3
assert_eq!(a_native.pow([0x3]), a.pow_by_constant(&[3])?.value()?);
assert_eq!(a_native.pow([0x3]), a.pow_by_constant(&[0x3])?.value()?);
assert!(cs.is_satisfied().unwrap());
// a * a * a = a^3
let mut constants = [F::zero(); 4];
@@ -322,12 +314,34 @@ pub(crate) mod tests {
];
let lookup_result = AF::two_bit_lookup(&bits, constants.as_ref())?;
assert_eq!(lookup_result.value()?, constants[2]);
assert!(cs.is_satisfied().unwrap());
let negone: F = UniformRand::rand(&mut test_rng());
let f = F::from(1u128 << 64);
let f_bits = algebra::BitIteratorLE::new(&[0u64, 1u64]).collect::<Vec<_>>();
let fv = AF::new_witness(cs.ns("alloc u128"), || Ok(f))?;
assert_eq!(fv.to_bits_le()?.value().unwrap()[..128], f_bits[..128]);
assert!(cs.is_satisfied().unwrap());
let n = AF::new_witness(cs.ns("alloc new var"), || Ok(negone)).unwrap();
let _ = n.to_bytes()?;
let _ = n.to_non_unique_bytes()?;
let r_native: F = UniformRand::rand(&mut test_rng());
let r = AF::new_witness(cs.ns("r_native"), || Ok(r_native)).unwrap();
let _ = r.to_non_unique_bits_le()?;
assert!(cs.is_satisfied().unwrap());
let _ = r.to_bits_le()?;
assert!(cs.is_satisfied().unwrap());
let bytes = r.to_non_unique_bytes()?;
assert_eq!(
algebra::to_bytes!(r_native).unwrap(),
bytes.value().unwrap()
);
assert!(cs.is_satisfied().unwrap());
let bytes = r.to_bytes()?;
assert_eq!(
algebra::to_bytes!(r_native).unwrap(),
bytes.value().unwrap()
);
assert!(cs.is_satisfied().unwrap());
let ab_false = &a + (AF::from(Boolean::Constant(false)) * b_native);
assert_eq!(ab_false.value()?, a_native);

View File

@@ -379,16 +379,16 @@ where
for<'b> &'b BF: FieldOpsBounds<'b, P::BaseField, BF>,
P: QuadExtVarParams<BF>,
{
fn to_bits(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_bits()?;
let mut c1 = self.c1.to_bits()?;
fn to_bits_le(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_bits_le()?;
let mut c1 = self.c1.to_bits_le()?;
c0.append(&mut c1);
Ok(c0)
}
fn to_non_unique_bits(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_non_unique_bits()?;
let mut c1 = self.c1.to_non_unique_bits()?;
fn to_non_unique_bits_le(&self) -> Result<Vec<Boolean<P::BasePrimeField>>, SynthesisError> {
let mut c0 = self.c0.to_non_unique_bits_le()?;
let mut c1 = self.c1.to_non_unique_bits_le()?;
c0.append(&mut c1);
Ok(c0)
}

View File

@@ -4,7 +4,7 @@ use algebra::{
short_weierstrass_jacobian::GroupAffine,
},
fields::Field,
BitIterator, One,
BitIteratorBE, One,
};
use r1cs_core::SynthesisError;
@@ -164,7 +164,7 @@ impl<P: Bls12Parameters> G2PreparedVar<P> {
let mut ell_coeffs = vec![];
let mut r = q.clone();
for i in BitIterator::new(P::X).skip(1) {
for i in BitIteratorBE::new(P::X).skip(1) {
ell_coeffs.push(Self::double(&mut r, &two_inv)?);
if i {

View File

@@ -3,7 +3,7 @@ use algebra::{
short_weierstrass_jacobian::{GroupAffine as SWAffine, GroupProjective as SWProjective},
SWModelParameters,
},
AffineCurve, BigInteger, BitIterator, Field, One, PrimeField, ProjectiveCurve, Zero,
AffineCurve, BigInteger, BitIteratorBE, Field, One, PrimeField, ProjectiveCurve, Zero,
};
use core::{borrow::Borrow, marker::PhantomData};
use r1cs_core::{ConstraintSystemRef, Namespace, SynthesisError};
@@ -255,22 +255,12 @@ where
fn enforce_prime_order(&self) -> Result<(), SynthesisError> {
let r_minus_1 = (-P::ScalarField::one()).into_repr();
let mut seen_one = false;
let mut result = Self::zero();
for b in BitIterator::new(r_minus_1) {
let old_seen_one = seen_one;
if seen_one {
result.double_in_place()?;
} else {
seen_one = b;
}
for b in BitIteratorBE::without_leading_zeros(r_minus_1) {
result.double_in_place()?;
if b {
result = if old_seen_one {
result + self
} else {
self.clone()
};
result += self;
}
}
self.negate()?.enforce_equal(&result)?;
@@ -527,10 +517,12 @@ where
power_of_2 += 1;
}
let cofactor_weight = BitIterator::new(cofactor.as_slice()).filter(|b| *b).count();
let cofactor_weight = BitIteratorBE::new(cofactor.as_slice())
.filter(|b| *b)
.count();
let modulus_minus_1 = (-P::ScalarField::one()).into_repr(); // r - 1
let modulus_minus_1_weight =
BitIterator::new(modulus_minus_1).filter(|b| *b).count();
BitIteratorBE::new(modulus_minus_1).filter(|b| *b).count();
// We pick the most efficient method of performing the prime order check:
// If the cofactor has lower hamming weight than the scalar field's modulus,
@@ -546,7 +538,10 @@ where
|| f().map(|g| g.borrow().into_affine().mul_by_cofactor_inv().into()),
mode,
)?;
(ge, BitIterator::new(cofactor.as_slice()))
(
ge,
BitIteratorBE::without_leading_zeros(cofactor.as_slice()),
)
} else {
let ge = Self::new_variable_omit_prime_order_check(
cs.ns("Witness without subgroup check with `r` check"),
@@ -555,37 +550,31 @@ where
let g = g.into_affine();
let mut power_of_two = P::ScalarField::one().into_repr();
power_of_two.muln(power_of_2);
let power_of_two_inv =
P::ScalarField::from(power_of_two).inverse().unwrap();
let power_of_two_inv = P::ScalarField::from_repr(power_of_two)
.and_then(|n| n.inverse())
.unwrap();
g.mul(power_of_two_inv)
})
},
mode,
)?;
(ge, BitIterator::new(modulus_minus_1.as_ref()))
(
ge,
BitIteratorBE::without_leading_zeros(modulus_minus_1.as_ref()),
)
};
// Remove the even part of the cofactor
for _ in 0..power_of_2 {
ge.double_in_place()?;
}
let mut seen_one = false;
let mut result = Self::zero();
for b in iter {
let old_seen_one = seen_one;
if seen_one {
result.double_in_place()?;
} else {
seen_one = b;
}
result.double_in_place()?;
if b {
result = if old_seen_one {
result + &ge
} else {
ge.clone()
};
result += &ge
}
}
if cofactor_weight < modulus_minus_1_weight {
@@ -616,23 +605,23 @@ where
F: FieldVar<P::BaseField, <P::BaseField as Field>::BasePrimeField>,
for<'a> &'a F: FieldOpsBounds<'a, P::BaseField, F>,
{
fn to_bits(
fn to_bits_le(
&self,
) -> Result<Vec<Boolean<<P::BaseField as Field>::BasePrimeField>>, SynthesisError> {
let g = self.to_affine()?;
let mut bits = g.x.to_bits()?;
let y_bits = g.y.to_bits()?;
let mut bits = g.x.to_bits_le()?;
let y_bits = g.y.to_bits_le()?;
bits.extend_from_slice(&y_bits);
bits.push(g.infinity);
Ok(bits)
}
fn to_non_unique_bits(
fn to_non_unique_bits_le(
&self,
) -> Result<Vec<Boolean<<P::BaseField as Field>::BasePrimeField>>, SynthesisError> {
let g = self.to_affine()?;
let mut bits = g.x.to_non_unique_bits()?;
let y_bits = g.y.to_non_unique_bits()?;
let mut bits = g.x.to_non_unique_bits_le()?;
let y_bits = g.y.to_non_unique_bits_le()?;
bits.extend_from_slice(&y_bits);
bits.push(g.infinity);
Ok(bits)
@@ -679,7 +668,7 @@ where
for<'a> &'a GG: GroupOpsBounds<'a, SWProjective<P>, GG>,
{
use crate::prelude::*;
use algebra::{test_rng, Group, UniformRand};
use algebra::{test_rng, BitIteratorLE, Group, UniformRand};
use r1cs_core::ConstraintSystem;
crate::groups::test::group_test::<SWProjective<P>, _, GG>()?;
@@ -739,11 +728,9 @@ where
let native_result = aa.into_affine().mul(scalar);
let native_result = native_result.into_affine();
let mut scalar: Vec<bool> = BitIterator::new(scalar.into_repr()).collect();
// Get the scalar bits into little-endian form.
scalar.reverse();
let scalar: Vec<bool> = BitIteratorLE::new(scalar.into_repr()).collect();
let input: Vec<Boolean<_>> = Vec::new_witness(cs.ns("bits"), || Ok(scalar)).unwrap();
let result = gadget_a.mul_bits(input.iter())?;
let result = gadget_a.scalar_mul_le(input.iter())?;
let result_val = result.value()?.into_affine();
assert_eq!(
result_val, native_result,

View File

@@ -3,7 +3,7 @@ use algebra::{
twisted_edwards_extended::{GroupAffine as TEAffine, GroupProjective as TEProjective},
AffineCurve, MontgomeryModelParameters, ProjectiveCurve, TEModelParameters,
},
BigInteger, BitIterator, Field, One, PrimeField, Zero,
BigInteger, BitIteratorBE, Field, One, PrimeField, Zero,
};
use r1cs_core::{ConstraintSystemRef, Namespace, SynthesisError};
@@ -328,22 +328,12 @@ where
fn enforce_prime_order(&self) -> Result<(), SynthesisError> {
let r_minus_1 = (-P::ScalarField::one()).into_repr();
let mut seen_one = false;
let mut result = Self::zero();
for b in BitIterator::new(r_minus_1) {
let old_seen_one = seen_one;
if seen_one {
result.double_in_place()?;
} else {
seen_one = b;
}
for b in BitIteratorBE::without_leading_zeros(r_minus_1) {
result.double_in_place()?;
if b {
result = if old_seen_one {
result + self
} else {
self.clone()
};
result += self;
}
}
self.negate()?.enforce_equal(&result)?;
@@ -398,7 +388,7 @@ where
Ok(Self::new(self.x.negate()?, self.y.clone()))
}
fn precomputed_base_scalar_mul<'a, I, B>(
fn precomputed_base_scalar_mul_le<'a, I, B>(
&mut self,
scalar_bits_with_base_powers: I,
) -> Result<(), SynthesisError>
@@ -477,7 +467,7 @@ where
acc_power += base_power;
}
let bits = bits.borrow().to_bits()?;
let bits = bits.borrow().to_bits_le()?;
if bits.len() != CHUNK_SIZE {
return Err(SynthesisError::Unsatisfiable);
}
@@ -552,10 +542,12 @@ where
power_of_2 += 1;
}
let cofactor_weight = BitIterator::new(cofactor.as_slice()).filter(|b| *b).count();
let cofactor_weight = BitIteratorBE::new(cofactor.as_slice())
.filter(|b| *b)
.count();
let modulus_minus_1 = (-P::ScalarField::one()).into_repr(); // r - 1
let modulus_minus_1_weight =
BitIterator::new(modulus_minus_1).filter(|b| *b).count();
BitIteratorBE::new(modulus_minus_1).filter(|b| *b).count();
// We pick the most efficient method of performing the prime order check:
// If the cofactor has lower hamming weight than the scalar field's modulus,
@@ -571,7 +563,10 @@ where
|| f().map(|g| g.borrow().into_affine().mul_by_cofactor_inv().into()),
mode,
)?;
(ge, BitIterator::new(cofactor.as_slice()))
(
ge,
BitIteratorBE::without_leading_zeros(cofactor.as_slice()),
)
} else {
let ge = Self::new_variable_omit_prime_order_check(
cs.ns("Witness without subgroup check with `r` check"),
@@ -580,37 +575,30 @@ where
let g = g.into_affine();
let mut power_of_two = P::ScalarField::one().into_repr();
power_of_two.muln(power_of_2);
let power_of_two_inv =
P::ScalarField::from(power_of_two).inverse().unwrap();
let power_of_two_inv = P::ScalarField::from_repr(power_of_two)
.and_then(|n| n.inverse())
.unwrap();
g.mul(power_of_two_inv)
})
},
mode,
)?;
(ge, BitIterator::new(modulus_minus_1.as_ref()))
(
ge,
BitIteratorBE::without_leading_zeros(modulus_minus_1.as_ref()),
)
};
// Remove the even part of the cofactor
for _ in 0..power_of_2 {
ge.double_in_place()?;
}
let mut seen_one = false;
let mut result = Self::zero();
for b in iter {
let old_seen_one = seen_one;
if seen_one {
result.double_in_place()?;
} else {
seen_one = b;
}
result.double_in_place()?;
if b {
result = if old_seen_one {
result + &ge
} else {
ge.clone()
};
result += &ge;
}
}
if cofactor_weight < modulus_minus_1_weight {
@@ -829,20 +817,20 @@ where
F: FieldVar<P::BaseField, <P::BaseField as Field>::BasePrimeField>,
for<'b> &'b F: FieldOpsBounds<'b, P::BaseField, F>,
{
fn to_bits(
fn to_bits_le(
&self,
) -> Result<Vec<Boolean<<P::BaseField as Field>::BasePrimeField>>, SynthesisError> {
let mut x_bits = self.x.to_bits()?;
let y_bits = self.y.to_bits()?;
let mut x_bits = self.x.to_bits_le()?;
let y_bits = self.y.to_bits_le()?;
x_bits.extend_from_slice(&y_bits);
Ok(x_bits)
}
fn to_non_unique_bits(
fn to_non_unique_bits_le(
&self,
) -> Result<Vec<Boolean<<P::BaseField as Field>::BasePrimeField>>, SynthesisError> {
let mut x_bits = self.x.to_non_unique_bits()?;
let y_bits = self.y.to_non_unique_bits()?;
let mut x_bits = self.x.to_non_unique_bits_le()?;
let y_bits = self.y.to_non_unique_bits_le()?;
x_bits.extend_from_slice(&y_bits);
Ok(x_bits)
@@ -884,7 +872,7 @@ where
for<'a> &'a GG: GroupOpsBounds<'a, TEProjective<P>, GG>,
{
use crate::prelude::*;
use algebra::{test_rng, Group, UniformRand};
use algebra::{test_rng, BitIteratorLE, Group, UniformRand};
use r1cs_core::ConstraintSystem;
crate::groups::test::group_test::<TEProjective<P>, _, GG>()?;
@@ -944,11 +932,9 @@ where
let native_result = AffineCurve::mul(&aa.into_affine(), scalar);
let native_result = native_result.into_affine();
let mut scalar: Vec<bool> = BitIterator::new(scalar.into_repr()).collect();
// Get the scalar bits into little-endian form.
scalar.reverse();
let scalar: Vec<bool> = BitIteratorLE::new(scalar.into_repr()).collect();
let input: Vec<Boolean<_>> = Vec::new_witness(cs.ns("bits"), || Ok(scalar)).unwrap();
let result = gadget_a.mul_bits(input.iter())?;
let result = gadget_a.scalar_mul_le(input.iter())?;
let result_val = result.value()?.into_affine();
assert_eq!(
result_val, native_result,

View File

@@ -72,24 +72,28 @@ pub trait CurveVar<C: ProjectiveCurve, ConstraintF: Field>:
fn negate(&self) -> Result<Self, SynthesisError>;
/// Inputs must be specified in *little-endian* form.
/// If the addition law is incomplete for the identity element,
/// `result` must not be the identity element.
fn mul_bits<'a>(
/// Computes `bits * self`, where `bits` is a little-endian
/// `Boolean` representation of a scalar.
fn scalar_mul_le<'a>(
&self,
bits: impl Iterator<Item = &'a Boolean<ConstraintF>>,
) -> Result<Self, SynthesisError> {
let mut power = self.clone();
let mut result = Self::zero();
let mut res = Self::zero();
let mut multiple = self.clone();
for bit in bits {
let new_encoded = result.clone() + &power;
result = bit.borrow().select(&new_encoded, &result)?;
power.double_in_place()?;
let tmp = res.clone() + &multiple;
res = bit.select(&tmp, &res)?;
multiple.double_in_place()?;
}
Ok(result)
Ok(res)
}
fn precomputed_base_scalar_mul<'a, I, B>(
/// Computes a `I * self` in place, where `I` is a `Boolean` *little-endian*
/// representation of the scalar.
///
/// The base powers are precomputed power-of-two multiples of a single
/// base.
fn precomputed_base_scalar_mul_le<'a, I, B>(
&mut self,
scalar_bits_with_base_powers: I,
) -> Result<(), SynthesisError>
@@ -117,7 +121,9 @@ pub trait CurveVar<C: ProjectiveCurve, ConstraintF: Field>:
Err(SynthesisError::AssignmentMissing)
}
fn precomputed_base_multiscalar_mul<'a, T, I, B>(
/// Computes a `\sum I_j * B_j`, where `I_j` is a `Boolean`
/// representation of the j-th scalar.
fn precomputed_base_multiscalar_mul_le<'a, T, I, B>(
bases: &[B],
scalars: I,
) -> Result<Self, SynthesisError>
@@ -130,8 +136,8 @@ pub trait CurveVar<C: ProjectiveCurve, ConstraintF: Field>:
// Compute ∏(h_i^{m_i}) for all i.
for (bits, base_powers) in scalars.zip(bases) {
let base_powers = base_powers.borrow();
let bits = bits.to_bits()?;
result.precomputed_base_scalar_mul(bits.iter().zip(base_powers))?;
let bits = bits.to_bits_le()?;
result.precomputed_base_scalar_mul_le(bits.iter().zip(base_powers))?;
}
Ok(result)
}
@@ -201,7 +207,9 @@ mod test {
assert_eq!(b2.value()?, b_b.value()?);
let _ = a.to_bytes()?;
assert!(cs.is_satisfied().unwrap());
let _ = a.to_non_unique_bytes()?;
assert!(cs.is_satisfied().unwrap());
let _ = b.to_bytes()?;
let _ = b.to_non_unique_bytes()?;

View File

@@ -8,7 +8,7 @@ use crate::{
};
use algebra::{
curves::bls12::{Bls12, Bls12Parameters, TwistType},
fields::BitIterator,
fields::BitIteratorBE,
};
use core::marker::PhantomData;
@@ -75,7 +75,7 @@ impl<P: Bls12Parameters> PG<Bls12<P>, P::Fp> for PairingVar<P> {
}
let mut f = Self::GTVar::one();
for i in BitIterator::new(P::X).skip(1) {
for i in BitIteratorBE::new(P::X).skip(1) {
f.square_in_place()?;
for &mut (p, ref mut coeffs) in pairs.iter_mut() {

View File

@@ -11,7 +11,7 @@ use crate::{
};
use algebra::{
curves::mnt4::{MNT4Parameters, MNT4},
fields::BitIterator,
fields::BitIteratorBE,
};
use core::marker::PhantomData;
@@ -100,18 +100,9 @@ impl<P: MNT4Parameters> PairingVar<P> {
let mut dbl_idx: usize = 0;
let mut add_idx: usize = 0;
let mut found_one = false;
for bit in BitIterator::new(P::ATE_LOOP_COUNT) {
// code below gets executed for all bits (EXCEPT the MSB itself) of
// mnt6_param_p (skipping leading zeros) in MSB to LSB order
if !found_one && bit {
found_one = true;
continue;
} else if !found_one {
continue;
}
// code below gets executed for all bits (EXCEPT the MSB itself) of
// mnt6_param_p (skipping leading zeros) in MSB to LSB order
for bit in BitIteratorBE::without_leading_zeros(P::ATE_LOOP_COUNT).skip(1) {
let dc = &q.double_coefficients[dbl_idx];
dbl_idx += 1;

View File

@@ -11,7 +11,7 @@ use crate::{
};
use algebra::{
curves::mnt6::{MNT6Parameters, MNT6},
fields::BitIterator,
fields::BitIteratorBE,
};
use core::marker::PhantomData;
@@ -96,18 +96,9 @@ impl<P: MNT6Parameters> PairingVar<P> {
let mut dbl_idx: usize = 0;
let mut add_idx: usize = 0;
let mut found_one = false;
for bit in BitIterator::new(P::ATE_LOOP_COUNT) {
// code below gets executed for all bits (EXCEPT the MSB itself) of
// mnt6_param_p (skipping leading zeros) in MSB to LSB order
if !found_one && bit {
found_one = true;
continue;
} else if !found_one {
continue;
}
// code below gets executed for all bits (EXCEPT the MSB itself) of
// mnt6_param_p (skipping leading zeros) in MSB to LSB order
for bit in BitIteratorBE::without_leading_zeros(P::ATE_LOOP_COUNT).skip(1) {
let dc = &q.double_coefficients[dbl_idx];
dbl_idx += 1;

View File

@@ -60,7 +60,7 @@ pub trait PairingVar<E: PairingEngine, ConstraintF: Field = <E as PairingEngine>
pub(crate) mod tests {
use crate::{prelude::*, Vec};
use algebra::{
test_rng, BitIterator, Field, PairingEngine, PrimeField, ProjectiveCurve, UniformRand,
test_rng, BitIteratorLE, Field, PairingEngine, PrimeField, ProjectiveCurve, UniformRand,
};
use r1cs_core::{ConstraintSystem, SynthesisError};
@@ -125,14 +125,14 @@ pub(crate) mod tests {
};
let (ans3_g, ans3_n) = {
let s_iter = BitIterator::new(s.into_repr())
let s_iter = BitIteratorLE::without_trailing_zeros(s.into_repr())
.map(Boolean::constant)
.collect::<Vec<_>>();
let mut ans_g = P::pairing(a_prep_g, b_prep_g)?;
let mut ans_n = E::pairing(a, b);
ans_n = ans_n.pow(s.into_repr());
ans_g = ans_g.pow(&s_iter)?;
ans_g = ans_g.pow_le(&s_iter)?;
(ans_g, ans_n)
};