Add solidity groth16, kzg10 and final decider verifiers in a dedicated workspace (#70)

* change: Refactor structure into workspace

* chore: Add empty readme

* change: Transform repo into workspace

* add: Create folding-verifier-solidity crate

* add: Include askama.toml for `sol` extension escaper

* add: Jordi's old Groth16 verifier .sol template and adapt it

* tmp: create simple template struct to test

* Update FoldingSchemes trait, fit Nova+CycleFold

- update lib.rs's `FoldingScheme` trait interface
- fit Nova+CycleFold into the `FoldingScheme` trait
- refactor `src/nova/*`

* chore: add serialization assets for testing

Now we include an `assets` folder with a serialized proof & vk for tests

* Add `examples` dir, with Nova's `FoldingScheme` example

* polishing

* expose poseidon_test_config outside tests

* change: Refactor structure into workspace

* chore: Add empty readme

* change: Transform repo into workspace

* add: Create folding-verifier-solidity crate

* add: Include askama.toml for `sol` extension escaper

* add: Jordi's old Groth16 verifier .sol template and adapt it

* tmp: create simple template struct to test

* feat: templating kzg working

* chore: add emv and revm

* feat: start evm file

* chore: add ark-poly-commit

* chore: move `commitment` to `folding-schemes`

* chore: update `.gitignore` to ignore generated contracts

* chore: update template with bn254 lib on it (avoids import), update for loop to account for whitespaces

* refactor: update template with no lib

* feat: add evm deploy code, compile and create kzg verifier

* chore: update `Cargo.toml` to have `folding-schemes` available with verifiers

* feat: start kzg prove and verify with sol

* chore: compute crs from kzg prover

* feat: evm kzg verification passing

* tmp

* change: Swap order of G2 coordinates within the template

* Update way to serialize proof with correct order

* chore: update `Cargo.toml`

* chore: add revm

* chore: add `save_solidity`

* refactor: verifiers in dedicated mod

* refactor: have dedicated `utils` module

* chore: expose modules

* chore: update verifier for kzg

* chore: rename templates

* fix: look for binary using also name of contract

* refactor: generate groth16 proof for sha256 pre-image, generate groth16 template with verifying key

* chore: template renaming

* fix: switch circuit for circuit that simply adds

* feat: generates test data on the fly

* feat: update to latest groth16 verifier

* refactor: rename folder, update `.gitignore`

* chore: update `Cargo.toml`

* chore: update templates extension to indicate that they are templates

* chore: rename templates, both files and structs

* fix: template inheritance working

* feat: template spdx and pragma statements

* feat: decider verifier compiles, update test for kzg10 and groth16 templates

* feat: parameterize which size of the crs should be stored on the contract

* chore: add comment on how the groth16 and kzg10 proofs will be linked together

* chore: cargo clippy run

* chore: cargo clippy tests

* chore: cargo fmt

* refactor: remove unused lifetime parameter

* chore: end merge

* chore: move examples to `folding-schemes` workspace

* get latest main changes

* fix: temp fix clippy warnings, will remove lints once not used in tests only

* fix: cargo clippy lint added on `code_size`

* fix: update path to test circuit and add step for installing solc

* chore: remove `save_solidity` steps

* fix: the borrowed expression implements the required traits

* chore: update `Cargo.toml`

* chore: remove extra `[patch.crates-io]`

* fix: update to patch at the workspace level and add comment explaining this

* refactor: correct `staticcall` with valid input/output sizes and change return syntax for pairing

* refactor: expose modules and remove `dead_code` calls

* chore: update `README.md`, add additional comments on `kzg10` template and update `groth16` template comments

* chore: be clearer on attributions on `kzg10`

---------

Co-authored-by: CPerezz <c.perezbaro@gmail.com>
Co-authored-by: arnaucube <root@arnaucube.com>
This commit is contained in:
Pierre
2024-02-09 08:19:25 +01:00
committed by GitHub
parent 97e973a685
commit 63dbbfe1bc
67 changed files with 1208 additions and 53 deletions

View File

@@ -0,0 +1,231 @@
use ark_ec::CurveGroup;
use ark_ff::PrimeField;
use ark_std::One;
use ark_std::Zero;
use std::ops::Add;
use std::sync::Arc;
use ark_std::{rand::Rng, UniformRand};
use super::utils::compute_sum_Mz;
use crate::ccs::CCS;
use crate::commitment::{
pedersen::{Params as PedersenParams, Pedersen},
CommitmentProver,
};
use crate::utils::hypercube::BooleanHypercube;
use crate::utils::mle::matrix_to_mle;
use crate::utils::mle::vec_to_mle;
use crate::utils::virtual_polynomial::VirtualPolynomial;
use crate::Error;
/// Witness for the LCCCS & CCCS, containing the w vector, and the r_w used as randomness in the Pedersen commitment.
#[derive(Debug, Clone)]
pub struct Witness<F: PrimeField> {
pub w: Vec<F>,
pub r_w: F, // randomness used in the Pedersen commitment of w
}
/// Committed CCS instance
#[derive(Debug, Clone)]
pub struct CCCS<C: CurveGroup> {
// Commitment to witness
pub C: C,
// Public input/output
pub x: Vec<C::ScalarField>,
}
impl<C: CurveGroup> CCS<C> {
pub fn to_cccs<R: Rng>(
&self,
rng: &mut R,
pedersen_params: &PedersenParams<C>,
z: &[C::ScalarField],
) -> Result<(CCCS<C>, Witness<C::ScalarField>), Error> {
let w: Vec<C::ScalarField> = z[(1 + self.l)..].to_vec();
let r_w = C::ScalarField::rand(rng);
let C = Pedersen::<C>::commit(pedersen_params, &w, &r_w)?;
Ok((
CCCS::<C> {
C,
x: z[1..(1 + self.l)].to_vec(),
},
Witness::<C::ScalarField> { w, r_w },
))
}
/// Computes q(x) = \sum^q c_i * \prod_{j \in S_i} ( \sum_{y \in {0,1}^s'} M_j(x, y) * z(y) )
/// polynomial over x
pub fn compute_q(&self, z: &Vec<C::ScalarField>) -> VirtualPolynomial<C::ScalarField> {
let z_mle = vec_to_mle(self.s_prime, z);
let mut q = VirtualPolynomial::<C::ScalarField>::new(self.s);
for i in 0..self.q {
let mut prod: VirtualPolynomial<C::ScalarField> =
VirtualPolynomial::<C::ScalarField>::new(self.s);
for j in self.S[i].clone() {
let M_j = matrix_to_mle(self.M[j].clone());
let sum_Mz = compute_sum_Mz(M_j, &z_mle, self.s_prime);
// Fold this sum into the running product
if prod.products.is_empty() {
// If this is the first time we are adding something to this virtual polynomial, we need to
// explicitly add the products using add_mle_list()
// XXX is this true? improve API
prod.add_mle_list([Arc::new(sum_Mz)], C::ScalarField::one())
.unwrap();
} else {
prod.mul_by_mle(Arc::new(sum_Mz), C::ScalarField::one())
.unwrap();
}
}
// Multiply by the product by the coefficient c_i
prod.scalar_mul(&self.c[i]);
// Add it to the running sum
q = q.add(&prod);
}
q
}
/// Computes Q(x) = eq(beta, x) * q(x)
/// = eq(beta, x) * \sum^q c_i * \prod_{j \in S_i} ( \sum_{y \in {0,1}^s'} M_j(x, y) * z(y) )
/// polynomial over x
pub fn compute_Q(
&self,
z: &Vec<C::ScalarField>,
beta: &[C::ScalarField],
) -> VirtualPolynomial<C::ScalarField> {
let q = self.compute_q(z);
q.build_f_hat(beta).unwrap()
}
}
impl<C: CurveGroup> CCCS<C> {
/// Perform the check of the CCCS instance described at section 4.1
pub fn check_relation(
&self,
pedersen_params: &PedersenParams<C>,
ccs: &CCS<C>,
w: &Witness<C::ScalarField>,
) -> Result<(), Error> {
// check that C is the commitment of w. Notice that this is not verifying a Pedersen
// opening, but checking that the Commmitment comes from committing to the witness.
if self.C != Pedersen::commit(pedersen_params, &w.w, &w.r_w)? {
return Err(Error::NotSatisfied);
}
// check CCCS relation
let z: Vec<C::ScalarField> =
[vec![C::ScalarField::one()], self.x.clone(), w.w.to_vec()].concat();
// A CCCS relation is satisfied if the q(x) multivariate polynomial evaluates to zero in the hypercube
let q_x = ccs.compute_q(&z);
for x in BooleanHypercube::new(ccs.s) {
if !q_x.evaluate(&x).unwrap().is_zero() {
return Err(Error::NotSatisfied);
}
}
Ok(())
}
}
#[cfg(test)]
pub mod tests {
use super::*;
use crate::ccs::tests::{get_test_ccs, get_test_z};
use ark_std::test_rng;
use ark_std::UniformRand;
use ark_pallas::{Fr, Projective};
/// Do some sanity checks on q(x). It's a multivariable polynomial and it should evaluate to zero inside the
/// hypercube, but to not-zero outside the hypercube.
#[test]
fn test_compute_q() {
let mut rng = test_rng();
let ccs = get_test_ccs::<Projective>();
let z = get_test_z(3);
let q = ccs.compute_q(&z);
// Evaluate inside the hypercube
for x in BooleanHypercube::new(ccs.s) {
assert_eq!(Fr::zero(), q.evaluate(&x).unwrap());
}
// Evaluate outside the hypercube
let beta: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
assert_ne!(Fr::zero(), q.evaluate(&beta).unwrap());
}
/// Perform some sanity checks on Q(x).
#[test]
fn test_compute_Q() {
let mut rng = test_rng();
let ccs: CCS<Projective> = get_test_ccs();
let z = get_test_z(3);
ccs.check_relation(&z).unwrap();
let beta: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Compute Q(x) = eq(beta, x) * q(x).
let Q = ccs.compute_Q(&z, &beta);
// Let's consider the multilinear polynomial G(x) = \sum_{y \in {0, 1}^s} eq(x, y) q(y)
// which interpolates the multivariate polynomial q(x) inside the hypercube.
//
// Observe that summing Q(x) inside the hypercube, directly computes G(\beta).
//
// Now, G(x) is multilinear and agrees with q(x) inside the hypercube. Since q(x) vanishes inside the
// hypercube, this means that G(x) also vanishes in the hypercube. Since G(x) is multilinear and vanishes
// inside the hypercube, this makes it the zero polynomial.
//
// Hence, evaluating G(x) at a random beta should give zero.
// Now sum Q(x) evaluations in the hypercube and expect it to be 0
let r = BooleanHypercube::new(ccs.s)
.map(|x| Q.evaluate(&x).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
assert_eq!(r, Fr::zero());
}
/// The polynomial G(x) (see above) interpolates q(x) inside the hypercube.
/// Summing Q(x) over the hypercube is equivalent to evaluating G(x) at some point.
/// This test makes sure that G(x) agrees with q(x) inside the hypercube, but not outside
#[test]
fn test_Q_against_q() {
let mut rng = test_rng();
let ccs: CCS<Projective> = get_test_ccs();
let z = get_test_z(3);
ccs.check_relation(&z).unwrap();
// Now test that if we create Q(x) with eq(d,y) where d is inside the hypercube, \sum Q(x) should be G(d) which
// should be equal to q(d), since G(x) interpolates q(x) inside the hypercube
let q = ccs.compute_q(&z);
for d in BooleanHypercube::new(ccs.s) {
let Q_at_d = ccs.compute_Q(&z, &d);
// Get G(d) by summing over Q_d(x) over the hypercube
let G_at_d = BooleanHypercube::new(ccs.s)
.map(|x| Q_at_d.evaluate(&x).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
assert_eq!(G_at_d, q.evaluate(&d).unwrap());
}
// Now test that they should disagree outside of the hypercube
let r: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
let Q_at_r = ccs.compute_Q(&z, &r);
// Get G(d) by summing over Q_d(x) over the hypercube
let G_at_r = BooleanHypercube::new(ccs.s)
.map(|x| Q_at_r.evaluate(&x).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
assert_ne!(G_at_r, q.evaluate(&r).unwrap());
}
}

View File

@@ -0,0 +1,318 @@
// hypernova nimfs verifier circuit
// see section 5 in https://eprint.iacr.org/2023/573.pdf
use crate::{ccs::CCS, folding::circuits::utils::EqEvalGadget};
use ark_ec::CurveGroup;
use ark_r1cs_std::{
alloc::AllocVar,
fields::{fp::FpVar, FieldVar},
ToBitsGadget,
};
use ark_relations::r1cs::{ConstraintSystemRef, SynthesisError};
use ark_std::Zero;
use std::marker::PhantomData;
/// Gadget to compute $\sum_{j \in [t]} \gamma^{j} \cdot e_1 \cdot \sigma_j + \gamma^{t+1} \cdot e_2 \cdot \sum_{i=1}^{q} c_i * \prod_{j \in S_i} \theta_j$.
/// This is the sum computed by the verifier and laid out in section 5, step 5 of "A multi-folding scheme for CCS".
pub struct ComputeCFromSigmasAndThetasGadget<C: CurveGroup> {
_c: PhantomData<C>,
}
impl<C: CurveGroup> ComputeCFromSigmasAndThetasGadget<C> {
/// Computes the sum $\sum_{j}^{j + n} \gamma^{j} \cdot eq_eval \cdot \sigma_{j}$, where $n$ is the length of the `sigmas` vector
/// It corresponds to the first term of the sum that $\mathcal{V}$ has to compute at section 5, step 5 of "A multi-folding scheme for CCS".
///
/// # Arguments
/// - `sigmas`: vector of $\sigma_j$ values
/// - `eq_eval`: the value of $\tilde{eq}(x_j, x^{\prime})$
/// - `gamma`: value $\gamma$
/// - `j`: the power at which we start to compute $\gamma^{j}$. This is needed in the context of multifolding.
///
/// # Notes
/// In the context of multifolding, `j` corresponds to `ccs.t` in `compute_c_from_sigmas_and_thetas`
fn sum_muls_gamma_pows_eq_sigma(
gamma: FpVar<C::ScalarField>,
eq_eval: FpVar<C::ScalarField>,
sigmas: Vec<FpVar<C::ScalarField>>,
j: FpVar<C::ScalarField>,
) -> Result<FpVar<C::ScalarField>, SynthesisError> {
let mut result = FpVar::<C::ScalarField>::zero();
let mut gamma_pow = gamma.pow_le(&j.to_bits_le()?)?;
for sigma in sigmas {
result += gamma_pow.clone() * eq_eval.clone() * sigma;
gamma_pow *= gamma.clone();
}
Ok(result)
}
/// Computes $\sum_{i=1}^{q} c_i * \prod_{j \in S_i} theta_j$
///
/// # Arguments
/// - `c_i`: vector of $c_i$ values
/// - `thetas`: vector of pre-processed $\thetas[j]$ values corresponding to a particular `ccs.S[i]`
///
/// # Notes
/// This is a part of the second term of the sum that $\mathcal{V}$ has to compute at section 5, step 5 of "A multi-folding scheme for CCS".
/// The first term is computed by `SumMulsGammaPowsEqSigmaGadget::sum_muls_gamma_pows_eq_sigma`.
/// This is a doct product between a vector of c_i values and a vector of pre-processed $\theta_j$ values, where $j$ is a value from $S_i$.
/// Hence, this requires some pre-processing of the $\theta_j$ values, before running this gadget.
fn sum_ci_mul_prod_thetaj(
c_i: Vec<FpVar<C::ScalarField>>,
thetas: Vec<Vec<FpVar<C::ScalarField>>>,
) -> Result<FpVar<C::ScalarField>, SynthesisError> {
let mut result = FpVar::<C::ScalarField>::zero();
for (i, c_i) in c_i.iter().enumerate() {
let prod = &thetas[i].iter().fold(FpVar::one(), |acc, e| acc * e);
result += c_i * prod;
}
Ok(result)
}
/// Computes the sum that the verifier has to compute at section 5, step 5 of "A multi-folding scheme for CCS".
///
/// # Arguments
/// - `cs`: constraint system
/// - `ccs`: the CCS instance
/// - `vec_sigmas`: vector of $\sigma_j$ values
/// - `vec_thetas`: vector of $\theta_j$ values
/// - `gamma`: value $\gamma$
/// - `beta`: vector of $\beta_j$ values
/// - `vec_r_x`: vector of $r_{x_j}$ values
/// - `vec_r_x_prime`: vector of $r_{x_j}^{\prime}$ values
///
/// # Notes
/// Arguments to this function are *almost* the same as the arguments to `compute_c_from_sigmas_and_thetas` in `utils.rs`.
#[allow(clippy::too_many_arguments)]
pub fn compute_c_from_sigmas_and_thetas(
cs: ConstraintSystemRef<C::ScalarField>,
ccs: &CCS<C>,
vec_sigmas: Vec<Vec<FpVar<C::ScalarField>>>,
vec_thetas: Vec<Vec<FpVar<C::ScalarField>>>,
gamma: FpVar<C::ScalarField>,
beta: Vec<FpVar<C::ScalarField>>,
vec_r_x: Vec<Vec<FpVar<C::ScalarField>>>,
vec_r_x_prime: Vec<FpVar<C::ScalarField>>,
) -> Result<FpVar<C::ScalarField>, SynthesisError> {
let mut c =
FpVar::<C::ScalarField>::new_witness(cs.clone(), || Ok(C::ScalarField::zero()))?;
let t = FpVar::<C::ScalarField>::new_witness(cs.clone(), || {
Ok(C::ScalarField::from(ccs.t as u64))
})?;
let mut e_lcccs = Vec::new();
for r_x in vec_r_x.iter() {
let e_1 = EqEvalGadget::eq_eval(r_x.to_vec(), vec_r_x_prime.to_vec())?;
e_lcccs.push(e_1);
}
for (i, sigmas) in vec_sigmas.iter().enumerate() {
let i_var = FpVar::<C::ScalarField>::new_witness(cs.clone(), || {
Ok(C::ScalarField::from(i as u64))
})?;
let pow = i_var * t.clone();
c += Self::sum_muls_gamma_pows_eq_sigma(
gamma.clone(),
e_lcccs[i].clone(),
sigmas.to_vec(),
pow,
)?;
}
let mu = FpVar::<C::ScalarField>::new_witness(cs.clone(), || {
Ok(C::ScalarField::from(vec_sigmas.len() as u64))
})?;
let e_2 = EqEvalGadget::eq_eval(beta, vec_r_x_prime)?;
for (k, thetas) in vec_thetas.iter().enumerate() {
// get prepared thetas. only step different from original `compute_c_from_sigmas_and_thetas`
let mut prepared_thetas = Vec::new();
for i in 0..ccs.q {
let prepared: Vec<FpVar<C::ScalarField>> =
ccs.S[i].iter().map(|j| thetas[*j].clone()).collect();
prepared_thetas.push(prepared.to_vec());
}
let c_i = Vec::<FpVar<C::ScalarField>>::new_witness(cs.clone(), || Ok(ccs.c.clone()))
.unwrap();
let lhs = Self::sum_ci_mul_prod_thetaj(c_i.clone(), prepared_thetas.clone())?;
// compute gamma^(t+1)
let pow = mu.clone() * t.clone()
+ FpVar::<C::ScalarField>::new_witness(cs.clone(), || {
Ok(C::ScalarField::from(k as u64))
})?;
let gamma_t1 = gamma.pow_le(&pow.to_bits_le()?)?;
c += gamma_t1.clone() * e_2.clone() * lhs.clone();
}
Ok(c)
}
}
#[cfg(test)]
mod tests {
use super::ComputeCFromSigmasAndThetasGadget;
use crate::{
ccs::{
tests::{get_test_ccs, get_test_z},
CCS,
},
commitment::pedersen::Pedersen,
folding::hypernova::utils::{
compute_c_from_sigmas_and_thetas, compute_sigmas_and_thetas, sum_ci_mul_prod_thetaj,
sum_muls_gamma_pows_eq_sigma,
},
utils::virtual_polynomial::eq_eval,
};
use ark_pallas::{Fr, Projective};
use ark_r1cs_std::{alloc::AllocVar, fields::fp::FpVar, R1CSVar};
use ark_relations::r1cs::ConstraintSystem;
use ark_std::{test_rng, UniformRand};
#[test]
pub fn test_sum_muls_gamma_pow_eq_sigma_gadget() {
let mut rng = test_rng();
let ccs: CCS<Projective> = get_test_ccs();
let z1 = get_test_z(3);
let z2 = get_test_z(4);
let gamma: Fr = Fr::rand(&mut rng);
let r_x_prime: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Initialize a multifolding object
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs_instance, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let sigmas_thetas =
compute_sigmas_and_thetas(&ccs, &[z1.clone()], &[z2.clone()], &r_x_prime);
let mut e_lcccs = Vec::new();
for r_x in &vec![lcccs_instance.r_x] {
e_lcccs.push(eq_eval(r_x, &r_x_prime).unwrap());
}
// Initialize cs and gamma
let cs = ConstraintSystem::<Fr>::new_ref();
let gamma_var = FpVar::<Fr>::new_witness(cs.clone(), || Ok(gamma)).unwrap();
for (i, sigmas) in sigmas_thetas.0.iter().enumerate() {
let expected =
sum_muls_gamma_pows_eq_sigma(gamma, e_lcccs[i], sigmas, (i * ccs.t) as u64);
let sigmas_var =
Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(sigmas.clone())).unwrap();
let eq_var = FpVar::<Fr>::new_witness(cs.clone(), || Ok(e_lcccs[i])).unwrap();
let pow =
FpVar::<Fr>::new_witness(cs.clone(), || Ok(Fr::from((i * ccs.t) as u64))).unwrap();
let computed =
ComputeCFromSigmasAndThetasGadget::<Projective>::sum_muls_gamma_pows_eq_sigma(
gamma_var.clone(),
eq_var,
sigmas_var,
pow,
)
.unwrap();
assert_eq!(expected, computed.value().unwrap());
}
}
#[test]
pub fn test_sum_ci_mul_prod_thetaj_gadget() {
let mut rng = test_rng();
let ccs: CCS<Projective> = get_test_ccs();
let z1 = get_test_z(3);
let z2 = get_test_z(4);
let r_x_prime: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Initialize a multifolding object
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs_instance, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let sigmas_thetas =
compute_sigmas_and_thetas(&ccs, &[z1.clone()], &[z2.clone()], &r_x_prime);
let mut e_lcccs = Vec::new();
for r_x in &vec![lcccs_instance.r_x] {
e_lcccs.push(eq_eval(r_x, &r_x_prime).unwrap());
}
// Initialize cs
let cs = ConstraintSystem::<Fr>::new_ref();
let vec_thetas = sigmas_thetas.1;
for thetas in vec_thetas.iter() {
// sum c_i * prod theta_j
let expected = sum_ci_mul_prod_thetaj(&ccs, thetas); // from `compute_c_from_sigmas_and_thetas`
let mut prepared_thetas = Vec::new();
for i in 0..ccs.q {
let prepared: Vec<Fr> = ccs.S[i].iter().map(|j| thetas[*j]).collect();
prepared_thetas
.push(Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(prepared)).unwrap());
}
let computed = ComputeCFromSigmasAndThetasGadget::<Projective>::sum_ci_mul_prod_thetaj(
Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(ccs.c.clone())).unwrap(),
prepared_thetas,
)
.unwrap();
assert_eq!(expected, computed.value().unwrap());
}
}
#[test]
pub fn test_compute_c_from_sigmas_and_thetas_gadget() {
let ccs: CCS<Projective> = get_test_ccs();
let z1 = get_test_z(3);
let z2 = get_test_z(4);
let mut rng = test_rng();
let gamma: Fr = Fr::rand(&mut rng);
let beta: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
let r_x_prime: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Initialize a multifolding object
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs_instance, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let sigmas_thetas =
compute_sigmas_and_thetas(&ccs, &[z1.clone()], &[z2.clone()], &r_x_prime);
let expected_c = compute_c_from_sigmas_and_thetas(
&ccs,
&sigmas_thetas,
gamma,
&beta,
&vec![lcccs_instance.r_x.clone()],
&r_x_prime,
);
let cs = ConstraintSystem::<Fr>::new_ref();
let mut vec_sigmas = Vec::new();
let mut vec_thetas = Vec::new();
for sigmas in sigmas_thetas.0 {
vec_sigmas
.push(Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(sigmas.clone())).unwrap());
}
for thetas in sigmas_thetas.1 {
vec_thetas
.push(Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(thetas.clone())).unwrap());
}
let vec_r_x =
vec![
Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(lcccs_instance.r_x.clone()))
.unwrap(),
];
let vec_r_x_prime =
Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(r_x_prime.clone())).unwrap();
let gamma_var = FpVar::<Fr>::new_witness(cs.clone(), || Ok(gamma)).unwrap();
let beta_var = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(beta.clone())).unwrap();
let computed_c = ComputeCFromSigmasAndThetasGadget::compute_c_from_sigmas_and_thetas(
cs,
&ccs,
vec_sigmas,
vec_thetas,
gamma_var,
beta_var,
vec_r_x,
vec_r_x_prime,
)
.unwrap();
assert_eq!(expected_c, computed_c.value().unwrap());
}
}

View File

@@ -0,0 +1,188 @@
use ark_ec::CurveGroup;
use ark_poly::DenseMultilinearExtension;
use ark_std::One;
use std::sync::Arc;
use ark_std::{rand::Rng, UniformRand};
use super::cccs::Witness;
use super::utils::{compute_all_sum_Mz_evals, compute_sum_Mz};
use crate::ccs::CCS;
use crate::commitment::{
pedersen::{Params as PedersenParams, Pedersen},
CommitmentProver,
};
use crate::utils::mle::{matrix_to_mle, vec_to_mle};
use crate::utils::virtual_polynomial::VirtualPolynomial;
use crate::Error;
/// Linearized Committed CCS instance
#[derive(Debug, Clone, Eq, PartialEq)]
pub struct LCCCS<C: CurveGroup> {
// Commitment to witness
pub C: C,
// Relaxation factor of z for folded LCCCS
pub u: C::ScalarField,
// Public input/output
pub x: Vec<C::ScalarField>,
// Random evaluation point for the v_i
pub r_x: Vec<C::ScalarField>,
// Vector of v_i
pub v: Vec<C::ScalarField>,
}
impl<C: CurveGroup> CCS<C> {
/// Compute v_j values of the linearized committed CCS form
/// Given `r`, compute: \sum_{y \in {0,1}^s'} M_j(r, y) * z(y)
fn compute_v_j(&self, z: &[C::ScalarField], r: &[C::ScalarField]) -> Vec<C::ScalarField> {
compute_all_sum_Mz_evals(&self.M, &z.to_vec(), r, self.s_prime)
}
pub fn to_lcccs<R: Rng>(
&self,
rng: &mut R,
pedersen_params: &PedersenParams<C>,
z: &[C::ScalarField],
) -> Result<(LCCCS<C>, Witness<C::ScalarField>), Error> {
let w: Vec<C::ScalarField> = z[(1 + self.l)..].to_vec();
let r_w = C::ScalarField::rand(rng);
let C = Pedersen::commit(pedersen_params, &w, &r_w)?;
let r_x: Vec<C::ScalarField> = (0..self.s).map(|_| C::ScalarField::rand(rng)).collect();
let v = self.compute_v_j(z, &r_x);
Ok((
LCCCS::<C> {
C,
u: C::ScalarField::one(),
x: z[1..(1 + self.l)].to_vec(),
r_x,
v,
},
Witness::<C::ScalarField> { w, r_w },
))
}
}
impl<C: CurveGroup> LCCCS<C> {
/// Compute all L_j(x) polynomials
pub fn compute_Ls(
&self,
ccs: &CCS<C>,
z: &Vec<C::ScalarField>,
) -> Vec<VirtualPolynomial<C::ScalarField>> {
let z_mle = vec_to_mle(ccs.s_prime, z);
// Convert all matrices to MLE
let M_x_y_mle: Vec<DenseMultilinearExtension<C::ScalarField>> =
ccs.M.clone().into_iter().map(matrix_to_mle).collect();
let mut vec_L_j_x = Vec::with_capacity(ccs.t);
for M_j in M_x_y_mle {
let sum_Mz = compute_sum_Mz(M_j, &z_mle, ccs.s_prime);
let sum_Mz_virtual =
VirtualPolynomial::new_from_mle(&Arc::new(sum_Mz.clone()), C::ScalarField::one());
let L_j_x = sum_Mz_virtual.build_f_hat(&self.r_x).unwrap();
vec_L_j_x.push(L_j_x);
}
vec_L_j_x
}
/// Perform the check of the LCCCS instance described at section 4.2
pub fn check_relation(
&self,
pedersen_params: &PedersenParams<C>,
ccs: &CCS<C>,
w: &Witness<C::ScalarField>,
) -> Result<(), Error> {
// check that C is the commitment of w. Notice that this is not verifying a Pedersen
// opening, but checking that the Commmitment comes from committing to the witness.
if self.C != Pedersen::commit(pedersen_params, &w.w, &w.r_w)? {
return Err(Error::NotSatisfied);
}
// check CCS relation
let z: Vec<C::ScalarField> = [vec![self.u], self.x.clone(), w.w.to_vec()].concat();
let computed_v = compute_all_sum_Mz_evals(&ccs.M, &z, &self.r_x, ccs.s_prime);
if computed_v != self.v {
return Err(Error::NotSatisfied);
}
Ok(())
}
}
#[cfg(test)]
pub mod tests {
use super::*;
use ark_std::Zero;
use crate::ccs::tests::{get_test_ccs, get_test_z};
use crate::utils::hypercube::BooleanHypercube;
use ark_std::test_rng;
use ark_pallas::{Fr, Projective};
#[test]
/// Test linearized CCCS v_j against the L_j(x)
fn test_lcccs_v_j() {
let mut rng = test_rng();
let ccs = get_test_ccs();
let z = get_test_z(3);
ccs.check_relation(&z.clone()).unwrap();
let pedersen_params = Pedersen::<Projective>::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z).unwrap();
// with our test vector comming from R1CS, v should have length 3
assert_eq!(lcccs.v.len(), 3);
let vec_L_j_x = lcccs.compute_Ls(&ccs, &z);
assert_eq!(vec_L_j_x.len(), lcccs.v.len());
for (v_i, L_j_x) in lcccs.v.into_iter().zip(vec_L_j_x) {
let sum_L_j_x = BooleanHypercube::new(ccs.s)
.map(|y| L_j_x.evaluate(&y).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
assert_eq!(v_i, sum_L_j_x);
}
}
/// Given a bad z, check that the v_j should not match with the L_j(x)
#[test]
fn test_bad_v_j() {
let mut rng = test_rng();
let ccs = get_test_ccs();
let z = get_test_z(3);
ccs.check_relation(&z.clone()).unwrap();
// Mutate z so that the relation does not hold
let mut bad_z = z.clone();
bad_z[3] = Fr::zero();
assert!(ccs.check_relation(&bad_z.clone()).is_err());
let pedersen_params = Pedersen::<Projective>::new_params(&mut rng, ccs.n - ccs.l - 1);
// Compute v_j with the right z
let (lcccs, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z).unwrap();
// with our test vector comming from R1CS, v should have length 3
assert_eq!(lcccs.v.len(), 3);
// Bad compute L_j(x) with the bad z
let vec_L_j_x = lcccs.compute_Ls(&ccs, &bad_z);
assert_eq!(vec_L_j_x.len(), lcccs.v.len());
// Make sure that the LCCCS is not satisfied given these L_j(x)
// i.e. summing L_j(x) over the hypercube should not give v_j for all j
let mut satisfied = true;
for (v_i, L_j_x) in lcccs.v.into_iter().zip(vec_L_j_x) {
let sum_L_j_x = BooleanHypercube::new(ccs.s)
.map(|y| L_j_x.evaluate(&y).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
if v_i != sum_L_j_x {
satisfied = false;
}
}
assert!(!satisfied);
}
}

View File

@@ -0,0 +1,6 @@
/// Implements the scheme described in [HyperNova](https://eprint.iacr.org/2023/573.pdf)
pub mod cccs;
pub mod circuit;
pub mod lcccs;
pub mod nimfs;
pub mod utils;

View File

@@ -0,0 +1,721 @@
use ark_crypto_primitives::sponge::Absorb;
use ark_ec::{CurveGroup, Group};
use ark_ff::{Field, PrimeField};
use ark_poly::univariate::DensePolynomial;
use ark_poly::{DenseUVPolynomial, Polynomial};
use ark_std::{One, Zero};
use super::cccs::{Witness, CCCS};
use super::lcccs::LCCCS;
use super::utils::{compute_c_from_sigmas_and_thetas, compute_g, compute_sigmas_and_thetas};
use crate::ccs::CCS;
use crate::transcript::Transcript;
use crate::utils::hypercube::BooleanHypercube;
use crate::utils::sum_check::structs::IOPProof as SumCheckProof;
use crate::utils::sum_check::{IOPSumCheck, SumCheck};
use crate::utils::virtual_polynomial::VPAuxInfo;
use crate::Error;
use std::fmt::Debug;
use std::marker::PhantomData;
/// Proof defines a multifolding proof
#[derive(Debug)]
pub struct Proof<C: CurveGroup> {
pub sc_proof: SumCheckProof<C::ScalarField>,
pub sigmas_thetas: SigmasThetas<C::ScalarField>,
}
#[derive(Debug)]
pub struct SigmasThetas<F: PrimeField>(pub Vec<Vec<F>>, pub Vec<Vec<F>>);
#[derive(Debug)]
/// Implements the Non-Interactive Multi Folding Scheme described in section 5 of
/// [HyperNova](https://eprint.iacr.org/2023/573.pdf)
pub struct NIMFS<C: CurveGroup, T: Transcript<C>> {
pub _c: PhantomData<C>,
pub _t: PhantomData<T>,
}
impl<C: CurveGroup, T: Transcript<C>> NIMFS<C, T>
where
<C as Group>::ScalarField: Absorb,
{
pub fn fold(
lcccs: &[LCCCS<C>],
cccs: &[CCCS<C>],
sigmas_thetas: &SigmasThetas<C::ScalarField>,
r_x_prime: Vec<C::ScalarField>,
rho: C::ScalarField,
) -> LCCCS<C> {
let (sigmas, thetas) = (sigmas_thetas.0.clone(), sigmas_thetas.1.clone());
let mut C_folded = C::zero();
let mut u_folded = C::ScalarField::zero();
let mut x_folded: Vec<C::ScalarField> = vec![C::ScalarField::zero(); lcccs[0].x.len()];
let mut v_folded: Vec<C::ScalarField> = vec![C::ScalarField::zero(); sigmas[0].len()];
for i in 0..(lcccs.len() + cccs.len()) {
let rho_i = rho.pow([i as u64]);
let c: C;
let u: C::ScalarField;
let x: Vec<C::ScalarField>;
let v: Vec<C::ScalarField>;
if i < lcccs.len() {
c = lcccs[i].C;
u = lcccs[i].u;
x = lcccs[i].x.clone();
v = sigmas[i].clone();
} else {
c = cccs[i - lcccs.len()].C;
u = C::ScalarField::one();
x = cccs[i - lcccs.len()].x.clone();
v = thetas[i - lcccs.len()].clone();
}
C_folded += c.mul(rho_i);
u_folded += rho_i * u;
x_folded = x_folded
.iter()
.zip(
x.iter()
.map(|x_i| *x_i * rho_i)
.collect::<Vec<C::ScalarField>>(),
)
.map(|(a_i, b_i)| *a_i + b_i)
.collect();
v_folded = v_folded
.iter()
.zip(
v.iter()
.map(|x_i| *x_i * rho_i)
.collect::<Vec<C::ScalarField>>(),
)
.map(|(a_i, b_i)| *a_i + b_i)
.collect();
}
LCCCS::<C> {
C: C_folded,
u: u_folded,
x: x_folded,
r_x: r_x_prime,
v: v_folded,
}
}
pub fn fold_witness(
w_lcccs: &[Witness<C::ScalarField>],
w_cccs: &[Witness<C::ScalarField>],
rho: C::ScalarField,
) -> Witness<C::ScalarField> {
let mut w_folded: Vec<C::ScalarField> = vec![C::ScalarField::zero(); w_lcccs[0].w.len()];
let mut r_w_folded = C::ScalarField::zero();
for i in 0..(w_lcccs.len() + w_cccs.len()) {
let rho_i = rho.pow([i as u64]);
let w: Vec<C::ScalarField>;
let r_w: C::ScalarField;
if i < w_lcccs.len() {
w = w_lcccs[i].w.clone();
r_w = w_lcccs[i].r_w;
} else {
w = w_cccs[i - w_lcccs.len()].w.clone();
r_w = w_cccs[i - w_lcccs.len()].r_w;
}
w_folded = w_folded
.iter()
.zip(
w.iter()
.map(|x_i| *x_i * rho_i)
.collect::<Vec<C::ScalarField>>(),
)
.map(|(a_i, b_i)| *a_i + b_i)
.collect();
r_w_folded += rho_i * r_w;
}
Witness {
w: w_folded,
r_w: r_w_folded,
}
}
/// Performs the multifolding prover. Given μ LCCCS instances and ν CCS instances, fold them
/// into a single LCCCS instance. Since this is the prover, also fold their witness.
/// Returns the final folded LCCCS, the folded witness, and the multifolding proof, which
/// contains the sumcheck proof and the helper sumcheck claim sigmas and thetas.
#[allow(clippy::type_complexity)]
pub fn prove(
transcript: &mut impl Transcript<C>,
ccs: &CCS<C>,
running_instances: &[LCCCS<C>],
new_instances: &[CCCS<C>],
w_lcccs: &[Witness<C::ScalarField>],
w_cccs: &[Witness<C::ScalarField>],
) -> Result<(Proof<C>, LCCCS<C>, Witness<C::ScalarField>), Error> {
// TODO appends to transcript
if running_instances.is_empty() {
return Err(Error::Empty);
}
if new_instances.is_empty() {
return Err(Error::Empty);
}
// construct the LCCCS z vector from the relaxation factor, public IO and witness
// XXX this deserves its own function in LCCCS
let mut z_lcccs = Vec::new();
for (i, running_instance) in running_instances.iter().enumerate() {
let z_1: Vec<C::ScalarField> = [
vec![running_instance.u],
running_instance.x.clone(),
w_lcccs[i].w.to_vec(),
]
.concat();
z_lcccs.push(z_1);
}
// construct the CCCS z vector from the public IO and witness
let mut z_cccs = Vec::new();
for (i, new_instance) in new_instances.iter().enumerate() {
let z_2: Vec<C::ScalarField> = [
vec![C::ScalarField::one()],
new_instance.x.clone(),
w_cccs[i].w.to_vec(),
]
.concat();
z_cccs.push(z_2);
}
// Step 1: Get some challenges
let gamma_scalar = C::ScalarField::from_le_bytes_mod_order(b"gamma");
let beta_scalar = C::ScalarField::from_le_bytes_mod_order(b"beta");
transcript.absorb(&gamma_scalar);
let gamma: C::ScalarField = transcript.get_challenge();
transcript.absorb(&beta_scalar);
let beta: Vec<C::ScalarField> = transcript.get_challenges(ccs.s);
// Compute g(x)
let g = compute_g(ccs, running_instances, &z_lcccs, &z_cccs, gamma, &beta);
// Step 3: Run the sumcheck prover
let sumcheck_proof = IOPSumCheck::<C, T>::prove(&g, transcript)
.map_err(|err| Error::SumCheckProveError(err.to_string()))?;
// Note: The following two "sanity checks" are done for this prototype, in a final version
// they should be removed.
//
// Sanity check 1: evaluate g(x) over x \in {0,1} (the boolean hypercube), and check that
// its sum is equal to the extracted_sum from the SumCheck.
//////////////////////////////////////////////////////////////////////
let mut g_over_bhc = C::ScalarField::zero();
for x in BooleanHypercube::new(ccs.s) {
g_over_bhc += g.evaluate(&x)?;
}
// note: this is the sum of g(x) over the whole boolean hypercube
let extracted_sum = IOPSumCheck::<C, T>::extract_sum(&sumcheck_proof);
if extracted_sum != g_over_bhc {
return Err(Error::NotEqual);
}
// Sanity check 2: expect \sum v_j * gamma^j to be equal to the sum of g(x) over the
// boolean hypercube (and also equal to the extracted_sum from the SumCheck).
let mut sum_v_j_gamma = C::ScalarField::zero();
for (i, running_instance) in running_instances.iter().enumerate() {
for j in 0..running_instance.v.len() {
let gamma_j = gamma.pow([(i * ccs.t + j) as u64]);
sum_v_j_gamma += running_instance.v[j] * gamma_j;
}
}
if g_over_bhc != sum_v_j_gamma {
return Err(Error::NotEqual);
}
if extracted_sum != sum_v_j_gamma {
return Err(Error::NotEqual);
}
//////////////////////////////////////////////////////////////////////
// Step 2: dig into the sumcheck and extract r_x_prime
let r_x_prime = sumcheck_proof.point.clone();
// Step 4: compute sigmas and thetas
let sigmas_thetas = compute_sigmas_and_thetas(ccs, &z_lcccs, &z_cccs, &r_x_prime);
// Step 6: Get the folding challenge
let rho_scalar = C::ScalarField::from_le_bytes_mod_order(b"rho");
transcript.absorb(&rho_scalar);
let rho: C::ScalarField = transcript.get_challenge();
// Step 7: Create the folded instance
let folded_lcccs = Self::fold(
running_instances,
new_instances,
&sigmas_thetas,
r_x_prime,
rho,
);
// Step 8: Fold the witnesses
let folded_witness = Self::fold_witness(w_lcccs, w_cccs, rho);
Ok((
Proof::<C> {
sc_proof: sumcheck_proof,
sigmas_thetas,
},
folded_lcccs,
folded_witness,
))
}
/// Performs the multifolding verifier. Given μ LCCCS instances and ν CCS instances, fold them
/// into a single LCCCS instance.
/// Returns the folded LCCCS instance.
pub fn verify(
transcript: &mut impl Transcript<C>,
ccs: &CCS<C>,
running_instances: &[LCCCS<C>],
new_instances: &[CCCS<C>],
proof: Proof<C>,
) -> Result<LCCCS<C>, Error> {
// TODO appends to transcript
if running_instances.is_empty() {
return Err(Error::Empty);
}
if new_instances.is_empty() {
return Err(Error::Empty);
}
// Step 1: Get some challenges
let gamma_scalar = C::ScalarField::from_le_bytes_mod_order(b"gamma");
transcript.absorb(&gamma_scalar);
let gamma: C::ScalarField = transcript.get_challenge();
let beta_scalar = C::ScalarField::from_le_bytes_mod_order(b"beta");
transcript.absorb(&beta_scalar);
let beta: Vec<C::ScalarField> = transcript.get_challenges(ccs.s);
let vp_aux_info = VPAuxInfo::<C::ScalarField> {
max_degree: ccs.d + 1,
num_variables: ccs.s,
phantom: PhantomData::<C::ScalarField>,
};
// Step 3: Start verifying the sumcheck
// First, compute the expected sumcheck sum: \sum gamma^j v_j
let mut sum_v_j_gamma = C::ScalarField::zero();
for (i, running_instance) in running_instances.iter().enumerate() {
for j in 0..running_instance.v.len() {
let gamma_j = gamma.pow([(i * ccs.t + j) as u64]);
sum_v_j_gamma += running_instance.v[j] * gamma_j;
}
}
// Verify the interactive part of the sumcheck
let sumcheck_subclaim =
IOPSumCheck::<C, T>::verify(sum_v_j_gamma, &proof.sc_proof, &vp_aux_info, transcript)
.map_err(|err| Error::SumCheckVerifyError(err.to_string()))?;
// Step 2: Dig into the sumcheck claim and extract the randomness used
let r_x_prime = sumcheck_subclaim.point.clone();
// Step 5: Finish verifying sumcheck (verify the claim c)
let c = compute_c_from_sigmas_and_thetas(
ccs,
&proof.sigmas_thetas,
gamma,
&beta,
&running_instances
.iter()
.map(|lcccs| lcccs.r_x.clone())
.collect(),
&r_x_prime,
);
// check that the g(r_x') from the sumcheck proof is equal to the computed c from sigmas&thetas
if c != sumcheck_subclaim.expected_evaluation {
return Err(Error::NotEqual);
}
// Sanity check: we can also compute g(r_x') from the proof last evaluation value, and
// should be equal to the previously obtained values.
let g_on_rxprime_from_sumcheck_last_eval =
DensePolynomial::from_coefficients_slice(&proof.sc_proof.proofs.last().unwrap().coeffs)
.evaluate(r_x_prime.last().unwrap());
if g_on_rxprime_from_sumcheck_last_eval != c {
return Err(Error::NotEqual);
}
if g_on_rxprime_from_sumcheck_last_eval != sumcheck_subclaim.expected_evaluation {
return Err(Error::NotEqual);
}
// Step 6: Get the folding challenge
let rho_scalar = C::ScalarField::from_le_bytes_mod_order(b"rho");
transcript.absorb(&rho_scalar);
let rho: C::ScalarField = transcript.get_challenge();
// Step 7: Compute the folded instance
Ok(Self::fold(
running_instances,
new_instances,
&proof.sigmas_thetas,
r_x_prime,
rho,
))
}
}
#[cfg(test)]
pub mod tests {
use super::*;
use crate::ccs::tests::{get_test_ccs, get_test_z};
use crate::transcript::poseidon::poseidon_test_config;
use crate::transcript::poseidon::PoseidonTranscript;
use ark_std::test_rng;
use ark_std::UniformRand;
use crate::commitment::pedersen::Pedersen;
use ark_pallas::{Fr, Projective};
#[test]
fn test_fold() {
let ccs = get_test_ccs();
let z1 = get_test_z::<Fr>(3);
let z2 = get_test_z::<Fr>(4);
ccs.check_relation(&z1).unwrap();
ccs.check_relation(&z2).unwrap();
let mut rng = test_rng();
let r_x_prime: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
let sigmas_thetas =
compute_sigmas_and_thetas(&ccs, &[z1.clone()], &[z2.clone()], &r_x_prime);
let pedersen_params = Pedersen::<Projective>::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs, w1) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let (cccs, w2) = ccs.to_cccs(&mut rng, &pedersen_params, &z2).unwrap();
lcccs.check_relation(&pedersen_params, &ccs, &w1).unwrap();
cccs.check_relation(&pedersen_params, &ccs, &w2).unwrap();
let mut rng = test_rng();
let rho = Fr::rand(&mut rng);
let folded = NIMFS::<Projective, PoseidonTranscript<Projective>>::fold(
&[lcccs],
&[cccs],
&sigmas_thetas,
r_x_prime,
rho,
);
let w_folded =
NIMFS::<Projective, PoseidonTranscript<Projective>>::fold_witness(&[w1], &[w2], rho);
// check lcccs relation
folded
.check_relation(&pedersen_params, &ccs, &w_folded)
.unwrap();
}
/// Perform multifolding of an LCCCS instance with a CCCS instance (as described in the paper)
#[test]
pub fn test_basic_multifolding() {
let mut rng = test_rng();
// Create a basic CCS circuit
let ccs = get_test_ccs::<Projective>();
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
// Generate a satisfying witness
let z_1 = get_test_z(3);
// Generate another satisfying witness
let z_2 = get_test_z(4);
// Create the LCCCS instance out of z_1
let (running_instance, w1) = ccs.to_lcccs(&mut rng, &pedersen_params, &z_1).unwrap();
// Create the CCCS instance out of z_2
let (new_instance, w2) = ccs.to_cccs(&mut rng, &pedersen_params, &z_2).unwrap();
// Prover's transcript
let poseidon_config = poseidon_test_config::<Fr>();
let mut transcript_p: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_p.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
// Run the prover side of the multifolding
let (proof, folded_lcccs, folded_witness) =
NIMFS::<Projective, PoseidonTranscript<Projective>>::prove(
&mut transcript_p,
&ccs,
&[running_instance.clone()],
&[new_instance.clone()],
&[w1],
&[w2],
)
.unwrap();
// Verifier's transcript
let mut transcript_v: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_v.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
// Run the verifier side of the multifolding
let folded_lcccs_v = NIMFS::<Projective, PoseidonTranscript<Projective>>::verify(
&mut transcript_v,
&ccs,
&[running_instance.clone()],
&[new_instance.clone()],
proof,
)
.unwrap();
assert_eq!(folded_lcccs, folded_lcccs_v);
// Check that the folded LCCCS instance is a valid instance with respect to the folded witness
folded_lcccs
.check_relation(&pedersen_params, &ccs, &folded_witness)
.unwrap();
}
/// Perform multiple steps of multifolding of an LCCCS instance with a CCCS instance
#[test]
pub fn test_multifolding_two_instances_multiple_steps() {
let mut rng = test_rng();
let ccs = get_test_ccs::<Projective>();
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
// LCCCS witness
let z_1 = get_test_z(2);
let (mut running_instance, mut w1) =
ccs.to_lcccs(&mut rng, &pedersen_params, &z_1).unwrap();
let poseidon_config = poseidon_test_config::<Fr>();
let mut transcript_p: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_p.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
let mut transcript_v: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_v.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
let n: usize = 10;
for i in 3..n {
println!("\niteration: i {}", i); // DBG
// CCS witness
let z_2 = get_test_z(i);
println!("z_2 {:?}", z_2); // DBG
let (new_instance, w2) = ccs.to_cccs(&mut rng, &pedersen_params, &z_2).unwrap();
// run the prover side of the multifolding
let (proof, folded_lcccs, folded_witness) =
NIMFS::<Projective, PoseidonTranscript<Projective>>::prove(
&mut transcript_p,
&ccs,
&[running_instance.clone()],
&[new_instance.clone()],
&[w1],
&[w2],
)
.unwrap();
// run the verifier side of the multifolding
let folded_lcccs_v = NIMFS::<Projective, PoseidonTranscript<Projective>>::verify(
&mut transcript_v,
&ccs,
&[running_instance.clone()],
&[new_instance.clone()],
proof,
)
.unwrap();
assert_eq!(folded_lcccs, folded_lcccs_v);
// check that the folded instance with the folded witness holds the LCCCS relation
println!("check_relation {}", i);
folded_lcccs
.check_relation(&pedersen_params, &ccs, &folded_witness)
.unwrap();
running_instance = folded_lcccs;
w1 = folded_witness;
}
}
/// Test that generates mu>1 and nu>1 instances, and folds them in a single multifolding step.
#[test]
pub fn test_multifolding_mu_nu_instances() {
let mut rng = test_rng();
// Create a basic CCS circuit
let ccs = get_test_ccs::<Projective>();
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let mu = 10;
let nu = 15;
// Generate a mu LCCCS & nu CCCS satisfying witness
let mut z_lcccs = Vec::new();
for i in 0..mu {
let z = get_test_z(i + 3);
z_lcccs.push(z);
}
let mut z_cccs = Vec::new();
for i in 0..nu {
let z = get_test_z(nu + i + 3);
z_cccs.push(z);
}
// Create the LCCCS instances out of z_lcccs
let mut lcccs_instances = Vec::new();
let mut w_lcccs = Vec::new();
for z_i in z_lcccs.iter() {
let (running_instance, w) = ccs.to_lcccs(&mut rng, &pedersen_params, z_i).unwrap();
lcccs_instances.push(running_instance);
w_lcccs.push(w);
}
// Create the CCCS instance out of z_cccs
let mut cccs_instances = Vec::new();
let mut w_cccs = Vec::new();
for z_i in z_cccs.iter() {
let (new_instance, w) = ccs.to_cccs(&mut rng, &pedersen_params, z_i).unwrap();
cccs_instances.push(new_instance);
w_cccs.push(w);
}
// Prover's transcript
let poseidon_config = poseidon_test_config::<Fr>();
let mut transcript_p: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_p.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
// Run the prover side of the multifolding
let (proof, folded_lcccs, folded_witness) =
NIMFS::<Projective, PoseidonTranscript<Projective>>::prove(
&mut transcript_p,
&ccs,
&lcccs_instances,
&cccs_instances,
&w_lcccs,
&w_cccs,
)
.unwrap();
// Verifier's transcript
let mut transcript_v: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_v.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
// Run the verifier side of the multifolding
let folded_lcccs_v = NIMFS::<Projective, PoseidonTranscript<Projective>>::verify(
&mut transcript_v,
&ccs,
&lcccs_instances,
&cccs_instances,
proof,
)
.unwrap();
assert_eq!(folded_lcccs, folded_lcccs_v);
// Check that the folded LCCCS instance is a valid instance with respect to the folded witness
folded_lcccs
.check_relation(&pedersen_params, &ccs, &folded_witness)
.unwrap();
}
/// Test that generates mu>1 and nu>1 instances, and folds them in a single multifolding step
/// and repeats the process doing multiple steps.
#[test]
pub fn test_multifolding_mu_nu_instances_multiple_steps() {
let mut rng = test_rng();
// Create a basic CCS circuit
let ccs = get_test_ccs::<Projective>();
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let poseidon_config = poseidon_test_config::<Fr>();
// Prover's transcript
let mut transcript_p: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_p.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
// Verifier's transcript
let mut transcript_v: PoseidonTranscript<Projective> =
PoseidonTranscript::<Projective>::new(&poseidon_config);
transcript_v.absorb(&Fr::from_le_bytes_mod_order(b"init init"));
let n_steps = 3;
// number of LCCCS & CCCS instances in each multifolding step
let mu = 10;
let nu = 15;
// Generate a mu LCCCS & nu CCCS satisfying witness, for each step
for step in 0..n_steps {
let mut z_lcccs = Vec::new();
for i in 0..mu {
let z = get_test_z(step + i + 3);
z_lcccs.push(z);
}
let mut z_cccs = Vec::new();
for i in 0..nu {
let z = get_test_z(nu + i + 3);
z_cccs.push(z);
}
// Create the LCCCS instances out of z_lcccs
let mut lcccs_instances = Vec::new();
let mut w_lcccs = Vec::new();
for z_i in z_lcccs.iter() {
let (running_instance, w) = ccs.to_lcccs(&mut rng, &pedersen_params, z_i).unwrap();
lcccs_instances.push(running_instance);
w_lcccs.push(w);
}
// Create the CCCS instance out of z_cccs
let mut cccs_instances = Vec::new();
let mut w_cccs = Vec::new();
for z_i in z_cccs.iter() {
let (new_instance, w) = ccs.to_cccs(&mut rng, &pedersen_params, z_i).unwrap();
cccs_instances.push(new_instance);
w_cccs.push(w);
}
// Run the prover side of the multifolding
let (proof, folded_lcccs, folded_witness) =
NIMFS::<Projective, PoseidonTranscript<Projective>>::prove(
&mut transcript_p,
&ccs,
&lcccs_instances,
&cccs_instances,
&w_lcccs,
&w_cccs,
)
.unwrap();
// Run the verifier side of the multifolding
let folded_lcccs_v = NIMFS::<Projective, PoseidonTranscript<Projective>>::verify(
&mut transcript_v,
&ccs,
&lcccs_instances,
&cccs_instances,
proof,
)
.unwrap();
assert_eq!(folded_lcccs, folded_lcccs_v);
// Check that the folded LCCCS instance is a valid instance with respect to the folded witness
folded_lcccs
.check_relation(&pedersen_params, &ccs, &folded_witness)
.unwrap();
}
}
}

View File

@@ -0,0 +1,382 @@
use ark_ec::CurveGroup;
use ark_ff::{Field, PrimeField};
use ark_poly::DenseMultilinearExtension;
use ark_poly::MultilinearExtension;
use ark_std::{One, Zero};
use std::ops::Add;
use crate::utils::multilinear_polynomial::fix_variables;
use crate::utils::multilinear_polynomial::scalar_mul;
use super::lcccs::LCCCS;
use super::nimfs::SigmasThetas;
use crate::ccs::CCS;
use crate::utils::hypercube::BooleanHypercube;
use crate::utils::mle::dense_vec_to_mle;
use crate::utils::mle::matrix_to_mle;
use crate::utils::vec::SparseMatrix;
use crate::utils::virtual_polynomial::{eq_eval, VirtualPolynomial};
/// Return a vector of evaluations p_j(r) = \sum_{y \in {0,1}^s'} M_j(r, y) * z(y) for all j values
/// in 0..self.t
pub fn compute_all_sum_Mz_evals<F: PrimeField>(
vec_M: &[SparseMatrix<F>],
z: &Vec<F>,
r: &[F],
s_prime: usize,
) -> Vec<F> {
// Convert z to MLE
let z_y_mle = dense_vec_to_mle(s_prime, z);
// Convert all matrices to MLE
let M_x_y_mle: Vec<DenseMultilinearExtension<F>> =
vec_M.iter().cloned().map(matrix_to_mle).collect();
let mut v = Vec::with_capacity(M_x_y_mle.len());
for M_i in M_x_y_mle {
let sum_Mz = compute_sum_Mz(M_i, &z_y_mle, s_prime);
let v_i = sum_Mz.evaluate(r).unwrap();
v.push(v_i);
}
v
}
/// Return the multilinear polynomial p(x) = \sum_{y \in {0,1}^s'} M_j(x, y) * z(y)
pub fn compute_sum_Mz<F: PrimeField>(
M_j: DenseMultilinearExtension<F>,
z: &DenseMultilinearExtension<F>,
s_prime: usize,
) -> DenseMultilinearExtension<F> {
let mut sum_Mz = DenseMultilinearExtension {
evaluations: vec![F::zero(); M_j.evaluations.len()],
num_vars: M_j.num_vars - s_prime,
};
let bhc = BooleanHypercube::new(s_prime);
for y in bhc.into_iter() {
// In a slightly counter-intuitive fashion fix_variables() fixes the right-most variables of the polynomial. So
// for a polynomial M(x,y) and a random field element r, if we do fix_variables(M,r) we will get M(x,r).
let M_j_y = fix_variables(&M_j, &y);
let z_y = z.evaluate(&y).unwrap();
let M_j_z = scalar_mul(&M_j_y, &z_y);
sum_Mz = sum_Mz.add(M_j_z);
}
sum_Mz
}
/// Compute the arrays of sigma_i and theta_i from step 4 corresponding to the LCCCS and CCCS
/// instances
pub fn compute_sigmas_and_thetas<C: CurveGroup>(
ccs: &CCS<C>,
z_lcccs: &[Vec<C::ScalarField>],
z_cccs: &[Vec<C::ScalarField>],
r_x_prime: &[C::ScalarField],
) -> SigmasThetas<C::ScalarField> {
let mut sigmas: Vec<Vec<C::ScalarField>> = Vec::new();
for z_lcccs_i in z_lcccs {
// sigmas
let sigma_i = compute_all_sum_Mz_evals(&ccs.M, z_lcccs_i, r_x_prime, ccs.s_prime);
sigmas.push(sigma_i);
}
let mut thetas: Vec<Vec<C::ScalarField>> = Vec::new();
for z_cccs_i in z_cccs {
// thetas
let theta_i = compute_all_sum_Mz_evals(&ccs.M, z_cccs_i, r_x_prime, ccs.s_prime);
thetas.push(theta_i);
}
SigmasThetas(sigmas, thetas)
}
/// Computes the sum $\sum_{j = 0}^{n} \gamma^{\text{pow} + j} \cdot eq_eval \cdot \sigma_{j}$
/// `pow` corresponds to `i * ccs.t` in `compute_c_from_sigmas_and_thetas`
pub fn sum_muls_gamma_pows_eq_sigma<F: PrimeField>(
gamma: F,
eq_eval: F,
sigmas: &[F],
pow: u64,
) -> F {
let mut result = F::zero();
for (j, sigma_j) in sigmas.iter().enumerate() {
let gamma_j = gamma.pow([(pow + (j as u64))]);
result += gamma_j * eq_eval * sigma_j;
}
result
}
/// Computes $\sum_{i=1}^{q} c_i * \prod_{j \in S_i} theta_j$
pub fn sum_ci_mul_prod_thetaj<C: CurveGroup>(
ccs: &CCS<C>,
thetas: &[C::ScalarField],
) -> C::ScalarField {
let mut result = C::ScalarField::zero();
for i in 0..ccs.q {
let mut prod = C::ScalarField::one();
for j in ccs.S[i].clone() {
prod *= thetas[j];
}
result += ccs.c[i] * prod;
}
result
}
/// Compute the right-hand-side of step 5 of the multifolding scheme
pub fn compute_c_from_sigmas_and_thetas<C: CurveGroup>(
ccs: &CCS<C>,
st: &SigmasThetas<C::ScalarField>,
gamma: C::ScalarField,
beta: &[C::ScalarField],
vec_r_x: &Vec<Vec<C::ScalarField>>,
r_x_prime: &[C::ScalarField],
) -> C::ScalarField {
let (vec_sigmas, vec_thetas) = (st.0.clone(), st.1.clone());
let mut c = C::ScalarField::zero();
let mut e_lcccs = Vec::new();
for r_x in vec_r_x {
e_lcccs.push(eq_eval(r_x, r_x_prime).unwrap());
}
for (i, sigmas) in vec_sigmas.iter().enumerate() {
// (sum gamma^j * e_i * sigma_j)
c += sum_muls_gamma_pows_eq_sigma(gamma, e_lcccs[i], sigmas, (i * ccs.t) as u64);
}
let mu = vec_sigmas.len();
let e2 = eq_eval(beta, r_x_prime).unwrap();
for (k, thetas) in vec_thetas.iter().enumerate() {
// + gamma^{t+1} * e2 * sum c_i * prod theta_j
let lhs = sum_ci_mul_prod_thetaj(ccs, thetas);
let gamma_t1 = gamma.pow([(mu * ccs.t + k) as u64]);
c += gamma_t1 * e2 * lhs;
}
c
}
/// Compute g(x) polynomial for the given inputs.
pub fn compute_g<C: CurveGroup>(
ccs: &CCS<C>,
running_instances: &[LCCCS<C>],
z_lcccs: &[Vec<C::ScalarField>],
z_cccs: &[Vec<C::ScalarField>],
gamma: C::ScalarField,
beta: &[C::ScalarField],
) -> VirtualPolynomial<C::ScalarField> {
let mu = running_instances.len();
let mut vec_Ls: Vec<VirtualPolynomial<C::ScalarField>> = Vec::new();
for (i, running_instance) in running_instances.iter().enumerate() {
let mut Ls = running_instance.compute_Ls(ccs, &z_lcccs[i]);
vec_Ls.append(&mut Ls);
}
let mut vec_Q: Vec<VirtualPolynomial<C::ScalarField>> = Vec::new();
// for (i, _) in cccs_instances.iter().enumerate() {
for z_cccs_i in z_cccs.iter() {
let Q = ccs.compute_Q(z_cccs_i, beta);
vec_Q.push(Q);
}
let mut g = vec_Ls[0].clone();
// note: the following two loops can be integrated in the previous two loops, but left
// separated for clarity in the PoC implementation.
for (j, L_j) in vec_Ls.iter_mut().enumerate().skip(1) {
let gamma_j = gamma.pow([j as u64]);
L_j.scalar_mul(&gamma_j);
g = g.add(L_j);
}
for (i, Q_i) in vec_Q.iter_mut().enumerate() {
let gamma_mut_i = gamma.pow([(mu * ccs.t + i) as u64]);
Q_i.scalar_mul(&gamma_mut_i);
g = g.add(Q_i);
}
g
}
#[cfg(test)]
pub mod tests {
use super::*;
use ark_pallas::{Fr, Projective};
use ark_std::test_rng;
use ark_std::One;
use ark_std::UniformRand;
use ark_std::Zero;
use crate::ccs::tests::{get_test_ccs, get_test_z};
use crate::commitment::pedersen::Pedersen;
use crate::utils::multilinear_polynomial::tests::fix_last_variables;
use crate::utils::virtual_polynomial::eq_eval;
#[test]
fn test_compute_sum_Mz_over_boolean_hypercube() {
let ccs = get_test_ccs::<Projective>();
let z = get_test_z(3);
ccs.check_relation(&z).unwrap();
let z_mle = dense_vec_to_mle(ccs.s_prime, &z);
// check that evaluating over all the values x over the boolean hypercube, the result of
// the next for loop is equal to 0
for x in BooleanHypercube::new(ccs.s) {
let mut r = Fr::zero();
for i in 0..ccs.q {
let mut Sj_prod = Fr::one();
for j in ccs.S[i].clone() {
let M_j = matrix_to_mle(ccs.M[j].clone());
let sum_Mz = compute_sum_Mz(M_j, &z_mle, ccs.s_prime);
let sum_Mz_x = sum_Mz.evaluate(&x).unwrap();
Sj_prod *= sum_Mz_x;
}
r += Sj_prod * ccs.c[i];
}
assert_eq!(r, Fr::zero());
}
}
/// Given M(x,y) matrix and a random field element `r`, test that ~M(r,y) is is an s'-variable polynomial which
/// compresses every column j of the M(x,y) matrix by performing a random linear combination between the elements
/// of the column and the values eq_i(r) where i is the row of that element
///
/// For example, for matrix M:
///
/// [2, 3, 4, 4
/// 4, 4, 3, 2
/// 2, 8, 9, 2
/// 9, 4, 2, 0]
///
/// The polynomial ~M(r,y) is a polynomial in F^2 which evaluates to the following values in the hypercube:
/// - M(00) = 2*eq_00(r) + 4*eq_10(r) + 2*eq_01(r) + 9*eq_11(r)
/// - M(10) = 3*eq_00(r) + 4*eq_10(r) + 8*eq_01(r) + 4*eq_11(r)
/// - M(01) = 4*eq_00(r) + 3*eq_10(r) + 9*eq_01(r) + 2*eq_11(r)
/// - M(11) = 4*eq_00(r) + 2*eq_10(r) + 2*eq_01(r) + 0*eq_11(r)
///
/// This is used by Hypernova in LCCCS to perform a verifier-chosen random linear combination between the columns
/// of the matrix and the z vector. This technique is also used extensively in "An Algebraic Framework for
/// Universal and Updatable SNARKs".
#[test]
fn test_compute_M_r_y_compression() {
let mut rng = test_rng();
// s = 2, s' = 3
let ccs = get_test_ccs::<Projective>();
let M = ccs.M[0].clone().to_dense();
let M_mle = matrix_to_mle(ccs.M[0].clone());
// Fix the polynomial ~M(r,y)
let r: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
let M_r_y = fix_last_variables(&M_mle, &r);
// compute M_r_y the other way around
for j in 0..M[0].len() {
// Go over every column of M
let column_j: Vec<Fr> = M.clone().iter().map(|x| x[j]).collect();
// and perform the random lincomb between the elements of the column and eq_i(r)
let rlc = BooleanHypercube::new(ccs.s)
.enumerate()
.map(|(i, x)| column_j[i] * eq_eval(&x, &r).unwrap())
.fold(Fr::zero(), |acc, result| acc + result);
assert_eq!(M_r_y.evaluations[j], rlc);
}
}
#[test]
fn test_compute_sigmas_and_thetas() {
let ccs = get_test_ccs();
let z1 = get_test_z(3);
let z2 = get_test_z(4);
ccs.check_relation(&z1).unwrap();
ccs.check_relation(&z2).unwrap();
let mut rng = test_rng();
let gamma: Fr = Fr::rand(&mut rng);
let beta: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
let r_x_prime: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Initialize a multifolding object
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs_instance, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let sigmas_thetas =
compute_sigmas_and_thetas(&ccs, &[z1.clone()], &[z2.clone()], &r_x_prime);
let g = compute_g(
&ccs,
&[lcccs_instance.clone()],
&[z1.clone()],
&[z2.clone()],
gamma,
&beta,
);
// we expect g(r_x_prime) to be equal to:
// c = (sum gamma^j * e1 * sigma_j) + gamma^{t+1} * e2 * sum c_i * prod theta_j
// from compute_c_from_sigmas_and_thetas
let expected_c = g.evaluate(&r_x_prime).unwrap();
let c = compute_c_from_sigmas_and_thetas::<Projective>(
&ccs,
&sigmas_thetas,
gamma,
&beta,
&vec![lcccs_instance.r_x],
&r_x_prime,
);
assert_eq!(c, expected_c);
}
#[test]
fn test_compute_g() {
let ccs = get_test_ccs();
let z1 = get_test_z(3);
let z2 = get_test_z(4);
ccs.check_relation(&z1).unwrap();
ccs.check_relation(&z2).unwrap();
let mut rng = test_rng(); // TMP
let gamma: Fr = Fr::rand(&mut rng);
let beta: Vec<Fr> = (0..ccs.s).map(|_| Fr::rand(&mut rng)).collect();
// Initialize a multifolding object
let pedersen_params = Pedersen::new_params(&mut rng, ccs.n - ccs.l - 1);
let (lcccs_instance, _) = ccs.to_lcccs(&mut rng, &pedersen_params, &z1).unwrap();
let mut sum_v_j_gamma = Fr::zero();
for j in 0..lcccs_instance.v.len() {
let gamma_j = gamma.pow([j as u64]);
sum_v_j_gamma += lcccs_instance.v[j] * gamma_j;
}
// Compute g(x) with that r_x
let g = compute_g::<Projective>(
&ccs,
&[lcccs_instance.clone()],
&[z1.clone()],
&[z2.clone()],
gamma,
&beta,
);
// evaluate g(x) over x \in {0,1}^s
let mut g_on_bhc = Fr::zero();
for x in BooleanHypercube::new(ccs.s) {
g_on_bhc += g.evaluate(&x).unwrap();
}
// evaluate sum_{j \in [t]} (gamma^j * Lj(x)) over x \in {0,1}^s
let mut sum_Lj_on_bhc = Fr::zero();
let vec_L = lcccs_instance.compute_Ls(&ccs, &z1);
for x in BooleanHypercube::new(ccs.s) {
for (j, Lj) in vec_L.iter().enumerate() {
let gamma_j = gamma.pow([j as u64]);
sum_Lj_on_bhc += Lj.evaluate(&x).unwrap() * gamma_j;
}
}
// Q(x) over bhc is assumed to be zero, as checked in the test 'test_compute_Q'
assert_ne!(g_on_bhc, Fr::zero());
// evaluating g(x) over the boolean hypercube should give the same result as evaluating the
// sum of gamma^j * Lj(x) over the boolean hypercube
assert_eq!(g_on_bhc, sum_Lj_on_bhc);
// evaluating g(x) over the boolean hypercube should give the same result as evaluating the
// sum of gamma^j * v_j over j \in [t]
assert_eq!(g_on_bhc, sum_v_j_gamma);
}
}