* Move r1cs and ccs to standalone folders * Simplify type bounds of SparseMatrixVar * Implement `EquivalenceGadget` trait for `FpVar` and `NonNativeUintVar`. Together with the existing `MatrixGadget` and `VectorGadget`, we can now use the same logic for checking R1CS satisfiability of `R1CSVar` both natively and non-natively. * Simplify trait bounds * Implement `ArithGadget` for `R1CSMatricesVar` and `CCSMatricesVar` * `PedersenGadget::commit` now takes slices as input * Structs for proofs and auxiliary values in protogalaxy * `u` in LCCCS should be `z[0]` * `Inputize` trait * Generic decider circuits * Verifier should check the commitments in committed instances * Update the comments according to the new docs * Fix examples * Add `DeciderEnabledNIFS::fold_group_elements_native` to wrap code for folding commitments * Fix incorrect endian * Format * Get rid of `unwrap` when possiblemain
@ -0,0 +1,35 @@ |
|||
use super::CCS;
|
|||
use crate::utils::gadgets::SparseMatrixVar;
|
|||
use ark_ff::PrimeField;
|
|||
use ark_r1cs_std::{
|
|||
alloc::{AllocVar, AllocationMode},
|
|||
fields::fp::FpVar,
|
|||
};
|
|||
use ark_relations::r1cs::{Namespace, SynthesisError};
|
|||
use ark_std::borrow::Borrow;
|
|||
/// CCSMatricesVar contains the matrices 'M' of the CCS without the rest of CCS parameters.
|
|||
///
|
|||
#[derive(Debug, Clone)]
|
|||
pub struct CCSMatricesVar<F: PrimeField> {
|
|||
// we only need native representation, so the constraint field==F
|
|||
pub M: Vec<SparseMatrixVar<FpVar<F>>>,
|
|||
}
|
|||
|
|||
impl<F: PrimeField> AllocVar<CCS<F>, F> for CCSMatricesVar<F> {
|
|||
fn new_variable<T: Borrow<CCS<F>>>(
|
|||
cs: impl Into<Namespace<F>>,
|
|||
f: impl FnOnce() -> Result<T, SynthesisError>,
|
|||
_mode: AllocationMode,
|
|||
) -> Result<Self, SynthesisError> {
|
|||
f().and_then(|val| {
|
|||
let cs = cs.into();
|
|||
let M: Vec<SparseMatrixVar<FpVar<F>>> = val
|
|||
.borrow()
|
|||
.M
|
|||
.iter()
|
|||
.map(|M| SparseMatrixVar::<FpVar<F>>::new_constant(cs.clone(), M.clone()))
|
|||
.collect::<Result<_, SynthesisError>>()?;
|
|||
Ok(Self { M })
|
|||
})
|
|||
}
|
|||
}
|
@ -0,0 +1,301 @@ |
|||
use crate::{
|
|||
arith::ArithGadget,
|
|||
utils::gadgets::{EquivalenceGadget, MatrixGadget, SparseMatrixVar, VectorGadget},
|
|||
};
|
|||
use ark_ff::PrimeField;
|
|||
use ark_r1cs_std::alloc::{AllocVar, AllocationMode};
|
|||
use ark_relations::r1cs::{Namespace, SynthesisError};
|
|||
use ark_std::{borrow::Borrow, marker::PhantomData, One};
|
|||
|
|||
use super::R1CS;
|
|||
|
|||
/// An in-circuit representation of the `R1CS` struct.
|
|||
///
|
|||
/// `M` is for the modulo operation involved in the satisfiability check when
|
|||
/// the underlying `FVar` is `NonNativeUintVar`.
|
|||
#[derive(Debug, Clone)]
|
|||
pub struct R1CSMatricesVar<M, FVar> {
|
|||
_m: PhantomData<M>,
|
|||
pub A: SparseMatrixVar<FVar>,
|
|||
pub B: SparseMatrixVar<FVar>,
|
|||
pub C: SparseMatrixVar<FVar>,
|
|||
}
|
|||
|
|||
impl<F: PrimeField, ConstraintF: PrimeField, FVar: AllocVar<F, ConstraintF>>
|
|||
AllocVar<R1CS<F>, ConstraintF> for R1CSMatricesVar<F, FVar>
|
|||
{
|
|||
fn new_variable<T: Borrow<R1CS<F>>>(
|
|||
cs: impl Into<Namespace<ConstraintF>>,
|
|||
f: impl FnOnce() -> Result<T, SynthesisError>,
|
|||
_mode: AllocationMode,
|
|||
) -> Result<Self, SynthesisError> {
|
|||
f().and_then(|val| {
|
|||
let cs = cs.into();
|
|||
|
|||
Ok(Self {
|
|||
_m: PhantomData,
|
|||
A: SparseMatrixVar::<FVar>::new_constant(cs.clone(), &val.borrow().A)?,
|
|||
B: SparseMatrixVar::<FVar>::new_constant(cs.clone(), &val.borrow().B)?,
|
|||
C: SparseMatrixVar::<FVar>::new_constant(cs.clone(), &val.borrow().C)?,
|
|||
})
|
|||
})
|
|||
}
|
|||
}
|
|||
|
|||
impl<M, FVar> R1CSMatricesVar<M, FVar>
|
|||
where
|
|||
SparseMatrixVar<FVar>: MatrixGadget<FVar>,
|
|||
[FVar]: VectorGadget<FVar>,
|
|||
{
|
|||
pub fn eval_at_z(&self, z: &[FVar]) -> Result<(Vec<FVar>, Vec<FVar>), SynthesisError> {
|
|||
// Multiply Cz by z[0] (u) here, allowing this method to be reused for
|
|||
// both relaxed and unrelaxed R1CS.
|
|||
let Az = self.A.mul_vector(z)?;
|
|||
let Bz = self.B.mul_vector(z)?;
|
|||
let Cz = self.C.mul_vector(z)?;
|
|||
let uCz = Cz.mul_scalar(&z[0])?;
|
|||
let AzBz = Az.hadamard(&Bz)?;
|
|||
Ok((AzBz, uCz))
|
|||
}
|
|||
}
|
|||
|
|||
impl<M, FVar, WVar: AsRef<[FVar]>, UVar: AsRef<[FVar]>> ArithGadget<WVar, UVar>
|
|||
for R1CSMatricesVar<M, FVar>
|
|||
where
|
|||
SparseMatrixVar<FVar>: MatrixGadget<FVar>,
|
|||
[FVar]: VectorGadget<FVar> + EquivalenceGadget<M>,
|
|||
FVar: Clone + One,
|
|||
{
|
|||
/// Evaluation is a tuple of two vectors (`AzBz` and `uCz`) instead of a
|
|||
/// single vector `AzBz - uCz`, because subtraction is not supported for
|
|||
/// `FVar = NonNativeUintVar`.
|
|||
type Evaluation = (Vec<FVar>, Vec<FVar>);
|
|||
|
|||
fn eval_relation(&self, w: &WVar, u: &UVar) -> Result<Self::Evaluation, SynthesisError> {
|
|||
self.eval_at_z(&[&[FVar::one()], u.as_ref(), w.as_ref()].concat())
|
|||
}
|
|||
|
|||
fn enforce_evaluation(
|
|||
_w: &WVar,
|
|||
_u: &UVar,
|
|||
(lhs, rhs): Self::Evaluation,
|
|||
) -> Result<(), SynthesisError> {
|
|||
lhs.enforce_equivalent(&rhs)
|
|||
}
|
|||
}
|
|||
|
|||
#[cfg(test)]
|
|||
pub mod tests {
|
|||
use std::cmp::max;
|
|||
|
|||
use ark_crypto_primitives::crh::{
|
|||
sha256::{
|
|||
constraints::{Sha256Gadget, UnitVar},
|
|||
Sha256,
|
|||
},
|
|||
CRHScheme, CRHSchemeGadget,
|
|||
};
|
|||
use ark_ec::CurveGroup;
|
|||
use ark_ff::BigInteger;
|
|||
use ark_pallas::{Fq, Fr, Projective};
|
|||
use ark_r1cs_std::{bits::uint8::UInt8, eq::EqGadget, fields::fp::FpVar};
|
|||
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystem, ConstraintSystemRef};
|
|||
use ark_std::{
|
|||
rand::{thread_rng, Rng},
|
|||
One, UniformRand,
|
|||
};
|
|||
use ark_vesta::{constraints::GVar as GVar2, Projective as Projective2};
|
|||
|
|||
use super::*;
|
|||
use crate::arith::{
|
|||
r1cs::{
|
|||
extract_r1cs, extract_w_x,
|
|||
tests::{get_test_r1cs, get_test_z},
|
|||
},
|
|||
Arith,
|
|||
};
|
|||
use crate::commitment::{pedersen::Pedersen, CommitmentScheme};
|
|||
use crate::folding::{
|
|||
circuits::{
|
|||
cyclefold::{CycleFoldCommittedInstanceVar, CycleFoldWitnessVar},
|
|||
nonnative::uint::NonNativeUintVar,
|
|||
},
|
|||
nova::{
|
|||
circuits::CommittedInstanceVar, decider_eth_circuit::WitnessVar, CommittedInstance,
|
|||
Witness,
|
|||
},
|
|||
};
|
|||
use crate::frontend::{
|
|||
utils::{CubicFCircuit, CustomFCircuit, WrapperCircuit},
|
|||
FCircuit,
|
|||
};
|
|||
|
|||
fn prepare_instances<C: CurveGroup, CS: CommitmentScheme<C>, R: Rng>(
|
|||
mut rng: R,
|
|||
r1cs: &R1CS<C::ScalarField>,
|
|||
z: &[C::ScalarField],
|
|||
) -> (Witness<C>, CommittedInstance<C>) {
|
|||
let (w, x) = r1cs.split_z(z);
|
|||
|
|||
let (cs_pp, _) = CS::setup(&mut rng, max(w.len(), r1cs.A.n_rows)).unwrap();
|
|||
|
|||
let mut w = Witness::new::<false>(w, r1cs.A.n_rows, &mut rng);
|
|||
w.E = r1cs.eval_at_z(z).unwrap();
|
|||
let mut u = w.commit::<CS, false>(&cs_pp, x).unwrap();
|
|||
u.u = z[0];
|
|||
|
|||
(w, u)
|
|||
}
|
|||
|
|||
#[test]
|
|||
fn test_relaxed_r1cs_small_gadget_handcrafted() {
|
|||
let rng = &mut thread_rng();
|
|||
|
|||
let r1cs: R1CS<Fr> = get_test_r1cs();
|
|||
let mut z = get_test_z(3);
|
|||
z[0] = Fr::rand(rng);
|
|||
let (w, u) = prepare_instances::<_, Pedersen<Projective>, _>(rng, &r1cs, &z);
|
|||
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
|
|||
let wVar = WitnessVar::new_witness(cs.clone(), || Ok(w)).unwrap();
|
|||
let uVar = CommittedInstanceVar::new_witness(cs.clone(), || Ok(u)).unwrap();
|
|||
let r1csVar =
|
|||
R1CSMatricesVar::<Fr, FpVar<Fr>>::new_witness(cs.clone(), || Ok(r1cs)).unwrap();
|
|||
|
|||
r1csVar.enforce_relation(&wVar, &uVar).unwrap();
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
}
|
|||
|
|||
// gets as input a circuit that implements the ConstraintSynthesizer trait, and that has been
|
|||
// initialized.
|
|||
fn test_relaxed_r1cs_gadget<CS: ConstraintSynthesizer<Fr>>(circuit: CS) {
|
|||
let rng = &mut thread_rng();
|
|||
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
|
|||
circuit.generate_constraints(cs.clone()).unwrap();
|
|||
cs.finalize();
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
|
|||
let cs = cs.into_inner().unwrap();
|
|||
|
|||
let r1cs = extract_r1cs::<Fr>(&cs).unwrap();
|
|||
let (w, x) = extract_w_x::<Fr>(&cs);
|
|||
r1cs.check_relation(&w, &x).unwrap();
|
|||
let mut z = [vec![Fr::one()], x, w].concat();
|
|||
z[0] = Fr::rand(rng);
|
|||
|
|||
let (w, u) = prepare_instances::<_, Pedersen<Projective>, _>(rng, &r1cs, &z);
|
|||
r1cs.check_relation(&w, &u).unwrap();
|
|||
|
|||
// set new CS for the circuit that checks the RelaxedR1CS of our original circuit
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
// prepare the inputs for our circuit
|
|||
let wVar = WitnessVar::new_witness(cs.clone(), || Ok(w)).unwrap();
|
|||
let uVar = CommittedInstanceVar::new_witness(cs.clone(), || Ok(u)).unwrap();
|
|||
let r1csVar =
|
|||
R1CSMatricesVar::<Fr, FpVar<Fr>>::new_witness(cs.clone(), || Ok(r1cs)).unwrap();
|
|||
|
|||
r1csVar.enforce_relation(&wVar, &uVar).unwrap();
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
}
|
|||
|
|||
#[test]
|
|||
fn test_relaxed_r1cs_small_gadget_arkworks() {
|
|||
let z_i = vec![Fr::from(3_u32)];
|
|||
let cubic_circuit = CubicFCircuit::<Fr>::new(()).unwrap();
|
|||
let circuit = WrapperCircuit::<Fr, CubicFCircuit<Fr>> {
|
|||
FC: cubic_circuit,
|
|||
z_i: Some(z_i.clone()),
|
|||
z_i1: Some(cubic_circuit.step_native(0, z_i, vec![]).unwrap()),
|
|||
};
|
|||
|
|||
test_relaxed_r1cs_gadget(circuit);
|
|||
}
|
|||
|
|||
struct Sha256TestCircuit<F: PrimeField> {
|
|||
_f: PhantomData<F>,
|
|||
pub x: Vec<u8>,
|
|||
pub y: Vec<u8>,
|
|||
}
|
|||
impl<F: PrimeField> ConstraintSynthesizer<F> for Sha256TestCircuit<F> {
|
|||
fn generate_constraints(self, cs: ConstraintSystemRef<F>) -> Result<(), SynthesisError> {
|
|||
let x = Vec::<UInt8<F>>::new_witness(cs.clone(), || Ok(self.x))?;
|
|||
let y = Vec::<UInt8<F>>::new_input(cs.clone(), || Ok(self.y))?;
|
|||
|
|||
let unitVar = UnitVar::default();
|
|||
let comp_y = <Sha256Gadget<F> as CRHSchemeGadget<Sha256, F>>::evaluate(&unitVar, &x)?;
|
|||
comp_y.0.enforce_equal(&y)?;
|
|||
Ok(())
|
|||
}
|
|||
}
|
|||
#[test]
|
|||
fn test_relaxed_r1cs_medium_gadget_arkworks() {
|
|||
let x = Fr::from(5_u32).into_bigint().to_bytes_le();
|
|||
let y = <Sha256 as CRHScheme>::evaluate(&(), x.clone()).unwrap();
|
|||
|
|||
let circuit = Sha256TestCircuit::<Fr> {
|
|||
_f: PhantomData,
|
|||
x,
|
|||
y,
|
|||
};
|
|||
test_relaxed_r1cs_gadget(circuit);
|
|||
}
|
|||
|
|||
#[test]
|
|||
fn test_relaxed_r1cs_custom_circuit() {
|
|||
let n_constraints = 10_000;
|
|||
let custom_circuit = CustomFCircuit::<Fr>::new(n_constraints).unwrap();
|
|||
let z_i = vec![Fr::from(5_u32)];
|
|||
let circuit = WrapperCircuit::<Fr, CustomFCircuit<Fr>> {
|
|||
FC: custom_circuit,
|
|||
z_i: Some(z_i.clone()),
|
|||
z_i1: Some(custom_circuit.step_native(0, z_i, vec![]).unwrap()),
|
|||
};
|
|||
test_relaxed_r1cs_gadget(circuit);
|
|||
}
|
|||
|
|||
#[test]
|
|||
fn test_relaxed_r1cs_nonnative_circuit() {
|
|||
let rng = &mut thread_rng();
|
|||
|
|||
let cs = ConstraintSystem::<Fq>::new_ref();
|
|||
// in practice we would use CycleFoldCircuit, but is a very big circuit (when computed
|
|||
// non-natively inside the RelaxedR1CS circuit), so in order to have a short test we use a
|
|||
// custom circuit.
|
|||
let custom_circuit = CustomFCircuit::<Fq>::new(10).unwrap();
|
|||
let z_i = vec![Fq::from(5_u32)];
|
|||
let circuit = WrapperCircuit::<Fq, CustomFCircuit<Fq>> {
|
|||
FC: custom_circuit,
|
|||
z_i: Some(z_i.clone()),
|
|||
z_i1: Some(custom_circuit.step_native(0, z_i, vec![]).unwrap()),
|
|||
};
|
|||
circuit.generate_constraints(cs.clone()).unwrap();
|
|||
cs.finalize();
|
|||
let cs = cs.into_inner().unwrap();
|
|||
let r1cs = extract_r1cs::<Fq>(&cs).unwrap();
|
|||
let (w, x) = extract_w_x::<Fq>(&cs);
|
|||
let z = [vec![Fq::rand(rng)], x, w].concat();
|
|||
|
|||
let (w, u) = prepare_instances::<_, Pedersen<Projective2>, _>(rng, &r1cs, &z);
|
|||
|
|||
// natively
|
|||
let cs = ConstraintSystem::<Fq>::new_ref();
|
|||
let wVar = WitnessVar::new_witness(cs.clone(), || Ok(&w)).unwrap();
|
|||
let uVar = CommittedInstanceVar::new_witness(cs.clone(), || Ok(&u)).unwrap();
|
|||
let r1csVar =
|
|||
R1CSMatricesVar::<Fq, FpVar<Fq>>::new_witness(cs.clone(), || Ok(&r1cs)).unwrap();
|
|||
r1csVar.enforce_relation(&wVar, &uVar).unwrap();
|
|||
|
|||
// non-natively
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
let wVar = CycleFoldWitnessVar::new_witness(cs.clone(), || Ok(w)).unwrap();
|
|||
let uVar =
|
|||
CycleFoldCommittedInstanceVar::<_, GVar2>::new_witness(cs.clone(), || Ok(u)).unwrap();
|
|||
let r1csVar =
|
|||
R1CSMatricesVar::<Fq, NonNativeUintVar<Fr>>::new_witness(cs.clone(), || Ok(r1cs))
|
|||
.unwrap();
|
|||
r1csVar.enforce_relation(&wVar, &uVar).unwrap();
|
|||
}
|
|||
}
|
@ -0,0 +1,211 @@ |
|||
use ark_crypto_primitives::sponge::{
|
|||
poseidon::constraints::PoseidonSpongeVar, CryptographicSponge,
|
|||
};
|
|||
use ark_ec::CurveGroup;
|
|||
use ark_ff::PrimeField;
|
|||
use ark_poly::Polynomial;
|
|||
use ark_r1cs_std::{
|
|||
fields::{fp::FpVar, FieldVar},
|
|||
poly::{domain::Radix2DomainVar, evaluations::univariate::EvaluationsVar},
|
|||
ToConstraintFieldGadget,
|
|||
};
|
|||
use ark_relations::r1cs::SynthesisError;
|
|||
use ark_std::log2;
|
|||
|
|||
use crate::folding::traits::{CommittedInstanceOps, CommittedInstanceVarOps, Dummy, WitnessOps};
|
|||
use crate::transcript::{Transcript, TranscriptVar};
|
|||
use crate::utils::vec::poly_from_vec;
|
|||
use crate::Error;
|
|||
use crate::{arith::Arith, folding::circuits::CF1};
|
|||
|
|||
pub mod off_chain;
|
|||
pub mod on_chain;
|
|||
|
|||
/// Gadget that computes the KZG challenges.
|
|||
/// It also offers the rust native implementation compatible with the gadget.
|
|||
pub struct KZGChallengesGadget {}
|
|||
|
|||
impl KZGChallengesGadget {
|
|||
pub fn get_challenges_native<
|
|||
C: CurveGroup,
|
|||
T: Transcript<CF1<C>>,
|
|||
U: CommittedInstanceOps<C>,
|
|||
>(
|
|||
transcript: &mut T,
|
|||
U_i: &U,
|
|||
) -> Vec<CF1<C>> {
|
|||
let mut challenges = vec![];
|
|||
for cm in U_i.get_commitments() {
|
|||
transcript.absorb_nonnative(&cm);
|
|||
challenges.push(transcript.get_challenge());
|
|||
}
|
|||
challenges
|
|||
}
|
|||
|
|||
pub fn get_challenges_gadget<
|
|||
C: CurveGroup,
|
|||
S: CryptographicSponge,
|
|||
T: TranscriptVar<CF1<C>, S>,
|
|||
U: CommittedInstanceVarOps<C>,
|
|||
>(
|
|||
transcript: &mut T,
|
|||
U_i: &U,
|
|||
) -> Result<Vec<FpVar<CF1<C>>>, SynthesisError> {
|
|||
let mut challenges = vec![];
|
|||
for cm in U_i.get_commitments() {
|
|||
transcript.absorb(&cm.to_constraint_field()?)?;
|
|||
challenges.push(transcript.get_challenge()?);
|
|||
}
|
|||
Ok(challenges)
|
|||
}
|
|||
}
|
|||
|
|||
/// Gadget that interpolates the polynomial from the given vector and returns
|
|||
/// its evaluation at the given point.
|
|||
/// It also offers the rust native implementation compatible with the gadget.
|
|||
pub struct EvalGadget {}
|
|||
|
|||
impl EvalGadget {
|
|||
pub fn evaluate_native<F: PrimeField>(v: &[F], point: F) -> Result<F, Error> {
|
|||
let mut v = v.to_vec();
|
|||
v.resize(v.len().next_power_of_two(), F::zero());
|
|||
|
|||
Ok(poly_from_vec(v)?.evaluate(&point))
|
|||
}
|
|||
|
|||
pub fn evaluate_gadget<F: PrimeField>(
|
|||
v: &[FpVar<F>],
|
|||
point: &FpVar<F>,
|
|||
) -> Result<FpVar<F>, SynthesisError> {
|
|||
let mut v = v.to_vec();
|
|||
v.resize(v.len().next_power_of_two(), FpVar::zero());
|
|||
let n = v.len() as u64;
|
|||
let gen = F::get_root_of_unity(n).ok_or(SynthesisError::PolynomialDegreeTooLarge)?;
|
|||
// `unwrap` below is safe because `Radix2DomainVar::new` only fails if
|
|||
// `offset.enforce_not_equal(&FpVar::zero())` returns an error.
|
|||
// But in our case, `offset` is `FpVar::one()`, i.e., both operands of
|
|||
// `enforce_not_equal` are constants.
|
|||
// Consequently, `FpVar`'s implementation of `enforce_not_equal` will
|
|||
// always return `Ok(())`.
|
|||
let domain = Radix2DomainVar::new(gen, log2(v.len()) as u64, FpVar::one()).unwrap();
|
|||
|
|||
let evaluations_var = EvaluationsVar::from_vec_and_domain(v, domain, true);
|
|||
evaluations_var.interpolate_and_evaluate(point)
|
|||
}
|
|||
}
|
|||
|
|||
/// This is a temporary workaround for step 6 (running NIFS.V for group elements
|
|||
/// in circuit) in an NIFS-agnostic way, because different folding schemes have
|
|||
/// different interfaces of folding verification now.
|
|||
///
|
|||
/// In the future, we may introduce a better solution that uses a trait for all
|
|||
/// folding schemes that specifies their native and in-circuit behaviors.
|
|||
pub trait DeciderEnabledNIFS<
|
|||
C: CurveGroup,
|
|||
RU: CommittedInstanceOps<C>, // Running instance
|
|||
IU: CommittedInstanceOps<C>, // Incoming instance
|
|||
W: WitnessOps<CF1<C>>,
|
|||
A: Arith<W, RU>,
|
|||
>
|
|||
{
|
|||
type ProofDummyCfg;
|
|||
type Proof: Dummy<Self::ProofDummyCfg>;
|
|||
type RandomnessDummyCfg;
|
|||
type Randomness: Dummy<Self::RandomnessDummyCfg>;
|
|||
|
|||
/// Fold the field elements in `U` and `u` inside the circuit.
|
|||
///
|
|||
/// `U_vec` is `U` expressed as a vector of `FpVar`s, which can be reused
|
|||
/// before or after calling this function to save constraints.
|
|||
#[allow(clippy::too_many_arguments)]
|
|||
fn fold_field_elements_gadget(
|
|||
arith: &A,
|
|||
transcript: &mut PoseidonSpongeVar<CF1<C>>,
|
|||
pp_hash: FpVar<CF1<C>>,
|
|||
U: RU::Var,
|
|||
U_vec: Vec<FpVar<CF1<C>>>,
|
|||
u: IU::Var,
|
|||
proof: Self::Proof,
|
|||
randomness: Self::Randomness,
|
|||
) -> Result<RU::Var, SynthesisError>;
|
|||
|
|||
/// Fold the group elements (i.e., commitments) in `U` and `u` outside the
|
|||
/// circuit.
|
|||
fn fold_group_elements_native(
|
|||
U_commitments: &[C],
|
|||
u_commitments: &[C],
|
|||
proof: Option<Self::Proof>,
|
|||
randomness: Self::Randomness,
|
|||
) -> Result<Vec<C>, Error>;
|
|||
}
|
|||
|
|||
#[cfg(test)]
|
|||
pub mod tests {
|
|||
use ark_crypto_primitives::sponge::{
|
|||
constraints::CryptographicSpongeVar, poseidon::PoseidonSponge,
|
|||
};
|
|||
use ark_pallas::{Fr, Projective};
|
|||
use ark_r1cs_std::{alloc::AllocVar, R1CSVar};
|
|||
use ark_relations::r1cs::ConstraintSystem;
|
|||
use ark_std::UniformRand;
|
|||
|
|||
use super::*;
|
|||
use crate::folding::nova::{circuits::CommittedInstanceVar, CommittedInstance};
|
|||
use crate::transcript::poseidon::poseidon_canonical_config;
|
|||
|
|||
// checks that the gadget and native implementations of the challenge computation match
|
|||
#[test]
|
|||
fn test_kzg_challenge_gadget() {
|
|||
let mut rng = ark_std::test_rng();
|
|||
let poseidon_config = poseidon_canonical_config::<Fr>();
|
|||
let mut transcript = PoseidonSponge::<Fr>::new(&poseidon_config);
|
|||
|
|||
let U_i = CommittedInstance::<Projective> {
|
|||
cmE: Projective::rand(&mut rng),
|
|||
u: Fr::rand(&mut rng),
|
|||
cmW: Projective::rand(&mut rng),
|
|||
x: vec![Fr::rand(&mut rng); 1],
|
|||
};
|
|||
|
|||
// compute the challenge natively
|
|||
let challenges = KZGChallengesGadget::get_challenges_native(&mut transcript, &U_i);
|
|||
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
let U_iVar =
|
|||
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(U_i.clone()))
|
|||
.unwrap();
|
|||
let mut transcript_var = PoseidonSpongeVar::<Fr>::new(cs.clone(), &poseidon_config);
|
|||
|
|||
let challenges_var =
|
|||
KZGChallengesGadget::get_challenges_gadget(&mut transcript_var, &U_iVar).unwrap();
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
|
|||
// check that the natively computed and in-circuit computed hashes match
|
|||
assert_eq!(challenges_var.value().unwrap(), challenges);
|
|||
}
|
|||
|
|||
#[test]
|
|||
fn test_polynomial_interpolation() {
|
|||
let mut rng = ark_std::test_rng();
|
|||
let n = 12;
|
|||
let l = 1 << n;
|
|||
|
|||
let v: Vec<Fr> = std::iter::repeat_with(|| Fr::rand(&mut rng))
|
|||
.take(l)
|
|||
.collect();
|
|||
let challenge = Fr::rand(&mut rng);
|
|||
|
|||
use ark_poly::Polynomial;
|
|||
let polynomial = poly_from_vec(v.to_vec()).unwrap();
|
|||
let eval = polynomial.evaluate(&challenge);
|
|||
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
let vVar = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(v)).unwrap();
|
|||
let challengeVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(challenge)).unwrap();
|
|||
|
|||
let evalVar = EvalGadget::evaluate_gadget(&vVar, &challengeVar).unwrap();
|
|||
|
|||
assert_eq!(evalVar.value().unwrap(), eval);
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
}
|
|||
}
|
@ -0,0 +1,316 @@ |
|||
/// This file implements a generic offchain decider circuit.
|
|||
/// For ethereum use cases, use the `GenericOnchainDeciderCircuit`.
|
|||
/// More details can be found at the documentation page:
|
|||
/// https://privacy-scaling-explorations.github.io/sonobe-docs/design/nova-decider-offchain.html
|
|||
use ark_crypto_primitives::sponge::{
|
|||
constraints::{AbsorbGadget, CryptographicSpongeVar},
|
|||
poseidon::{constraints::PoseidonSpongeVar, PoseidonConfig},
|
|||
Absorb,
|
|||
};
|
|||
use ark_ec::CurveGroup;
|
|||
use ark_r1cs_std::{
|
|||
alloc::AllocVar, eq::EqGadget, fields::fp::FpVar, prelude::CurveVar, ToConstraintFieldGadget,
|
|||
};
|
|||
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystemRef, SynthesisError};
|
|||
use ark_std::{marker::PhantomData, Zero};
|
|||
|
|||
use crate::{
|
|||
arith::{
|
|||
r1cs::{circuits::R1CSMatricesVar, R1CS},
|
|||
Arith, ArithGadget,
|
|||
},
|
|||
folding::{
|
|||
circuits::{
|
|||
cyclefold::{
|
|||
CycleFoldCommittedInstance, CycleFoldCommittedInstanceVar, CycleFoldWitness,
|
|||
},
|
|||
decider::{EvalGadget, KZGChallengesGadget},
|
|||
nonnative::affine::NonNativeAffineVar,
|
|||
CF1, CF2,
|
|||
},
|
|||
nova::{circuits::CommittedInstanceVar, decider_eth_circuit::WitnessVar},
|
|||
traits::{CommittedInstanceOps, CommittedInstanceVarOps, Dummy, WitnessOps, WitnessVarOps},
|
|||
},
|
|||
};
|
|||
|
|||
use super::DeciderEnabledNIFS;
|
|||
|
|||
/// Circuit that implements part of the in-circuit checks needed for the offchain verification over
|
|||
/// the Curve2's BaseField (=Curve1's ScalarField).
|
|||
pub struct GenericOffchainDeciderCircuit1<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1>, // Running instance
|
|||
IU: CommittedInstanceOps<C1>, // Incoming instance
|
|||
W: WitnessOps<CF1<C1>>, // Witness
|
|||
A: Arith<W, RU>, // Constraint system
|
|||
AVar: ArithGadget<W::Var, RU::Var>, // In-circuit representation of `A`
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
> {
|
|||
pub _gc2: PhantomData<GC2>,
|
|||
pub _avar: PhantomData<AVar>,
|
|||
/// Constraint system of the Augmented Function circuit
|
|||
pub arith: A,
|
|||
pub poseidon_config: PoseidonConfig<CF1<C1>>,
|
|||
/// public params hash
|
|||
pub pp_hash: CF1<C1>,
|
|||
pub i: CF1<C1>,
|
|||
/// initial state
|
|||
pub z_0: Vec<CF1<C1>>,
|
|||
/// current i-th state
|
|||
pub z_i: Vec<CF1<C1>>,
|
|||
/// Folding scheme instances
|
|||
pub U_i: RU,
|
|||
pub W_i: W,
|
|||
pub u_i: IU,
|
|||
pub w_i: W,
|
|||
pub U_i1: RU,
|
|||
pub W_i1: W,
|
|||
|
|||
/// Helper for folding verification
|
|||
pub proof: D::Proof,
|
|||
pub randomness: D::Randomness,
|
|||
|
|||
/// CycleFold running instance
|
|||
pub cf_U_i: CycleFoldCommittedInstance<C2>,
|
|||
|
|||
/// KZG challenges
|
|||
pub kzg_challenges: Vec<CF1<C1>>,
|
|||
pub kzg_evaluations: Vec<CF1<C1>>,
|
|||
}
|
|||
|
|||
impl<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup<ScalarField = CF2<C1>, BaseField = CF1<C1>>,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1> + for<'a> Dummy<&'a A>,
|
|||
IU: CommittedInstanceOps<C1> + for<'a> Dummy<&'a A>,
|
|||
W: WitnessOps<CF1<C1>> + for<'a> Dummy<&'a A>,
|
|||
A: Arith<W, RU>,
|
|||
AVar: ArithGadget<W::Var, RU::Var> + AllocVar<A, CF1<C1>>,
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
>
|
|||
Dummy<(
|
|||
A,
|
|||
&R1CS<CF1<C2>>,
|
|||
PoseidonConfig<CF1<C1>>,
|
|||
D::ProofDummyCfg,
|
|||
D::RandomnessDummyCfg,
|
|||
usize,
|
|||
usize,
|
|||
)> for GenericOffchainDeciderCircuit1<C1, C2, GC2, RU, IU, W, A, AVar, D>
|
|||
{
|
|||
fn dummy(
|
|||
(
|
|||
arith,
|
|||
cf_arith,
|
|||
poseidon_config,
|
|||
proof_config,
|
|||
randomness_config,
|
|||
state_len,
|
|||
num_commitments,
|
|||
): (
|
|||
A,
|
|||
&R1CS<CF1<C2>>,
|
|||
PoseidonConfig<CF1<C1>>,
|
|||
D::ProofDummyCfg,
|
|||
D::RandomnessDummyCfg,
|
|||
usize,
|
|||
usize,
|
|||
),
|
|||
) -> Self {
|
|||
Self {
|
|||
_gc2: PhantomData,
|
|||
_avar: PhantomData,
|
|||
poseidon_config,
|
|||
pp_hash: Zero::zero(),
|
|||
i: Zero::zero(),
|
|||
z_0: vec![Zero::zero(); state_len],
|
|||
z_i: vec![Zero::zero(); state_len],
|
|||
U_i: RU::dummy(&arith),
|
|||
W_i: W::dummy(&arith),
|
|||
u_i: IU::dummy(&arith),
|
|||
w_i: W::dummy(&arith),
|
|||
U_i1: RU::dummy(&arith),
|
|||
W_i1: W::dummy(&arith),
|
|||
proof: D::Proof::dummy(proof_config),
|
|||
randomness: D::Randomness::dummy(randomness_config),
|
|||
cf_U_i: CycleFoldCommittedInstance::dummy(cf_arith),
|
|||
kzg_challenges: vec![Zero::zero(); num_commitments],
|
|||
kzg_evaluations: vec![Zero::zero(); num_commitments],
|
|||
arith,
|
|||
}
|
|||
}
|
|||
}
|
|||
|
|||
impl<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup<ScalarField = CF2<C1>, BaseField = CF1<C1>>,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1>,
|
|||
IU: CommittedInstanceOps<C1>,
|
|||
W: WitnessOps<CF1<C1>>,
|
|||
A: Arith<W, RU>,
|
|||
AVar: ArithGadget<W::Var, RU::Var> + AllocVar<A, CF1<C1>>,
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
> ConstraintSynthesizer<CF1<C1>>
|
|||
for GenericOffchainDeciderCircuit1<C1, C2, GC2, RU, IU, W, A, AVar, D>
|
|||
where
|
|||
RU::Var: AbsorbGadget<CF1<C1>> + CommittedInstanceVarOps<C1, PointVar = NonNativeAffineVar<C1>>,
|
|||
CF1<C1>: Absorb,
|
|||
{
|
|||
fn generate_constraints(self, cs: ConstraintSystemRef<CF1<C1>>) -> Result<(), SynthesisError> {
|
|||
let arith = AVar::new_witness(cs.clone(), || Ok(&self.arith))?;
|
|||
|
|||
let pp_hash = FpVar::new_input(cs.clone(), || Ok(self.pp_hash))?;
|
|||
let i = FpVar::new_input(cs.clone(), || Ok(self.i))?;
|
|||
let z_0 = Vec::new_input(cs.clone(), || Ok(self.z_0))?;
|
|||
let z_i = Vec::new_input(cs.clone(), || Ok(self.z_i))?;
|
|||
|
|||
let u_i = IU::Var::new_witness(cs.clone(), || Ok(self.u_i))?;
|
|||
let U_i = RU::Var::new_witness(cs.clone(), || Ok(self.U_i))?;
|
|||
// here (U_i1, W_i1) = NIFS.P( (U_i,W_i), (u_i,w_i))
|
|||
let U_i1_commitments = Vec::<NonNativeAffineVar<C1>>::new_input(cs.clone(), || {
|
|||
Ok(self.U_i1.get_commitments())
|
|||
})?;
|
|||
let U_i1 = RU::Var::new_witness(cs.clone(), || Ok(self.U_i1))?;
|
|||
let W_i1 = W::Var::new_witness(cs.clone(), || Ok(self.W_i1))?;
|
|||
U_i1.get_commitments().enforce_equal(&U_i1_commitments)?;
|
|||
|
|||
let cf_U_i =
|
|||
CycleFoldCommittedInstanceVar::<C2, GC2>::new_input(cs.clone(), || Ok(self.cf_U_i))?;
|
|||
|
|||
// allocate the inputs for the checks 7.1 and 7.2
|
|||
let kzg_challenges = Vec::new_input(cs.clone(), || Ok(self.kzg_challenges))?;
|
|||
let kzg_evaluations = Vec::new_input(cs.clone(), || Ok(self.kzg_evaluations))?;
|
|||
|
|||
// `sponge` is for digest computation.
|
|||
let sponge = PoseidonSpongeVar::new(cs.clone(), &self.poseidon_config);
|
|||
// `transcript` is for challenge generation.
|
|||
let mut transcript = sponge.clone();
|
|||
// notice that the `pp_hash` is absorbed inside the ChallengeGadget::get_challenge_gadget call
|
|||
|
|||
// 1. enforce `U_{i+1}` and `W_{i+1}` satisfy `arith`
|
|||
arith.enforce_relation(&W_i1, &U_i1)?;
|
|||
|
|||
// 2. enforce `u_i` is an incoming instance
|
|||
u_i.enforce_incoming()?;
|
|||
|
|||
// 3. u_i.x[0] == H(i, z_0, z_i, U_i), u_i.x[1] == H(cf_U_i)
|
|||
let (u_i_x, U_i_vec) = U_i.hash(&sponge, &pp_hash, &i, &z_0, &z_i)?;
|
|||
let (cf_u_i_x, _) = cf_U_i.hash(&sponge, pp_hash.clone())?;
|
|||
u_i.get_public_inputs().enforce_equal(&[u_i_x, cf_u_i_x])?;
|
|||
|
|||
// 6.1. partially enforce `NIFS.V(U_i, u_i) = U_{i+1}`.
|
|||
D::fold_field_elements_gadget(
|
|||
&self.arith,
|
|||
&mut transcript,
|
|||
pp_hash,
|
|||
U_i,
|
|||
U_i_vec,
|
|||
u_i,
|
|||
self.proof,
|
|||
self.randomness,
|
|||
)?
|
|||
.enforce_partial_equal(&U_i1)?;
|
|||
|
|||
// 7.1. compute and check KZG challenges
|
|||
KZGChallengesGadget::get_challenges_gadget(&mut transcript, &U_i1)?
|
|||
.enforce_equal(&kzg_challenges)?;
|
|||
|
|||
// 7.2. check the claimed evaluations
|
|||
for (((v, _r), c), e) in W_i1
|
|||
.get_openings()
|
|||
.iter()
|
|||
.zip(&kzg_challenges)
|
|||
.zip(&kzg_evaluations)
|
|||
{
|
|||
// The randomness `_r` is currently not used.
|
|||
EvalGadget::evaluate_gadget(v, c)?.enforce_equal(e)?;
|
|||
}
|
|||
|
|||
Ok(())
|
|||
}
|
|||
}
|
|||
|
|||
/// Circuit that implements part of the in-circuit checks needed for the offchain verification over
|
|||
/// the Curve1's BaseField (=Curve2's ScalarField).
|
|||
pub struct GenericOffchainDeciderCircuit2<C2: CurveGroup> {
|
|||
/// R1CS of the CycleFold circuit
|
|||
pub cf_arith: R1CS<CF1<C2>>,
|
|||
pub poseidon_config: PoseidonConfig<CF1<C2>>,
|
|||
/// public params hash
|
|||
pub pp_hash: CF1<C2>,
|
|||
|
|||
/// CycleFold running instance
|
|||
pub cf_U_i: CycleFoldCommittedInstance<C2>,
|
|||
pub cf_W_i: CycleFoldWitness<C2>,
|
|||
|
|||
/// KZG challenges
|
|||
pub kzg_challenges: Vec<CF1<C2>>,
|
|||
pub kzg_evaluations: Vec<CF1<C2>>,
|
|||
}
|
|||
|
|||
impl<C2: CurveGroup> Dummy<(R1CS<CF1<C2>>, PoseidonConfig<CF1<C2>>, usize)>
|
|||
for GenericOffchainDeciderCircuit2<C2>
|
|||
{
|
|||
fn dummy(
|
|||
(cf_arith, poseidon_config, num_commitments): (
|
|||
R1CS<CF1<C2>>,
|
|||
PoseidonConfig<CF1<C2>>,
|
|||
usize,
|
|||
),
|
|||
) -> Self {
|
|||
Self {
|
|||
poseidon_config,
|
|||
pp_hash: Zero::zero(),
|
|||
cf_U_i: CycleFoldCommittedInstance::dummy(&cf_arith),
|
|||
cf_W_i: CycleFoldWitness::dummy(&cf_arith),
|
|||
kzg_challenges: vec![Zero::zero(); num_commitments],
|
|||
kzg_evaluations: vec![Zero::zero(); num_commitments],
|
|||
cf_arith,
|
|||
}
|
|||
}
|
|||
}
|
|||
|
|||
impl<C2: CurveGroup> ConstraintSynthesizer<CF1<C2>> for GenericOffchainDeciderCircuit2<C2> {
|
|||
fn generate_constraints(self, cs: ConstraintSystemRef<CF1<C2>>) -> Result<(), SynthesisError> {
|
|||
let cf_r1cs = R1CSMatricesVar::<CF1<C2>, FpVar<CF1<C2>>>::new_witness(cs.clone(), || {
|
|||
Ok(self.cf_arith.clone())
|
|||
})?;
|
|||
|
|||
let pp_hash = FpVar::new_input(cs.clone(), || Ok(self.pp_hash))?;
|
|||
|
|||
let cf_U_i = CommittedInstanceVar::new_input(cs.clone(), || Ok(self.cf_U_i))?;
|
|||
let cf_W_i = WitnessVar::new_witness(cs.clone(), || Ok(self.cf_W_i))?;
|
|||
|
|||
// allocate the inputs for the checks 4.1 and 4.2
|
|||
let kzg_challenges = Vec::new_input(cs.clone(), || Ok(self.kzg_challenges))?;
|
|||
let kzg_evaluations = Vec::new_input(cs.clone(), || Ok(self.kzg_evaluations))?;
|
|||
|
|||
// `transcript` is for challenge generation.
|
|||
let mut transcript = PoseidonSpongeVar::new(cs.clone(), &self.poseidon_config);
|
|||
transcript.absorb(&pp_hash)?;
|
|||
|
|||
// 5. enforce `cf_U_i` and `cf_W_i` satisfy `cf_r1cs`
|
|||
cf_r1cs.enforce_relation(&cf_W_i, &cf_U_i)?;
|
|||
|
|||
// 4.1. compute and check KZG challenges
|
|||
KZGChallengesGadget::get_challenges_gadget(&mut transcript, &cf_U_i)?
|
|||
.enforce_equal(&kzg_challenges)?;
|
|||
|
|||
// 4.2. check the claimed evaluations
|
|||
for (((v, _r), c), e) in cf_W_i
|
|||
.get_openings()
|
|||
.iter()
|
|||
.zip(&kzg_challenges)
|
|||
.zip(&kzg_evaluations)
|
|||
{
|
|||
// The randomness `_r` is currently not used.
|
|||
EvalGadget::evaluate_gadget(v, c)?.enforce_equal(e)?;
|
|||
}
|
|||
|
|||
Ok(())
|
|||
}
|
|||
}
|
@ -0,0 +1,325 @@ |
|||
/// This file implements the onchain (Ethereum's EVM) decider circuit. For non-ethereum use cases,
|
|||
/// other more efficient approaches can be used.
|
|||
use ark_crypto_primitives::sponge::{
|
|||
constraints::{AbsorbGadget, CryptographicSpongeVar},
|
|||
poseidon::{constraints::PoseidonSpongeVar, PoseidonConfig},
|
|||
Absorb,
|
|||
};
|
|||
use ark_ec::CurveGroup;
|
|||
use ark_r1cs_std::{
|
|||
alloc::AllocVar, eq::EqGadget, fields::fp::FpVar, prelude::CurveVar, ToConstraintFieldGadget,
|
|||
};
|
|||
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystemRef, SynthesisError};
|
|||
use ark_std::{marker::PhantomData, Zero};
|
|||
|
|||
use crate::{
|
|||
arith::{r1cs::R1CS, Arith, ArithGadget},
|
|||
commitment::pedersen::Params as PedersenParams,
|
|||
folding::{
|
|||
circuits::{
|
|||
cyclefold::{
|
|||
CycleFoldCommittedInstance, CycleFoldCommittedInstanceVar, CycleFoldWitness,
|
|||
},
|
|||
decider::{EvalGadget, KZGChallengesGadget},
|
|||
nonnative::affine::NonNativeAffineVar,
|
|||
CF1, CF2,
|
|||
},
|
|||
traits::{CommittedInstanceOps, CommittedInstanceVarOps, Dummy, WitnessOps, WitnessVarOps},
|
|||
},
|
|||
};
|
|||
|
|||
use super::DeciderEnabledNIFS;
|
|||
|
|||
/// A generic circuit tailored for the onchain (Ethereum's EVM) verification of
|
|||
/// IVC proofs, where we support IVC built upon any folding scheme.
|
|||
///
|
|||
/// Specifically, `GenericDeciderEthCircuit` implements the in-circuit version
|
|||
/// of the IVC verification algorithm, which essentially checks the following:
|
|||
/// - `R_arith(W_i, U_i)`:
|
|||
/// The running instance `U_i` and witness `W_i` satisfy `arith`,
|
|||
/// and the commitments in `U_i` open to the values in `W_i`.
|
|||
/// - `R_arith(w_i, u_i)`:
|
|||
/// The incoming instance `u_i` and witness `w_i` satisfy `arith`,
|
|||
/// and the commitments in `u_i` open to the values in `w_i`.
|
|||
/// - `R_cf_arith(cf_W_i, cf_U_i)`:
|
|||
/// The CycleFold instance `cf_U_i` and witness `cf_W_i` satisfy `cf_arith`,
|
|||
/// and the commitments in `cf_U_i` open to the values in `cf_W_i`.
|
|||
/// - `u_i` contains the correct hash of the initial and final states.
|
|||
///
|
|||
/// To reduce the number of relation checks, the prover, before invoking the
|
|||
/// circuit, further folds `U_i, u_i` into `U_{i+1}`, and `W_i, w_i` into
|
|||
/// `W_{i+1}`.
|
|||
/// Now, the circuit only needs to perform two relation checks, i.e.,
|
|||
/// `R_arith(W_{i+1}, U_{i+1})` and `R_cf_arith(cf_W_i, cf_U_i)`, plus a few
|
|||
/// constraints for enforcing the correct hash in `u_i` and the correct folding
|
|||
/// from `U_i, u_i` to `U_{i+1}`.
|
|||
///
|
|||
/// We further reduce the circuit size by avoiding the non-native commitment
|
|||
/// checks involved in `R_arith(W_{i+1}, U_{i+1})`.
|
|||
/// Now, we now only check the satisfiability of the constraint system `arith`
|
|||
/// with the witness `W_{i+1}` and instance `U_{i+1}` in the circuit, but the
|
|||
/// actual commitment checks are done with the help of KZG.
|
|||
///
|
|||
/// For more details, see [https://privacy-scaling-explorations.github.io/sonobe-docs/design/nova-decider-onchain.html].
|
|||
pub struct GenericOnchainDeciderCircuit<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1>, // Running instance
|
|||
IU: CommittedInstanceOps<C1>, // Incoming instance
|
|||
W: WitnessOps<CF1<C1>>, // Witness
|
|||
A: Arith<W, RU>, // Constraint system
|
|||
AVar: ArithGadget<W::Var, RU::Var>, // In-circuit representation of `A`
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
> {
|
|||
pub _gc2: PhantomData<GC2>,
|
|||
pub _avar: PhantomData<AVar>,
|
|||
/// Constraint system of the Augmented Function circuit
|
|||
pub arith: A,
|
|||
/// R1CS of the CycleFold circuit
|
|||
pub cf_arith: R1CS<CF1<C2>>,
|
|||
/// CycleFold PedersenParams over C2
|
|||
pub cf_pedersen_params: PedersenParams<C2>,
|
|||
pub poseidon_config: PoseidonConfig<CF1<C1>>,
|
|||
/// public params hash
|
|||
pub pp_hash: CF1<C1>,
|
|||
pub i: CF1<C1>,
|
|||
/// initial state
|
|||
pub z_0: Vec<CF1<C1>>,
|
|||
/// current i-th state
|
|||
pub z_i: Vec<CF1<C1>>,
|
|||
/// Folding scheme instances
|
|||
pub U_i: RU,
|
|||
pub W_i: W,
|
|||
pub u_i: IU,
|
|||
pub w_i: W,
|
|||
pub U_i1: RU,
|
|||
pub W_i1: W,
|
|||
|
|||
/// Helper for folding verification
|
|||
pub proof: D::Proof,
|
|||
pub randomness: D::Randomness,
|
|||
|
|||
/// CycleFold running instance
|
|||
pub cf_U_i: CycleFoldCommittedInstance<C2>,
|
|||
pub cf_W_i: CycleFoldWitness<C2>,
|
|||
|
|||
/// KZG challenges
|
|||
pub kzg_challenges: Vec<CF1<C1>>,
|
|||
pub kzg_evaluations: Vec<CF1<C1>>,
|
|||
}
|
|||
|
|||
impl<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup<ScalarField = CF2<C1>, BaseField = CF1<C1>>,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1> + for<'a> Dummy<&'a A>,
|
|||
IU: CommittedInstanceOps<C1> + for<'a> Dummy<&'a A>,
|
|||
W: WitnessOps<CF1<C1>> + for<'a> Dummy<&'a A>,
|
|||
A: Arith<W, RU>,
|
|||
AVar: ArithGadget<W::Var, RU::Var> + AllocVar<A, CF1<C1>>,
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
>
|
|||
Dummy<(
|
|||
A,
|
|||
R1CS<CF1<C2>>,
|
|||
PedersenParams<C2>,
|
|||
PoseidonConfig<CF1<C1>>,
|
|||
D::ProofDummyCfg,
|
|||
D::RandomnessDummyCfg,
|
|||
usize,
|
|||
usize,
|
|||
)> for GenericOnchainDeciderCircuit<C1, C2, GC2, RU, IU, W, A, AVar, D>
|
|||
{
|
|||
fn dummy(
|
|||
(
|
|||
arith,
|
|||
cf_arith,
|
|||
cf_pedersen_params,
|
|||
poseidon_config,
|
|||
proof_config,
|
|||
randomness_config,
|
|||
state_len,
|
|||
num_commitments,
|
|||
): (
|
|||
A,
|
|||
R1CS<CF1<C2>>,
|
|||
PedersenParams<C2>,
|
|||
PoseidonConfig<CF1<C1>>,
|
|||
D::ProofDummyCfg,
|
|||
D::RandomnessDummyCfg,
|
|||
usize,
|
|||
usize,
|
|||
),
|
|||
) -> Self {
|
|||
Self {
|
|||
_gc2: PhantomData,
|
|||
_avar: PhantomData,
|
|||
cf_pedersen_params,
|
|||
poseidon_config,
|
|||
pp_hash: Zero::zero(),
|
|||
i: Zero::zero(),
|
|||
z_0: vec![Zero::zero(); state_len],
|
|||
z_i: vec![Zero::zero(); state_len],
|
|||
U_i: RU::dummy(&arith),
|
|||
W_i: W::dummy(&arith),
|
|||
u_i: IU::dummy(&arith),
|
|||
w_i: W::dummy(&arith),
|
|||
U_i1: RU::dummy(&arith),
|
|||
W_i1: W::dummy(&arith),
|
|||
proof: D::Proof::dummy(proof_config),
|
|||
randomness: D::Randomness::dummy(randomness_config),
|
|||
cf_U_i: CycleFoldCommittedInstance::dummy(&cf_arith),
|
|||
cf_W_i: CycleFoldWitness::dummy(&cf_arith),
|
|||
kzg_challenges: vec![Zero::zero(); num_commitments],
|
|||
kzg_evaluations: vec![Zero::zero(); num_commitments],
|
|||
arith,
|
|||
cf_arith,
|
|||
}
|
|||
}
|
|||
}
|
|||
|
|||
impl<
|
|||
C1: CurveGroup,
|
|||
C2: CurveGroup<ScalarField = CF2<C1>, BaseField = CF1<C1>>,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
RU: CommittedInstanceOps<C1>,
|
|||
IU: CommittedInstanceOps<C1>,
|
|||
W: WitnessOps<CF1<C1>>,
|
|||
A: Arith<W, RU>,
|
|||
AVar: ArithGadget<W::Var, RU::Var> + AllocVar<A, CF1<C1>>,
|
|||
D: DeciderEnabledNIFS<C1, RU, IU, W, A>,
|
|||
> ConstraintSynthesizer<CF1<C1>>
|
|||
for GenericOnchainDeciderCircuit<C1, C2, GC2, RU, IU, W, A, AVar, D>
|
|||
where
|
|||
RU::Var: AbsorbGadget<CF1<C1>> + CommittedInstanceVarOps<C1, PointVar = NonNativeAffineVar<C1>>,
|
|||
CF1<C1>: Absorb,
|
|||
{
|
|||
fn generate_constraints(self, cs: ConstraintSystemRef<CF1<C1>>) -> Result<(), SynthesisError> {
|
|||
let arith = AVar::new_witness(cs.clone(), || Ok(&self.arith))?;
|
|||
|
|||
let pp_hash = FpVar::new_input(cs.clone(), || Ok(self.pp_hash))?;
|
|||
let i = FpVar::new_input(cs.clone(), || Ok(self.i))?;
|
|||
let z_0 = Vec::new_input(cs.clone(), || Ok(self.z_0))?;
|
|||
let z_i = Vec::new_input(cs.clone(), || Ok(self.z_i))?;
|
|||
|
|||
let u_i = IU::Var::new_witness(cs.clone(), || Ok(self.u_i))?;
|
|||
let U_i = RU::Var::new_witness(cs.clone(), || Ok(self.U_i))?;
|
|||
// here (U_i1, W_i1) = NIFS.P( (U_i,W_i), (u_i,w_i))
|
|||
let U_i1_commitments = Vec::<NonNativeAffineVar<C1>>::new_input(cs.clone(), || {
|
|||
Ok(self.U_i1.get_commitments())
|
|||
})?;
|
|||
let U_i1 = RU::Var::new_witness(cs.clone(), || Ok(self.U_i1))?;
|
|||
let W_i1 = W::Var::new_witness(cs.clone(), || Ok(self.W_i1))?;
|
|||
U_i1.get_commitments().enforce_equal(&U_i1_commitments)?;
|
|||
|
|||
let cf_U_i =
|
|||
CycleFoldCommittedInstanceVar::<C2, GC2>::new_witness(cs.clone(), || Ok(self.cf_U_i))?;
|
|||
|
|||
// allocate the inputs for the check 7.1 and 7.2
|
|||
let kzg_challenges = Vec::new_input(cs.clone(), || Ok(self.kzg_challenges))?;
|
|||
let kzg_evaluations = Vec::new_input(cs.clone(), || Ok(self.kzg_evaluations))?;
|
|||
|
|||
// `sponge` is for digest computation.
|
|||
let sponge = PoseidonSpongeVar::new(cs.clone(), &self.poseidon_config);
|
|||
// `transcript` is for challenge generation.
|
|||
let mut transcript = sponge.clone();
|
|||
|
|||
// NOTE: we use the same enumeration as in
|
|||
// https://privacy-scaling-explorations.github.io/sonobe-docs/design/nova-decider-onchain.html
|
|||
// in order to make it easier to reason about.
|
|||
|
|||
// 1. enforce `U_{i+1}` and `W_{i+1}` satisfy `arith`
|
|||
arith.enforce_relation(&W_i1, &U_i1)?;
|
|||
|
|||
// 2. enforce `u_i` is an incoming instance
|
|||
u_i.enforce_incoming()?;
|
|||
|
|||
// 3. u_i.x[0] == H(i, z_0, z_i, U_i), u_i.x[1] == H(cf_U_i)
|
|||
let (u_i_x, U_i_vec) = U_i.hash(&sponge, &pp_hash, &i, &z_0, &z_i)?;
|
|||
let (cf_u_i_x, _) = cf_U_i.hash(&sponge, pp_hash.clone())?;
|
|||
u_i.get_public_inputs().enforce_equal(&[u_i_x, cf_u_i_x])?;
|
|||
|
|||
#[cfg(feature = "light-test")]
|
|||
log::warn!("[WARNING]: Running with the 'light-test' feature, skipping the big part of the DeciderEthCircuit.\n Only for testing purposes.");
|
|||
|
|||
// The following two checks (and their respective allocations) are disabled for normal
|
|||
// tests since they take several millions of constraints and would take several minutes
|
|||
// (and RAM) to run the test. It is active by default, and not active only when
|
|||
// 'light-test' feature is used.
|
|||
#[cfg(not(feature = "light-test"))]
|
|||
{
|
|||
// imports here instead of at the top of the file, so we avoid having multiple
|
|||
// `#[cfg(not(test))]`
|
|||
use crate::{
|
|||
arith::r1cs::circuits::R1CSMatricesVar,
|
|||
commitment::pedersen::PedersenGadget,
|
|||
folding::circuits::{
|
|||
cyclefold::CycleFoldWitnessVar, nonnative::uint::NonNativeUintVar,
|
|||
},
|
|||
};
|
|||
use ark_r1cs_std::ToBitsGadget;
|
|||
let cf_W_i = CycleFoldWitnessVar::<C2>::new_witness(cs.clone(), || Ok(self.cf_W_i))?;
|
|||
// 4. check Pedersen commitments of cf_U_i.{cmE, cmW}
|
|||
let H = GC2::constant(self.cf_pedersen_params.h);
|
|||
let G = self
|
|||
.cf_pedersen_params
|
|||
.generators
|
|||
.iter()
|
|||
.map(|&g| GC2::constant(g.into()))
|
|||
.collect::<Vec<_>>();
|
|||
let cf_W_i_E_bits = cf_W_i
|
|||
.E
|
|||
.iter()
|
|||
.map(|E_i| E_i.to_bits_le())
|
|||
.collect::<Result<Vec<_>, _>>()?;
|
|||
let cf_W_i_W_bits = cf_W_i
|
|||
.W
|
|||
.iter()
|
|||
.map(|W_i| W_i.to_bits_le())
|
|||
.collect::<Result<Vec<_>, _>>()?;
|
|||
PedersenGadget::<C2, GC2>::commit(&H, &G, &cf_W_i_E_bits, &cf_W_i.rE.to_bits_le()?)?
|
|||
.enforce_equal(&cf_U_i.cmE)?;
|
|||
PedersenGadget::<C2, GC2>::commit(&H, &G, &cf_W_i_W_bits, &cf_W_i.rW.to_bits_le()?)?
|
|||
.enforce_equal(&cf_U_i.cmW)?;
|
|||
|
|||
let cf_r1cs = R1CSMatricesVar::<CF1<C2>, NonNativeUintVar<CF2<C2>>>::new_constant(
|
|||
ConstraintSystemRef::None,
|
|||
self.cf_arith,
|
|||
)?;
|
|||
|
|||
// 5. enforce `cf_U_i` and `cf_W_i` satisfy `cf_r1cs`
|
|||
cf_r1cs.enforce_relation(&cf_W_i, &cf_U_i)?;
|
|||
}
|
|||
|
|||
// 6.1. partially enforce `NIFS.V(U_i, u_i) = U_{i+1}`.
|
|||
D::fold_field_elements_gadget(
|
|||
&self.arith,
|
|||
&mut transcript,
|
|||
pp_hash,
|
|||
U_i,
|
|||
U_i_vec,
|
|||
u_i,
|
|||
self.proof,
|
|||
self.randomness,
|
|||
)?
|
|||
.enforce_partial_equal(&U_i1)?;
|
|||
|
|||
// 7.1. compute and check KZG challenges
|
|||
KZGChallengesGadget::get_challenges_gadget(&mut transcript, &U_i1)?
|
|||
.enforce_equal(&kzg_challenges)?;
|
|||
|
|||
// 7.2. check the claimed evaluations
|
|||
for (((v, _r), c), e) in W_i1
|
|||
.get_openings()
|
|||
.iter()
|
|||
.zip(&kzg_challenges)
|
|||
.zip(&kzg_evaluations)
|
|||
{
|
|||
// The randomness `_r` is currently not used.
|
|||
EvalGadget::evaluate_gadget(v, c)?.enforce_equal(e)?;
|
|||
}
|
|||
|
|||
Ok(())
|
|||
}
|
|||
}
|
@ -0,0 +1,252 @@ |
|||
/// This file implements the onchain (Ethereum's EVM) decider circuit. For non-ethereum use cases,
|
|||
/// other more efficient approaches can be used.
|
|||
use ark_crypto_primitives::sponge::{
|
|||
constraints::CryptographicSpongeVar,
|
|||
poseidon::{constraints::PoseidonSpongeVar, PoseidonSponge},
|
|||
Absorb, CryptographicSponge,
|
|||
};
|
|||
use ark_ec::CurveGroup;
|
|||
use ark_ff::PrimeField;
|
|||
use ark_r1cs_std::{
|
|||
alloc::{AllocVar, AllocationMode},
|
|||
eq::EqGadget,
|
|||
fields::fp::FpVar,
|
|||
groups::CurveVar,
|
|||
ToConstraintFieldGadget,
|
|||
};
|
|||
use ark_relations::r1cs::{Namespace, SynthesisError};
|
|||
use ark_std::{borrow::Borrow, marker::PhantomData};
|
|||
|
|||
use crate::{
|
|||
arith::r1cs::{circuits::R1CSMatricesVar, R1CS},
|
|||
commitment::{pedersen::Params as PedersenParams, CommitmentScheme},
|
|||
folding::{
|
|||
circuits::{
|
|||
decider::{
|
|||
on_chain::GenericOnchainDeciderCircuit, DeciderEnabledNIFS, EvalGadget,
|
|||
KZGChallengesGadget,
|
|||
},
|
|||
CF1, CF2,
|
|||
},
|
|||
traits::{WitnessOps, WitnessVarOps},
|
|||
},
|
|||
frontend::FCircuit,
|
|||
Error,
|
|||
};
|
|||
|
|||
use super::{
|
|||
circuits::FoldingGadget,
|
|||
constants::{INCOMING, RUNNING},
|
|||
folding::{Folding, ProtoGalaxyProof},
|
|||
CommittedInstance, CommittedInstanceVar, ProtoGalaxy, Witness,
|
|||
};
|
|||
|
|||
/// In-circuit representation of the Witness associated to the CommittedInstance.
|
|||
#[derive(Debug, Clone)]
|
|||
pub struct WitnessVar<F: PrimeField> {
|
|||
pub W: Vec<FpVar<F>>,
|
|||
pub rW: FpVar<F>,
|
|||
}
|
|||
|
|||
impl<F: PrimeField> AllocVar<Witness<F>, F> for WitnessVar<F> {
|
|||
fn new_variable<T: Borrow<Witness<F>>>(
|
|||
cs: impl Into<Namespace<F>>,
|
|||
f: impl FnOnce() -> Result<T, SynthesisError>,
|
|||
mode: AllocationMode,
|
|||
) -> Result<Self, SynthesisError> {
|
|||
f().and_then(|val| {
|
|||
let cs = cs.into();
|
|||
|
|||
let W = Vec::new_variable(cs.clone(), || Ok(val.borrow().w.to_vec()), mode)?;
|
|||
let rW = FpVar::new_variable(cs.clone(), || Ok(val.borrow().r_w), mode)?;
|
|||
|
|||
Ok(Self { W, rW })
|
|||
})
|
|||
}
|
|||
}
|
|||
|
|||
impl<F: PrimeField> WitnessVarOps<F> for WitnessVar<F> {
|
|||
fn get_openings(&self) -> Vec<(&[FpVar<F>], FpVar<F>)> {
|
|||
vec![(&self.W, self.rW.clone())]
|
|||
}
|
|||
}
|
|||
|
|||
pub type DeciderEthCircuit<C1, C2, GC2> = GenericOnchainDeciderCircuit<
|
|||
C1,
|
|||
C2,
|
|||
GC2,
|
|||
CommittedInstance<C1, RUNNING>,
|
|||
CommittedInstance<C1, INCOMING>,
|
|||
Witness<CF1<C1>>,
|
|||
R1CS<CF1<C1>>,
|
|||
R1CSMatricesVar<CF1<C1>, FpVar<CF1<C1>>>,
|
|||
DeciderProtoGalaxyGadget,
|
|||
>;
|
|||
|
|||
/// returns an instance of the DeciderEthCircuit from the given ProtoGalaxy struct
|
|||
impl<
|
|||
C1: CurveGroup,
|
|||
GC1: CurveVar<C1, CF2<C1>> + ToConstraintFieldGadget<CF2<C1>>,
|
|||
C2: CurveGroup,
|
|||
GC2: CurveVar<C2, CF2<C2>> + ToConstraintFieldGadget<CF2<C2>>,
|
|||
FC: FCircuit<C1::ScalarField>,
|
|||
CS1: CommitmentScheme<C1, false>,
|
|||
// enforce that the CS2 is Pedersen commitment scheme, since we're at Ethereum's EVM decider
|
|||
CS2: CommitmentScheme<C2, false, ProverParams = PedersenParams<C2>>,
|
|||
> TryFrom<ProtoGalaxy<C1, GC1, C2, GC2, FC, CS1, CS2>> for DeciderEthCircuit<C1, C2, GC2>
|
|||
where
|
|||
CF1<C1>: Absorb,
|
|||
{
|
|||
type Error = Error;
|
|||
|
|||
fn try_from(protogalaxy: ProtoGalaxy<C1, GC1, C2, GC2, FC, CS1, CS2>) -> Result<Self, Error> {
|
|||
let mut transcript = PoseidonSponge::<C1::ScalarField>::new(&protogalaxy.poseidon_config);
|
|||
|
|||
let (U_i1, W_i1, proof, aux) = Folding::prove(
|
|||
&mut transcript,
|
|||
&protogalaxy.r1cs,
|
|||
&protogalaxy.U_i,
|
|||
&protogalaxy.W_i,
|
|||
&[protogalaxy.u_i.clone()],
|
|||
&[protogalaxy.w_i.clone()],
|
|||
)?;
|
|||
|
|||
// compute the KZG challenges used as inputs in the circuit
|
|||
let kzg_challenges = KZGChallengesGadget::get_challenges_native(&mut transcript, &U_i1);
|
|||
|
|||
// get KZG evals
|
|||
let kzg_evaluations = W_i1
|
|||
.get_openings()
|
|||
.iter()
|
|||
.zip(&kzg_challenges)
|
|||
.map(|((v, _), &c)| EvalGadget::evaluate_native(v, c))
|
|||
.collect::<Result<Vec<_>, _>>()?;
|
|||
|
|||
Ok(Self {
|
|||
_gc2: PhantomData,
|
|||
_avar: PhantomData,
|
|||
arith: protogalaxy.r1cs,
|
|||
cf_arith: protogalaxy.cf_r1cs,
|
|||
cf_pedersen_params: protogalaxy.cf_cs_params,
|
|||
poseidon_config: protogalaxy.poseidon_config,
|
|||
pp_hash: protogalaxy.pp_hash,
|
|||
i: protogalaxy.i,
|
|||
z_0: protogalaxy.z_0,
|
|||
z_i: protogalaxy.z_i,
|
|||
U_i: protogalaxy.U_i,
|
|||
W_i: protogalaxy.W_i,
|
|||
u_i: protogalaxy.u_i,
|
|||
w_i: protogalaxy.w_i,
|
|||
U_i1,
|
|||
W_i1,
|
|||
proof,
|
|||
randomness: aux.L_X_evals,
|
|||
cf_U_i: protogalaxy.cf_U_i,
|
|||
cf_W_i: protogalaxy.cf_W_i,
|
|||
kzg_challenges,
|
|||
kzg_evaluations,
|
|||
})
|
|||
}
|
|||
}
|
|||
|
|||
pub struct DeciderProtoGalaxyGadget;
|
|||
|
|||
impl<C: CurveGroup>
|
|||
DeciderEnabledNIFS<
|
|||
C,
|
|||
CommittedInstance<C, RUNNING>,
|
|||
CommittedInstance<C, INCOMING>,
|
|||
Witness<CF1<C>>,
|
|||
R1CS<CF1<C>>,
|
|||
> for DeciderProtoGalaxyGadget
|
|||
{
|
|||
type Proof = ProtoGalaxyProof<CF1<C>>;
|
|||
type ProofDummyCfg = (usize, usize, usize);
|
|||
type Randomness = Vec<CF1<C>>;
|
|||
type RandomnessDummyCfg = usize;
|
|||
|
|||
fn fold_field_elements_gadget(
|
|||
_arith: &R1CS<CF1<C>>,
|
|||
transcript: &mut PoseidonSpongeVar<CF1<C>>,
|
|||
_pp_hash: FpVar<CF1<C>>,
|
|||
U: CommittedInstanceVar<C, RUNNING>,
|
|||
_U_vec: Vec<FpVar<CF1<C>>>,
|
|||
u: CommittedInstanceVar<C, INCOMING>,
|
|||
proof: Self::Proof,
|
|||
randomness: Self::Randomness,
|
|||
) -> Result<CommittedInstanceVar<C, RUNNING>, SynthesisError> {
|
|||
let cs = transcript.cs();
|
|||
let F_coeffs = Vec::new_witness(cs.clone(), || Ok(&proof.F_coeffs[..]))?;
|
|||
let K_coeffs = Vec::new_witness(cs.clone(), || Ok(&proof.K_coeffs[..]))?;
|
|||
let randomness = Vec::new_input(cs.clone(), || Ok(randomness))?;
|
|||
|
|||
let (U_next, L_X_evals) =
|
|||
FoldingGadget::fold_committed_instance(transcript, &U, &[u], F_coeffs, K_coeffs)?;
|
|||
L_X_evals.enforce_equal(&randomness)?;
|
|||
|
|||
Ok(U_next)
|
|||
}
|
|||
|
|||
fn fold_group_elements_native(
|
|||
U_commitments: &[C],
|
|||
u_commitments: &[C],
|
|||
_: Option<Self::Proof>,
|
|||
L_X_evals: Self::Randomness,
|
|||
) -> Result<Vec<C>, Error> {
|
|||
let U_phi = U_commitments[0];
|
|||
let u_phi = u_commitments[0];
|
|||
Ok(vec![U_phi * L_X_evals[0] + u_phi * L_X_evals[1]])
|
|||
}
|
|||
}
|
|||
|
|||
#[cfg(test)]
|
|||
pub mod tests {
|
|||
use ark_bn254::{constraints::GVar, Fr, G1Projective as Projective};
|
|||
use ark_grumpkin::{constraints::GVar as GVar2, Projective as Projective2};
|
|||
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystem};
|
|||
|
|||
use super::*;
|
|||
use crate::commitment::pedersen::Pedersen;
|
|||
use crate::folding::protogalaxy::ProtoGalaxy;
|
|||
use crate::frontend::{utils::CubicFCircuit, FCircuit};
|
|||
use crate::transcript::poseidon::poseidon_canonical_config;
|
|||
use crate::FoldingScheme;
|
|||
|
|||
#[test]
|
|||
fn test_decider_circuit() {
|
|||
let mut rng = ark_std::test_rng();
|
|||
let poseidon_config = poseidon_canonical_config::<Fr>();
|
|||
|
|||
let F_circuit = CubicFCircuit::<Fr>::new(()).unwrap();
|
|||
let z_0 = vec![Fr::from(3_u32)];
|
|||
|
|||
type PG = ProtoGalaxy<
|
|||
Projective,
|
|||
GVar,
|
|||
Projective2,
|
|||
GVar2,
|
|||
CubicFCircuit<Fr>,
|
|||
Pedersen<Projective>,
|
|||
Pedersen<Projective2>,
|
|||
>;
|
|||
let pg_params = PG::preprocess(&mut rng, &(poseidon_config, F_circuit)).unwrap();
|
|||
|
|||
// generate a Nova instance and do a step of it
|
|||
let mut protogalaxy = PG::init(&pg_params, F_circuit, z_0.clone()).unwrap();
|
|||
protogalaxy.prove_step(&mut rng, vec![], None).unwrap();
|
|||
|
|||
let ivc_proof = protogalaxy.ivc_proof();
|
|||
PG::verify(pg_params.1, ivc_proof).unwrap();
|
|||
|
|||
// load the DeciderEthCircuit from the generated Nova instance
|
|||
let decider_circuit =
|
|||
DeciderEthCircuit::<Projective, Projective2, GVar2>::try_from(protogalaxy).unwrap();
|
|||
|
|||
let cs = ConstraintSystem::<Fr>::new_ref();
|
|||
|
|||
// generate the constraints and check that are satisfied by the inputs
|
|||
decider_circuit.generate_constraints(cs.clone()).unwrap();
|
|||
assert!(cs.is_satisfied().unwrap());
|
|||
dbg!(cs.num_constraints());
|
|||
}
|
|||
}
|