Browse Source

Add NIFSGadgetTrait, implement Mova's NIFSGadget, adapt Nova NIFSGadget into NIFSGadgetTrait (#173)

* add new NIFSGadgetTrait

* implement Mova's NIFSGadget

* refactor Nova NIFSGadget to fit into the new NIFSGadgetTrait

* abstract NIFSGadget related tests for all implementors of
  NIFSGadgetTrait to avoid duplicated code in the tests between the
  different Nova variants gadget tests

* frontends/noir update mimc usage since it has been migrated from
  noir's std into it's own repo
main
arnaucube 1 month ago
committed by GitHub
parent
commit
e1183877e7
No known key found for this signature in database GPG Key ID: B5690EEEBB952194
20 changed files with 782 additions and 405 deletions
  1. +1
    -1
      examples/full_flow.rs
  2. +2
    -7
      examples/noir_full_flow.rs
  3. +2
    -2
      folding-schemes/src/arith/r1cs/circuits.rs
  4. +1
    -1
      folding-schemes/src/folding/circuits/decider/mod.rs
  5. +1
    -1
      folding-schemes/src/folding/circuits/decider/off_chain.rs
  6. +32
    -343
      folding-schemes/src/folding/nova/circuits.rs
  7. +0
    -1
      folding-schemes/src/folding/nova/decider.rs
  8. +1
    -0
      folding-schemes/src/folding/nova/decider_circuits.rs
  9. +1
    -3
      folding-schemes/src/folding/nova/decider_eth.rs
  10. +6
    -16
      folding-schemes/src/folding/nova/decider_eth_circuit.rs
  11. +3
    -3
      folding-schemes/src/folding/nova/mod.rs
  12. +178
    -10
      folding-schemes/src/folding/nova/nifs/mod.rs
  13. +62
    -4
      folding-schemes/src/folding/nova/nifs/nova.rs
  14. +237
    -0
      folding-schemes/src/folding/nova/nifs/nova_circuits.rs
  15. +28
    -5
      folding-schemes/src/folding/nova/nifs/ova.rs
  16. +224
    -0
      folding-schemes/src/folding/nova/nifs/ova_circuits.rs
  17. +1
    -1
      folding-schemes/src/folding/nova/traits.rs
  18. +1
    -1
      frontends/src/noir/test_folder/test_mimc/Nargo.toml
  19. +1
    -3
      frontends/src/noir/test_folder/test_mimc/src/main.nr
  20. +0
    -3
      solidity-verifiers/templates/nova_cyclefold_decider.askama.sol

+ 1
- 1
examples/full_flow.rs

@ -78,7 +78,7 @@ impl FCircuit for CubicFCircuit {
} }
fn main() { fn main() {
let n_steps = 10;
let n_steps = 5;
// set the initial state // set the initial state
let z_0 = vec![Fr::from(3_u32)]; let z_0 = vec![Fr::from(3_u32)];

+ 2
- 7
examples/noir_full_flow.rs

@ -28,7 +28,7 @@ use folding_schemes::{
Decider, FoldingScheme, Decider, FoldingScheme,
}; };
use frontends::noir::{load_noir_circuit, NoirFCircuit}; use frontends::noir::{load_noir_circuit, NoirFCircuit};
use std::{env, time::Instant};
use std::time::Instant;
use solidity_verifiers::{ use solidity_verifiers::{
evm::{compile_solidity, Evm}, evm::{compile_solidity, Evm},
@ -42,12 +42,7 @@ fn main() {
let z_0 = vec![Fr::from(1)]; let z_0 = vec![Fr::from(1)];
// initialize the noir fcircuit // initialize the noir fcircuit
let cur_path = env::current_dir().unwrap();
let circuit_path = format!(
"{}/frontends/src/noir/test_folder/test_mimc/target/test_mimc.json",
cur_path.to_str().unwrap()
);
let circuit_path = format!("./frontends/src/noir/test_folder/test_mimc/target/test_mimc.json",);
let circuit = load_noir_circuit(circuit_path).unwrap(); let circuit = load_noir_circuit(circuit_path).unwrap();
let f_circuit = NoirFCircuit { let f_circuit = NoirFCircuit {

+ 2
- 2
folding-schemes/src/arith/r1cs/circuits.rs

@ -121,8 +121,8 @@ pub mod tests {
nonnative::uint::NonNativeUintVar, nonnative::uint::NonNativeUintVar,
}, },
nova::{ nova::{
circuits::CommittedInstanceVar, decider_eth_circuit::WitnessVar, CommittedInstance,
Witness,
decider_eth_circuit::WitnessVar, nifs::nova_circuits::CommittedInstanceVar,
CommittedInstance, Witness,
}, },
}; };
use crate::frontend::{ use crate::frontend::{

+ 1
- 1
folding-schemes/src/folding/circuits/decider/mod.rs

@ -150,7 +150,7 @@ pub mod tests {
use ark_std::UniformRand; use ark_std::UniformRand;
use super::*; use super::*;
use crate::folding::nova::{circuits::CommittedInstanceVar, CommittedInstance};
use crate::folding::nova::{nifs::nova_circuits::CommittedInstanceVar, CommittedInstance};
use crate::transcript::poseidon::poseidon_canonical_config; use crate::transcript::poseidon::poseidon_canonical_config;
// checks that the gadget and native implementations of the challenge computation match // checks that the gadget and native implementations of the challenge computation match

+ 1
- 1
folding-schemes/src/folding/circuits/decider/off_chain.rs

@ -28,7 +28,7 @@ use crate::{
nonnative::affine::NonNativeAffineVar, nonnative::affine::NonNativeAffineVar,
CF1, CF2, CF1, CF2,
}, },
nova::{circuits::CommittedInstanceVar, decider_eth_circuit::WitnessVar},
nova::{decider_eth_circuit::WitnessVar, nifs::nova_circuits::CommittedInstanceVar},
traits::{CommittedInstanceOps, CommittedInstanceVarOps, Dummy, WitnessOps, WitnessVarOps}, traits::{CommittedInstanceOps, CommittedInstanceVarOps, Dummy, WitnessOps, WitnessVarOps},
}, },
}; };

+ 32
- 343
folding-schemes/src/folding/nova/circuits.rs

@ -1,25 +1,30 @@
/// contains [Nova](https://eprint.iacr.org/2021/370.pdf) related circuits /// contains [Nova](https://eprint.iacr.org/2021/370.pdf) related circuits
use ark_crypto_primitives::sponge::{ use ark_crypto_primitives::sponge::{
constraints::{AbsorbGadget, CryptographicSpongeVar},
poseidon::{constraints::PoseidonSpongeVar, PoseidonConfig},
Absorb, CryptographicSponge,
constraints::CryptographicSpongeVar,
poseidon::{constraints::PoseidonSpongeVar, PoseidonConfig, PoseidonSponge},
Absorb,
}; };
use ark_ec::{CurveGroup, Group}; use ark_ec::{CurveGroup, Group};
use ark_ff::PrimeField; use ark_ff::PrimeField;
use ark_r1cs_std::{ use ark_r1cs_std::{
alloc::{AllocVar, AllocationMode},
alloc::AllocVar,
boolean::Boolean, boolean::Boolean,
eq::EqGadget, eq::EqGadget,
fields::{fp::FpVar, FieldVar}, fields::{fp::FpVar, FieldVar},
prelude::CurveVar, prelude::CurveVar,
uint8::UInt8,
R1CSVar, ToConstraintFieldGadget, R1CSVar, ToConstraintFieldGadget,
}; };
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystemRef, Namespace, SynthesisError};
use ark_relations::r1cs::{ConstraintSynthesizer, ConstraintSystemRef, SynthesisError};
use ark_std::{fmt::Debug, One, Zero}; use ark_std::{fmt::Debug, One, Zero};
use core::{borrow::Borrow, marker::PhantomData};
use core::marker::PhantomData;
use super::{CommittedInstance, NovaCycleFoldConfig};
use super::{
nifs::{
nova_circuits::{CommittedInstanceVar, NIFSGadget},
NIFSGadgetTrait,
},
CommittedInstance, NovaCycleFoldConfig,
};
use crate::folding::circuits::{ use crate::folding::circuits::{
cyclefold::{ cyclefold::{
CycleFoldChallengeGadget, CycleFoldCommittedInstance, CycleFoldCommittedInstanceVar, CycleFoldChallengeGadget, CycleFoldCommittedInstance, CycleFoldCommittedInstanceVar,
@ -28,180 +33,9 @@ use crate::folding::circuits::{
nonnative::{affine::NonNativeAffineVar, uint::NonNativeUintVar}, nonnative::{affine::NonNativeAffineVar, uint::NonNativeUintVar},
CF1, CF2, CF1, CF2,
}; };
use crate::folding::traits::{CommittedInstanceVarOps, Dummy};
use crate::frontend::FCircuit; use crate::frontend::FCircuit;
use crate::transcript::{AbsorbNonNativeGadget, Transcript, TranscriptVar};
use crate::{
constants::NOVA_N_BITS_RO,
folding::traits::{CommittedInstanceVarOps, Dummy},
};
/// CommittedInstanceVar contains the u, x, cmE and cmW values which are folded on the main Nova
/// constraints field (E1::Fr, where E1 is the main curve). The peculiarity is that cmE and cmW are
/// represented non-natively over the constraint field.
#[derive(Debug, Clone)]
pub struct CommittedInstanceVar<C: CurveGroup> {
pub u: FpVar<C::ScalarField>,
pub x: Vec<FpVar<C::ScalarField>>,
pub cmE: NonNativeAffineVar<C>,
pub cmW: NonNativeAffineVar<C>,
}
impl<C> AllocVar<CommittedInstance<C>, CF1<C>> for CommittedInstanceVar<C>
where
C: CurveGroup,
{
fn new_variable<T: Borrow<CommittedInstance<C>>>(
cs: impl Into<Namespace<CF1<C>>>,
f: impl FnOnce() -> Result<T, SynthesisError>,
mode: AllocationMode,
) -> Result<Self, SynthesisError> {
f().and_then(|val| {
let cs = cs.into();
let u = FpVar::<C::ScalarField>::new_variable(cs.clone(), || Ok(val.borrow().u), mode)?;
let x: Vec<FpVar<C::ScalarField>> =
Vec::new_variable(cs.clone(), || Ok(val.borrow().x.clone()), mode)?;
let cmE =
NonNativeAffineVar::<C>::new_variable(cs.clone(), || Ok(val.borrow().cmE), mode)?;
let cmW =
NonNativeAffineVar::<C>::new_variable(cs.clone(), || Ok(val.borrow().cmW), mode)?;
Ok(Self { u, x, cmE, cmW })
})
}
}
impl<C: CurveGroup> AbsorbGadget<C::ScalarField> for CommittedInstanceVar<C> {
fn to_sponge_bytes(&self) -> Result<Vec<UInt8<C::ScalarField>>, SynthesisError> {
FpVar::batch_to_sponge_bytes(&self.to_sponge_field_elements()?)
}
fn to_sponge_field_elements(&self) -> Result<Vec<FpVar<C::ScalarField>>, SynthesisError> {
Ok([
vec![self.u.clone()],
self.x.clone(),
self.cmE.to_constraint_field()?,
self.cmW.to_constraint_field()?,
]
.concat())
}
}
impl<C: CurveGroup> CommittedInstanceVarOps<C> for CommittedInstanceVar<C> {
type PointVar = NonNativeAffineVar<C>;
fn get_commitments(&self) -> Vec<Self::PointVar> {
vec![self.cmW.clone(), self.cmE.clone()]
}
fn get_public_inputs(&self) -> &[FpVar<CF1<C>>] {
&self.x
}
fn enforce_incoming(&self) -> Result<(), SynthesisError> {
let zero = NonNativeUintVar::new_constant(ConstraintSystemRef::None, CF2::<C>::zero())?;
self.cmE.x.enforce_equal_unaligned(&zero)?;
self.cmE.y.enforce_equal_unaligned(&zero)?;
self.u.enforce_equal(&FpVar::one())
}
fn enforce_partial_equal(&self, other: &Self) -> Result<(), SynthesisError> {
self.u.enforce_equal(&other.u)?;
self.x.enforce_equal(&other.x)
}
}
/// Implements the circuit that does the checks of the Non-Interactive Folding Scheme Verifier
/// described in section 4 of [Nova](https://eprint.iacr.org/2021/370.pdf), where the cmE & cmW checks are
/// delegated to the NIFSCycleFoldGadget.
pub struct NIFSGadget<C: CurveGroup> {
_c: PhantomData<C>,
}
impl<C: CurveGroup> NIFSGadget<C> {
pub fn fold_committed_instance(
r: FpVar<CF1<C>>,
ci1: CommittedInstanceVar<C>, // U_i
ci2: CommittedInstanceVar<C>, // u_i
) -> Result<CommittedInstanceVar<C>, SynthesisError> {
Ok(CommittedInstanceVar {
cmE: NonNativeAffineVar::new_constant(ConstraintSystemRef::None, C::zero())?,
cmW: NonNativeAffineVar::new_constant(ConstraintSystemRef::None, C::zero())?,
// ci3.u = ci1.u + r * ci2.u
u: ci1.u + &r * ci2.u,
// ci3.x = ci1.x + r * ci2.x
x: ci1
.x
.iter()
.zip(ci2.x)
.map(|(a, b)| a + &r * &b)
.collect::<Vec<FpVar<CF1<C>>>>(),
})
}
/// Implements the constraints for NIFS.V for u and x, since cm(E) and cm(W) are delegated to
/// the CycleFold circuit.
pub fn verify(
r: FpVar<CF1<C>>,
ci1: CommittedInstanceVar<C>, // U_i
ci2: CommittedInstanceVar<C>, // u_i
ci3: CommittedInstanceVar<C>, // U_{i+1}
) -> Result<(), SynthesisError> {
let ci = Self::fold_committed_instance(r, ci1, ci2)?;
ci.u.enforce_equal(&ci3.u)?;
ci.x.enforce_equal(&ci3.x)?;
Ok(())
}
}
/// ChallengeGadget computes the RO challenge used for the Nova instances NIFS, it contains a
/// rust-native and a in-circuit compatible versions.
pub struct ChallengeGadget<C: CurveGroup, CI: Absorb> {
_c: PhantomData<C>,
_ci: PhantomData<CI>,
}
impl<C: CurveGroup, CI: Absorb> ChallengeGadget<C, CI>
where
<C as Group>::ScalarField: Absorb,
{
pub fn get_challenge_native<T: Transcript<C::ScalarField>>(
transcript: &mut T,
pp_hash: C::ScalarField, // public params hash
U_i: &CI,
u_i: &CI,
cmT: Option<&C>,
) -> Vec<bool> {
transcript.absorb(&pp_hash);
transcript.absorb(&U_i);
transcript.absorb(&u_i);
// in the Nova case we absorb the cmT, in Ova case we don't since it is not used.
if let Some(cmT_value) = cmT {
transcript.absorb_nonnative(cmT_value);
}
transcript.squeeze_bits(NOVA_N_BITS_RO)
}
// compatible with the native get_challenge_native
pub fn get_challenge_gadget<S: CryptographicSponge, T: TranscriptVar<CF1<C>, S>>(
transcript: &mut T,
pp_hash: FpVar<CF1<C>>, // public params hash
U_i_vec: Vec<FpVar<CF1<C>>>, // apready processed input, so we don't have to recompute these values
u_i: CommittedInstanceVar<C>,
cmT: Option<NonNativeAffineVar<C>>,
) -> Result<Vec<Boolean<C::ScalarField>>, SynthesisError> {
transcript.absorb(&pp_hash)?;
transcript.absorb(&U_i_vec)?;
transcript.absorb(&u_i)?;
// in the Nova case we absorb the cmT, in Ova case we don't since it is not used.
if let Some(cmT_value) = cmT {
transcript.absorb_nonnative(&cmT_value)?;
}
transcript.squeeze_bits(NOVA_N_BITS_RO)
}
}
use crate::transcript::AbsorbNonNativeGadget;
/// `AugmentedFCircuit` enhances the original step function `F`, so that it can /// `AugmentedFCircuit` enhances the original step function `F`, so that it can
/// be used in recursive arguments such as IVC. /// be used in recursive arguments such as IVC.
@ -364,32 +198,33 @@ where
x: vec![u_i_x, cf_u_i_x], x: vec![u_i_x, cf_u_i_x],
}; };
// P.3. nifs.verify, obtains U_{i+1} by folding u_i & U_i .
// compute r = H(u_i, U_i, cmT)
let r_bits = ChallengeGadget::<C1, CommittedInstance<C1>>::get_challenge_gadget(
// P.3. nifs.verify, obtains U_{i+1} by folding u_i & U_i.
// Notice that NIFSGadget::verify does not fold cmE & cmW.
// We set `U_i1.cmE` and `U_i1.cmW` to unconstrained witnesses `U_i1_cmE` and `U_i1_cmW`
// respectively.
// The correctness of them will be checked on the other curve.
let (mut U_i1, r_bits) = NIFSGadget::<
C1,
PoseidonSponge<C1::ScalarField>,
PoseidonSpongeVar<C1::ScalarField>,
>::verify(
&mut transcript, &mut transcript,
pp_hash.clone(), pp_hash.clone(),
U_i.clone(),
U_i_vec, U_i_vec,
u_i.clone(), u_i.clone(),
Some(cmT.clone()), Some(cmT.clone()),
)?; )?;
let r = Boolean::le_bits_to_fp_var(&r_bits)?;
// Also convert r_bits to a `NonNativeFieldVar`
U_i1.cmE = U_i1_cmE;
U_i1.cmW = U_i1_cmW;
// convert r_bits to a `NonNativeFieldVar`
let r_nonnat = { let r_nonnat = {
let mut bits = r_bits; let mut bits = r_bits;
bits.resize(C1::BaseField::MODULUS_BIT_SIZE as usize, Boolean::FALSE); bits.resize(C1::BaseField::MODULUS_BIT_SIZE as usize, Boolean::FALSE);
NonNativeUintVar::from(&bits) NonNativeUintVar::from(&bits)
}; };
// Notice that NIFSGadget::fold_committed_instance does not fold cmE & cmW.
// We set `U_i1.cmE` and `U_i1.cmW` to unconstrained witnesses `U_i1_cmE` and `U_i1_cmW`
// respectively.
// The correctness of them will be checked on the other curve.
let mut U_i1 = NIFSGadget::<C1>::fold_committed_instance(r, U_i.clone(), u_i.clone())?;
U_i1.cmE = U_i1_cmE;
U_i1.cmW = U_i1_cmW;
// P.4.a compute and check the first output of F' // P.4.a compute and check the first output of F'
// get z_{i+1} from the F circuit // get z_{i+1} from the F circuit
@ -507,160 +342,14 @@ where
pub mod tests { pub mod tests {
use super::*; use super::*;
use ark_bn254::{Fr, G1Projective as Projective}; use ark_bn254::{Fr, G1Projective as Projective};
use ark_crypto_primitives::sponge::poseidon::PoseidonSponge;
use ark_crypto_primitives::sponge::{poseidon::PoseidonSponge, CryptographicSponge};
use ark_ff::BigInteger; use ark_ff::BigInteger;
use ark_relations::r1cs::ConstraintSystem; use ark_relations::r1cs::ConstraintSystem;
use ark_std::UniformRand; use ark_std::UniformRand;
use crate::commitment::pedersen::Pedersen;
use crate::folding::nova::nifs::{nova::NIFS, NIFSTrait};
use crate::folding::traits::CommittedInstanceOps;
use crate::folding::nova::nifs::nova::ChallengeGadget;
use crate::transcript::poseidon::poseidon_canonical_config; use crate::transcript::poseidon::poseidon_canonical_config;
#[test]
fn test_committed_instance_var() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
};
let cs = ConstraintSystem::<Fr>::new_ref();
let ciVar =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci.clone())).unwrap();
assert_eq!(ciVar.u.value().unwrap(), ci.u);
assert_eq!(ciVar.x.value().unwrap(), ci.x);
// the values cmE and cmW are checked in the CycleFold's circuit
// CommittedInstanceInCycleFoldVar in
// cyclefold::tests::test_committed_instance_cyclefold_var
}
#[test]
fn test_nifs_gadget() {
let mut rng = ark_std::test_rng();
// prepare the committed instances to test in-circuit
let ci: Vec<CommittedInstance<Projective>> = (0..2)
.into_iter()
.map(|_| CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
})
.collect();
let (ci1, ci2) = (ci[0].clone(), ci[1].clone());
let pp_hash = Fr::rand(&mut rng);
let cmT = Projective::rand(&mut rng);
let poseidon_config = poseidon_canonical_config::<Fr>();
let mut transcript = PoseidonSponge::<Fr>::new(&poseidon_config);
let (ci3, r_bits) = NIFS::<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>::verify(
&mut transcript,
pp_hash,
&ci1,
&ci2,
&cmT,
)
.unwrap();
let r_Fr = Fr::from_bigint(BigInteger::from_bits_le(&r_bits)).unwrap();
let cs = ConstraintSystem::<Fr>::new_ref();
let rVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(r_Fr)).unwrap();
let ci1Var =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci1.clone()))
.unwrap();
let ci2Var =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci2.clone()))
.unwrap();
let ci3Var =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci3.clone()))
.unwrap();
NIFSGadget::<Projective>::verify(
rVar.clone(),
ci1Var.clone(),
ci2Var.clone(),
ci3Var.clone(),
)
.unwrap();
assert!(cs.is_satisfied().unwrap());
}
/// test that checks the native CommittedInstance.to_sponge_{bytes,field_elements}
/// vs the R1CS constraints version
#[test]
pub fn test_committed_instance_to_sponge_preimage() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
};
let bytes = ci.to_sponge_bytes_as_vec();
let field_elements = ci.to_sponge_field_elements_as_vec();
let cs = ConstraintSystem::<Fr>::new_ref();
let ciVar =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci.clone())).unwrap();
let bytes_var = ciVar.to_sponge_bytes().unwrap();
let field_elements_var = ciVar.to_sponge_field_elements().unwrap();
assert!(cs.is_satisfied().unwrap());
// check that the natively computed and in-circuit computed hashes match
assert_eq!(bytes_var.value().unwrap(), bytes);
assert_eq!(field_elements_var.value().unwrap(), field_elements);
}
#[test]
fn test_committed_instance_hash() {
let mut rng = ark_std::test_rng();
let poseidon_config = poseidon_canonical_config::<Fr>();
let sponge = PoseidonSponge::<Fr>::new(&poseidon_config);
let pp_hash = Fr::from(42u32); // only for test
let i = Fr::from(3_u32);
let z_0 = vec![Fr::from(3_u32)];
let z_i = vec![Fr::from(3_u32)];
let ci = 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 CommittedInstance hash natively
let h = ci.hash(&sponge, pp_hash, i, &z_0, &z_i);
let cs = ConstraintSystem::<Fr>::new_ref();
let pp_hashVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(pp_hash)).unwrap();
let iVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(i)).unwrap();
let z_0Var = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(z_0.clone())).unwrap();
let z_iVar = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(z_i.clone())).unwrap();
let ciVar =
CommittedInstanceVar::<Projective>::new_witness(cs.clone(), || Ok(ci.clone())).unwrap();
let sponge = PoseidonSpongeVar::<Fr>::new(cs.clone(), &poseidon_config);
// compute the CommittedInstance hash in-circuit
let (hVar, _) = ciVar
.hash(&sponge, &pp_hashVar, &iVar, &z_0Var, &z_iVar)
.unwrap();
assert!(cs.is_satisfied().unwrap());
// check that the natively computed and in-circuit computed hashes match
assert_eq!(hVar.value().unwrap(), h);
}
// checks that the gadget and native implementations of the challenge computation match // checks that the gadget and native implementations of the challenge computation match
#[test] #[test]
fn test_challenge_gadget() { fn test_challenge_gadget() {

+ 0
- 1
folding-schemes/src/folding/nova/decider.rs

@ -288,7 +288,6 @@ where
&proof.cs1_challenges, &proof.cs1_challenges,
&proof.cs1_proofs.iter().map(|p| p.eval).collect::<Vec<_>>(), &proof.cs1_proofs.iter().map(|p| p.eval).collect::<Vec<_>>(),
&proof.cmT.inputize(), &proof.cmT.inputize(),
&[proof.r],
] ]
.concat(); .concat();

+ 1
- 0
folding-schemes/src/folding/nova/decider_circuits.rs

@ -55,6 +55,7 @@ impl<
> TryFrom<Nova<C1, GC1, C2, GC2, FC, CS1, CS2, H>> for DeciderCircuit1<C1, C2, GC2> > TryFrom<Nova<C1, GC1, C2, GC2, FC, CS1, CS2, H>> for DeciderCircuit1<C1, C2, GC2>
where where
CF1<C1>: Absorb, CF1<C1>: Absorb,
<C1 as CurveGroup>::BaseField: PrimeField,
{ {
type Error = Error; type Error = Error;

+ 1
- 3
folding-schemes/src/folding/nova/decider_eth.rs

@ -22,8 +22,7 @@ use crate::commitment::{
pedersen::Params as PedersenParams, pedersen::Params as PedersenParams,
CommitmentScheme, CommitmentScheme,
}; };
use crate::folding::circuits::decider::DeciderEnabledNIFS;
use crate::folding::circuits::CF2;
use crate::folding::circuits::{decider::DeciderEnabledNIFS, CF2};
use crate::folding::traits::{Inputize, WitnessOps}; use crate::folding::traits::{Inputize, WitnessOps};
use crate::frontend::FCircuit; use crate::frontend::FCircuit;
use crate::Error; use crate::Error;
@ -228,7 +227,6 @@ where
&proof.kzg_challenges, &proof.kzg_challenges,
&proof.kzg_proofs.iter().map(|p| p.eval).collect::<Vec<_>>(), &proof.kzg_proofs.iter().map(|p| p.eval).collect::<Vec<_>>(),
&proof.cmT.inputize(), &proof.cmT.inputize(),
&[proof.r],
] ]
.concat(); .concat();

+ 6
- 16
folding-schemes/src/folding/nova/decider_eth_circuit.rs

@ -11,8 +11,6 @@ use ark_ec::CurveGroup;
use ark_ff::{BigInteger, PrimeField}; use ark_ff::{BigInteger, PrimeField};
use ark_r1cs_std::{ use ark_r1cs_std::{
alloc::{AllocVar, AllocationMode}, alloc::{AllocVar, AllocationMode},
boolean::Boolean,
eq::EqGadget,
fields::fp::FpVar, fields::fp::FpVar,
prelude::CurveVar, prelude::CurveVar,
ToConstraintFieldGadget, ToConstraintFieldGadget,
@ -21,8 +19,8 @@ use ark_relations::r1cs::{Namespace, SynthesisError};
use ark_std::{borrow::Borrow, marker::PhantomData}; use ark_std::{borrow::Borrow, marker::PhantomData};
use super::{ use super::{
circuits::{ChallengeGadget, CommittedInstanceVar, NIFSGadget},
nifs::{nova::NIFS, NIFSTrait},
nifs::nova_circuits::{CommittedInstanceVar, NIFSGadget},
nifs::{nova::NIFS, NIFSGadgetTrait, NIFSTrait},
CommittedInstance, Nova, Witness, CommittedInstance, Nova, Witness,
}; };
use crate::commitment::{pedersen::Params as PedersenParams, CommitmentScheme}; use crate::commitment::{pedersen::Params as PedersenParams, CommitmentScheme};
@ -108,6 +106,7 @@ impl<
> TryFrom<Nova<C1, GC1, C2, GC2, FC, CS1, CS2, H>> for DeciderEthCircuit<C1, C2, GC2> > TryFrom<Nova<C1, GC1, C2, GC2, FC, CS1, CS2, H>> for DeciderEthCircuit<C1, C2, GC2>
where where
CF1<C1>: Absorb, CF1<C1>: Absorb,
<C1 as CurveGroup>::BaseField: PrimeField,
{ {
type Error = Error; type Error = Error;
@ -187,21 +186,12 @@ where
U_vec: Vec<FpVar<CF1<C>>>, U_vec: Vec<FpVar<CF1<C>>>,
u: CommittedInstanceVar<C>, u: CommittedInstanceVar<C>,
proof: C, proof: C,
randomness: CF1<C>,
_randomness: CF1<C>,
) -> Result<CommittedInstanceVar<C>, SynthesisError> { ) -> Result<CommittedInstanceVar<C>, SynthesisError> {
let cs = transcript.cs(); let cs = transcript.cs();
let cmT = NonNativeAffineVar::new_input(cs.clone(), || Ok(proof))?; let cmT = NonNativeAffineVar::new_input(cs.clone(), || Ok(proof))?;
let r = FpVar::new_input(cs.clone(), || Ok(randomness))?;
let r_bits = ChallengeGadget::<C, CommittedInstance<C>>::get_challenge_gadget(
transcript,
pp_hash,
U_vec,
u.clone(),
Some(cmT),
)?;
Boolean::le_bits_to_fp_var(&r_bits)?.enforce_equal(&r)?;
NIFSGadget::<C>::fold_committed_instance(r, U, u)
let (new_U, _) = NIFSGadget::verify(transcript, pp_hash, U, U_vec, u, Some(cmT))?;
Ok(new_U)
} }
fn fold_group_elements_native( fn fold_group_elements_native(

+ 3
- 3
folding-schemes/src/folding/nova/mod.rs

@ -17,7 +17,6 @@ use ark_std::fmt::Debug;
use ark_std::rand::RngCore; use ark_std::rand::RngCore;
use ark_std::{One, UniformRand, Zero}; use ark_std::{One, UniformRand, Zero};
use core::marker::PhantomData; use core::marker::PhantomData;
use decider_eth_circuit::WitnessVar;
use crate::folding::circuits::cyclefold::{ use crate::folding::circuits::cyclefold::{
fold_cyclefold_circuit, CycleFoldCircuit, CycleFoldCommittedInstance, CycleFoldConfig, fold_cyclefold_circuit, CycleFoldCircuit, CycleFoldCommittedInstance, CycleFoldConfig,
@ -38,6 +37,7 @@ use crate::{
utils::{get_cm_coordinates, pp_hash}, utils::{get_cm_coordinates, pp_hash},
}; };
use crate::{arith::Arith, commitment::CommitmentScheme}; use crate::{arith::Arith, commitment::CommitmentScheme};
use decider_eth_circuit::WitnessVar;
pub mod circuits; pub mod circuits;
pub mod traits; pub mod traits;
@ -46,8 +46,8 @@ pub mod zk;
// NIFS related: // NIFS related:
pub mod nifs; pub mod nifs;
use circuits::{AugmentedFCircuit, CommittedInstanceVar};
use nifs::{nova::NIFS, NIFSTrait};
use circuits::AugmentedFCircuit;
use nifs::{nova::NIFS, nova_circuits::CommittedInstanceVar, NIFSTrait};
// offchain decider // offchain decider
pub mod decider; pub mod decider;

+ 178
- 10
folding-schemes/src/folding/nova/nifs/mod.rs

@ -1,18 +1,28 @@
/// This module defines the NIFSTrait, which is set to implement the NIFS (Non-Interactive Folding
/// Scheme) by the various schemes (Nova, Mova, Ova).
use ark_crypto_primitives::sponge::Absorb;
/// This module defines the traits related to the NIFS (Non-Interactive Folding Scheme).
/// - NIFSTrait, which implements the NIFS interface
/// - NIFSGadget, which implements the NIFS in-circuit
/// both traits implemented by the various Nova variants schemes; ie.
/// [Nova](https://eprint.iacr.org/2021/370.pdf), [Ova](https://hackmd.io/V4838nnlRKal9ZiTHiGYzw),
/// [Mova](https://eprint.iacr.org/2024/1220.pdf).
use ark_crypto_primitives::sponge::{constraints::AbsorbGadget, Absorb, CryptographicSponge};
use ark_ec::CurveGroup; use ark_ec::CurveGroup;
use ark_r1cs_std::{alloc::AllocVar, boolean::Boolean, fields::fp::FpVar};
use ark_relations::r1cs::SynthesisError;
use ark_std::fmt::Debug; use ark_std::fmt::Debug;
use ark_std::rand::RngCore; use ark_std::rand::RngCore;
use crate::arith::r1cs::R1CS; use crate::arith::r1cs::R1CS;
use crate::commitment::CommitmentScheme; use crate::commitment::CommitmentScheme;
use crate::transcript::Transcript;
use crate::folding::circuits::CF1;
use crate::folding::traits::{CommittedInstanceOps, CommittedInstanceVarOps};
use crate::transcript::{Transcript, TranscriptVar};
use crate::Error; use crate::Error;
pub mod mova; pub mod mova;
pub mod nova; pub mod nova;
pub mod nova_circuits;
pub mod ova; pub mod ova;
pub mod ova_circuits;
pub mod pointvsline; pub mod pointvsline;
/// Defines the NIFS (Non-Interactive Folding Scheme) trait, initially defined in /// Defines the NIFS (Non-Interactive Folding Scheme) trait, initially defined in
@ -27,7 +37,7 @@ pub trait NIFSTrait<
const H: bool = false, const H: bool = false,
> >
{ {
type CommittedInstance: Debug + Clone + Absorb;
type CommittedInstance: Debug + Clone + Absorb; // + CommittedInstanceOps<C>;
type Witness: Debug + Clone; type Witness: Debug + Clone;
type ProverAux: Debug + Clone; // Prover's aux params. eg. in Nova is T type ProverAux: Debug + Clone; // Prover's aux params. eg. in Nova is T
type Proof: Debug + Clone; // proof. eg. in Nova is cmT type Proof: Debug + Clone; // proof. eg. in Nova is cmT
@ -83,17 +93,56 @@ pub trait NIFSTrait<
) -> Result<(Self::CommittedInstance, Vec<bool>), Error>; ) -> Result<(Self::CommittedInstance, Vec<bool>), Error>;
} }
/// Defines the NIFS (Non-Interactive Folding Scheme) Gadget trait, which specifies the in-circuit
/// logic of the NIFS.Verify defined in [Nova](https://eprint.iacr.org/2021/370.pdf) and it's
/// variants [Ova](https://hackmd.io/V4838nnlRKal9ZiTHiGYzw) and
/// [Mova](https://eprint.iacr.org/2024/1220.pdf).
pub trait NIFSGadgetTrait<C: CurveGroup, S: CryptographicSponge, T: TranscriptVar<CF1<C>, S>> {
type CommittedInstance: Debug + Clone + Absorb + CommittedInstanceOps<C>;
type CommittedInstanceVar: Debug
+ Clone
+ AbsorbGadget<C::ScalarField>
+ AllocVar<Self::CommittedInstance, CF1<C>>
+ CommittedInstanceVarOps<C>;
type Proof: Debug + Clone;
type ProofVar: Debug + Clone + AllocVar<Self::Proof, CF1<C>>;
/// Implements the constraints for NIFS.V for u and x, since cm(E) and cm(W) are delegated to
/// the CycleFold circuit.
#[allow(clippy::type_complexity)]
fn verify(
transcript: &mut T,
pp_hash: FpVar<CF1<C>>,
U_i: Self::CommittedInstanceVar,
// U_i_vec is passed to reuse the already computed U_i_vec from previous methods
U_i_vec: Vec<FpVar<CF1<C>>>,
u_i: Self::CommittedInstanceVar,
proof: Option<Self::ProofVar>,
) -> Result<(Self::CommittedInstanceVar, Vec<Boolean<CF1<C>>>), SynthesisError>;
}
/// These tests are the generic tests so that in the tests of Nova, Mova, Ova, we just need to
/// instantiate these tests to test both the NIFSTrait and NIFSGadgetTrait implementations for each
/// of the schemes.
#[cfg(test)] #[cfg(test)]
pub mod tests { pub mod tests {
use super::*;
use crate::transcript::poseidon::poseidon_canonical_config;
use ark_crypto_primitives::sponge::{
constraints::{AbsorbGadget, CryptographicSpongeVar},
poseidon::constraints::PoseidonSpongeVar,
Absorb,
};
use ark_crypto_primitives::sponge::{poseidon::PoseidonSponge, CryptographicSponge}; use ark_crypto_primitives::sponge::{poseidon::PoseidonSponge, CryptographicSponge};
use ark_pallas::{Fr, Projective}; 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}; use ark_std::{test_rng, UniformRand};
use super::NIFSTrait; use super::NIFSTrait;
use super::*;
use crate::arith::r1cs::tests::{get_test_r1cs, get_test_z}; use crate::arith::r1cs::tests::{get_test_r1cs, get_test_z};
use crate::commitment::pedersen::Pedersen; use crate::commitment::pedersen::Pedersen;
use crate::folding::traits::{CommittedInstanceOps, CommittedInstanceVarOps};
use crate::transcript::poseidon::poseidon_canonical_config;
/// Test method used to test the different implementations of the NIFSTrait (ie. Nova, Mova, /// Test method used to test the different implementations of the NIFSTrait (ie. Nova, Mova,
/// Ova). Runs a loop using the NIFS trait, and returns the last Witness and CommittedInstance /// Ova). Runs a loop using the NIFS trait, and returns the last Witness and CommittedInstance
@ -101,9 +150,7 @@ pub mod tests {
pub(crate) fn test_nifs_opt< pub(crate) fn test_nifs_opt<
N: NIFSTrait<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>, N: NIFSTrait<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
>() -> (N::Witness, N::CommittedInstance) { >() -> (N::Witness, N::CommittedInstance) {
let r1cs = get_test_r1cs();
let z = get_test_z(3);
let (w, x) = r1cs.split_z(&z);
let r1cs: R1CS<Fr> = get_test_r1cs();
let mut rng = ark_std::test_rng(); let mut rng = ark_std::test_rng();
let (pedersen_params, _) = Pedersen::<Projective>::setup(&mut rng, r1cs.A.n_cols).unwrap(); let (pedersen_params, _) = Pedersen::<Projective>::setup(&mut rng, r1cs.A.n_cols).unwrap();
@ -114,6 +161,8 @@ pub mod tests {
let pp_hash = Fr::rand(&mut rng); let pp_hash = Fr::rand(&mut rng);
// prepare the running instance // prepare the running instance
let z = get_test_z(3);
let (w, x) = r1cs.split_z(&z);
let mut W_i = N::new_witness(w.clone(), r1cs.A.n_rows, test_rng()); let mut W_i = N::new_witness(w.clone(), r1cs.A.n_rows, test_rng());
let mut U_i = N::new_instance(&mut rng, &pedersen_params, &W_i, x, vec![]).unwrap(); let mut U_i = N::new_instance(&mut rng, &pedersen_params, &W_i, x, vec![]).unwrap();
@ -149,4 +198,123 @@ pub mod tests {
(W_i, U_i) (W_i, U_i)
} }
/// Test method used to test the different implementations of the NIFSGadgetTrait (ie. Nova,
/// Mova, Ova). It returns the last Witness and CommittedInstance so that it can be checked at
/// the parent test that their values match.
pub(crate) fn test_nifs_gadget_opt<N, NG>(
ci: Vec<NG::CommittedInstance>,
proof: NG::Proof,
) -> Result<(NG::CommittedInstance, NG::CommittedInstanceVar), Error>
where
N: NIFSTrait<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NG: NIFSGadgetTrait<
Projective,
PoseidonSponge<Fr>,
PoseidonSpongeVar<Fr>,
CommittedInstance = N::CommittedInstance, // constrain that N::CI==NG::CI
Proof = N::Proof, // constrain that N::Proof==NG::Proof
>,
{
let mut rng = ark_std::test_rng();
let (U_i, u_i) = (ci[0].clone(), ci[1].clone());
let pp_hash = Fr::rand(&mut rng);
let poseidon_config = poseidon_canonical_config::<Fr>();
let mut transcript = PoseidonSponge::<Fr>::new(&poseidon_config);
let (ci3, _) = N::verify(&mut transcript, pp_hash, &U_i, &u_i, &proof)?;
let cs = ConstraintSystem::<Fr>::new_ref();
let mut transcriptVar = PoseidonSpongeVar::<Fr>::new(cs.clone(), &poseidon_config);
let pp_hashVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(pp_hash))?;
let ci1Var = NG::CommittedInstanceVar::new_witness(cs.clone(), || Ok(U_i.clone()))?;
let ci2Var = NG::CommittedInstanceVar::new_witness(cs.clone(), || Ok(u_i.clone()))?;
let proofVar = NG::ProofVar::new_witness(cs.clone(), || Ok(proof))?;
let ci1Var_vec = ci1Var.to_sponge_field_elements()?;
let (out, _) = NG::verify(
&mut transcriptVar,
pp_hashVar,
ci1Var.clone(),
ci1Var_vec,
ci2Var.clone(),
Some(proofVar.clone()),
)?;
assert!(cs.is_satisfied()?);
// return the NIFS.V and the NIFSGadget.V obtained values, so that they are checked at the
// parent test
Ok((ci3, out))
}
/// test that checks the native CommittedInstance.to_sponge_{bytes,field_elements}
/// vs the R1CS constraints version
pub(crate) fn test_committed_instance_to_sponge_preimage_opt<N, NG>(ci: N::CommittedInstance)
where
N: NIFSTrait<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NG: NIFSGadgetTrait<
Projective,
PoseidonSponge<Fr>,
PoseidonSpongeVar<Fr>,
CommittedInstance = N::CommittedInstance, // constrain that N::CI==NG::CI
>,
{
let bytes = ci.to_sponge_bytes_as_vec();
let field_elements = ci.to_sponge_field_elements_as_vec();
let cs = ConstraintSystem::<Fr>::new_ref();
let ciVar = NG::CommittedInstanceVar::new_witness(cs.clone(), || Ok(ci.clone())).unwrap();
let bytes_var = ciVar.to_sponge_bytes().unwrap();
let field_elements_var = ciVar.to_sponge_field_elements().unwrap();
assert!(cs.is_satisfied().unwrap());
// check that the natively computed and in-circuit computed hashes match
assert_eq!(bytes_var.value().unwrap(), bytes);
assert_eq!(field_elements_var.value().unwrap(), field_elements);
}
pub(crate) fn test_committed_instance_hash_opt<N, NG>(ci: NG::CommittedInstance)
where
N: NIFSTrait<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NG: NIFSGadgetTrait<
Projective,
PoseidonSponge<Fr>,
PoseidonSpongeVar<Fr>,
CommittedInstance = N::CommittedInstance, // constrain that N::CI==NG::CI
>,
N::CommittedInstance: CommittedInstanceOps<Projective>,
{
let poseidon_config = poseidon_canonical_config::<Fr>();
let sponge = PoseidonSponge::<Fr>::new(&poseidon_config);
let pp_hash = Fr::from(42u32); // only for test
let i = Fr::from(3_u32);
let z_0 = vec![Fr::from(3_u32)];
let z_i = vec![Fr::from(3_u32)];
// compute the CommittedInstance hash natively
let h = ci.hash(&sponge, pp_hash, i, &z_0, &z_i);
let cs = ConstraintSystem::<Fr>::new_ref();
let pp_hashVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(pp_hash)).unwrap();
let iVar = FpVar::<Fr>::new_witness(cs.clone(), || Ok(i)).unwrap();
let z_0Var = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(z_0.clone())).unwrap();
let z_iVar = Vec::<FpVar<Fr>>::new_witness(cs.clone(), || Ok(z_i.clone())).unwrap();
let ciVar = NG::CommittedInstanceVar::new_witness(cs.clone(), || Ok(ci.clone())).unwrap();
let sponge = PoseidonSpongeVar::<Fr>::new(cs.clone(), &poseidon_config);
// compute the CommittedInstance hash in-circuit
let (hVar, _) = ciVar
.hash(&sponge, &pp_hashVar, &iVar, &z_0Var, &z_iVar)
.unwrap();
assert!(cs.is_satisfied().unwrap());
// check that the natively computed and in-circuit computed hashes match
assert_eq!(hVar.value().unwrap(), h);
}
} }

+ 62
- 4
folding-schemes/src/folding/nova/nifs/nova.rs

@ -1,8 +1,10 @@
/// This module contains the implementation the NIFSTrait for the /// This module contains the implementation the NIFSTrait for the
/// [Nova](https://eprint.iacr.org/2021/370.pdf) NIFS (Non-Interactive Folding Scheme). /// [Nova](https://eprint.iacr.org/2021/370.pdf) NIFS (Non-Interactive Folding Scheme).
use ark_crypto_primitives::sponge::Absorb;
use ark_crypto_primitives::sponge::{constraints::AbsorbGadget, Absorb, CryptographicSponge};
use ark_ec::{CurveGroup, Group}; use ark_ec::{CurveGroup, Group};
use ark_ff::{BigInteger, PrimeField}; use ark_ff::{BigInteger, PrimeField};
use ark_r1cs_std::{boolean::Boolean, fields::fp::FpVar};
use ark_relations::r1cs::SynthesisError;
use ark_std::rand::RngCore; use ark_std::rand::RngCore;
use ark_std::Zero; use ark_std::Zero;
use std::marker::PhantomData; use std::marker::PhantomData;
@ -10,13 +12,69 @@ use std::marker::PhantomData;
use super::NIFSTrait; use super::NIFSTrait;
use crate::arith::r1cs::R1CS; use crate::arith::r1cs::R1CS;
use crate::commitment::CommitmentScheme; use crate::commitment::CommitmentScheme;
use crate::folding::circuits::cyclefold::{CycleFoldCommittedInstance, CycleFoldWitness};
use crate::folding::nova::circuits::ChallengeGadget;
use crate::constants::NOVA_N_BITS_RO;
use crate::folding::circuits::{
cyclefold::{CycleFoldCommittedInstance, CycleFoldWitness},
nonnative::affine::NonNativeAffineVar,
CF1,
};
use crate::folding::nova::{CommittedInstance, Witness}; use crate::folding::nova::{CommittedInstance, Witness};
use crate::transcript::Transcript;
use crate::transcript::{Transcript, TranscriptVar};
use crate::utils::vec::{hadamard, mat_vec_mul, vec_add, vec_scalar_mul, vec_sub}; use crate::utils::vec::{hadamard, mat_vec_mul, vec_add, vec_scalar_mul, vec_sub};
use crate::Error; use crate::Error;
/// ChallengeGadget computes the RO challenge used for the Nova instances NIFS, it contains a
/// rust-native and a in-circuit compatible versions.
pub struct ChallengeGadget<C: CurveGroup, CI: Absorb> {
_c: PhantomData<C>,
_ci: PhantomData<CI>,
}
impl<C: CurveGroup, CI: Absorb> ChallengeGadget<C, CI>
where
C: CurveGroup,
// <C as CurveGroup>::BaseField: PrimeField,
<C as Group>::ScalarField: Absorb,
{
pub fn get_challenge_native<T: Transcript<C::ScalarField>>(
transcript: &mut T,
pp_hash: C::ScalarField, // public params hash
U_i: &CI,
u_i: &CI,
cmT: Option<&C>,
) -> Vec<bool> {
transcript.absorb(&pp_hash);
transcript.absorb(&U_i);
transcript.absorb(&u_i);
// in the Nova case we absorb the cmT, in Ova case we don't since it is not used.
if let Some(cmT_value) = cmT {
transcript.absorb_nonnative(cmT_value);
}
transcript.squeeze_bits(NOVA_N_BITS_RO)
}
// compatible with the native get_challenge_native
pub fn get_challenge_gadget<
S: CryptographicSponge,
T: TranscriptVar<CF1<C>, S>,
CIVar: AbsorbGadget<CF1<C>>,
>(
transcript: &mut T,
pp_hash: FpVar<CF1<C>>, // public params hash
U_i_vec: Vec<FpVar<CF1<C>>>, // apready processed input, so we don't have to recompute these values
u_i: CIVar,
cmT: Option<NonNativeAffineVar<C>>,
) -> Result<Vec<Boolean<C::ScalarField>>, SynthesisError> {
transcript.absorb(&pp_hash)?;
transcript.absorb(&U_i_vec)?;
transcript.absorb(&u_i)?;
// in the Nova case we absorb the cmT, in Ova case we don't since it is not used.
if let Some(cmT_value) = cmT {
transcript.absorb_nonnative(&cmT_value)?;
}
transcript.squeeze_bits(NOVA_N_BITS_RO)
}
}
/// Implements the Non-Interactive Folding Scheme described in section 4 of /// Implements the Non-Interactive Folding Scheme described in section 4 of
/// [Nova](https://eprint.iacr.org/2021/370.pdf). /// [Nova](https://eprint.iacr.org/2021/370.pdf).
/// `H` specifies whether the NIFS will use a blinding factor /// `H` specifies whether the NIFS will use a blinding factor

+ 237
- 0
folding-schemes/src/folding/nova/nifs/nova_circuits.rs

@ -0,0 +1,237 @@
/// contains [Nova](https://eprint.iacr.org/2021/370.pdf) NIFS related circuits
use ark_crypto_primitives::sponge::{constraints::AbsorbGadget, Absorb, CryptographicSponge};
use ark_ec::{CurveGroup, Group};
use ark_r1cs_std::{
alloc::{AllocVar, AllocationMode},
boolean::Boolean,
eq::EqGadget,
fields::{fp::FpVar, FieldVar},
uint8::UInt8,
ToConstraintFieldGadget,
};
use ark_relations::r1cs::{ConstraintSystemRef, Namespace, SynthesisError};
use ark_std::{fmt::Debug, Zero};
use core::{borrow::Borrow, marker::PhantomData};
use super::NIFSGadgetTrait;
use crate::folding::circuits::{
nonnative::{affine::NonNativeAffineVar, uint::NonNativeUintVar},
CF1, CF2,
};
use crate::folding::nova::CommittedInstance;
use crate::folding::traits::CommittedInstanceVarOps;
use crate::transcript::TranscriptVar;
use super::nova::ChallengeGadget;
/// CommittedInstanceVar contains the u, x, cmE and cmW values which are folded on the main Nova
/// constraints field (E1::Fr, where E1 is the main curve). The peculiarity is that cmE and cmW are
/// represented non-natively over the constraint field.
#[derive(Debug, Clone)]
pub struct CommittedInstanceVar<C: CurveGroup> {
pub u: FpVar<C::ScalarField>,
pub x: Vec<FpVar<C::ScalarField>>,
pub cmE: NonNativeAffineVar<C>,
pub cmW: NonNativeAffineVar<C>,
}
impl<C> AllocVar<CommittedInstance<C>, CF1<C>> for CommittedInstanceVar<C>
where
C: CurveGroup,
{
fn new_variable<T: Borrow<CommittedInstance<C>>>(
cs: impl Into<Namespace<CF1<C>>>,
f: impl FnOnce() -> Result<T, SynthesisError>,
mode: AllocationMode,
) -> Result<Self, SynthesisError> {
f().and_then(|val| {
let cs = cs.into();
let u = FpVar::<C::ScalarField>::new_variable(cs.clone(), || Ok(val.borrow().u), mode)?;
let x: Vec<FpVar<C::ScalarField>> =
Vec::new_variable(cs.clone(), || Ok(val.borrow().x.clone()), mode)?;
let cmE =
NonNativeAffineVar::<C>::new_variable(cs.clone(), || Ok(val.borrow().cmE), mode)?;
let cmW =
NonNativeAffineVar::<C>::new_variable(cs.clone(), || Ok(val.borrow().cmW), mode)?;
Ok(Self { u, x, cmE, cmW })
})
}
}
impl<C> AbsorbGadget<C::ScalarField> for CommittedInstanceVar<C>
where
C: CurveGroup,
{
fn to_sponge_bytes(&self) -> Result<Vec<UInt8<C::ScalarField>>, SynthesisError> {
FpVar::batch_to_sponge_bytes(&self.to_sponge_field_elements()?)
}
fn to_sponge_field_elements(&self) -> Result<Vec<FpVar<C::ScalarField>>, SynthesisError> {
Ok([
vec![self.u.clone()],
self.x.clone(),
self.cmE.to_constraint_field()?,
self.cmW.to_constraint_field()?,
]
.concat())
}
}
impl<C: CurveGroup> CommittedInstanceVarOps<C> for CommittedInstanceVar<C> {
type PointVar = NonNativeAffineVar<C>;
fn get_commitments(&self) -> Vec<Self::PointVar> {
vec![self.cmW.clone(), self.cmE.clone()]
}
fn get_public_inputs(&self) -> &[FpVar<CF1<C>>] {
&self.x
}
fn enforce_incoming(&self) -> Result<(), SynthesisError> {
let zero = NonNativeUintVar::new_constant(ConstraintSystemRef::None, CF2::<C>::zero())?;
self.cmE.x.enforce_equal_unaligned(&zero)?;
self.cmE.y.enforce_equal_unaligned(&zero)?;
self.u.enforce_equal(&FpVar::one())
}
fn enforce_partial_equal(&self, other: &Self) -> Result<(), SynthesisError> {
self.u.enforce_equal(&other.u)?;
self.x.enforce_equal(&other.x)
}
}
/// Implements the circuit that does the checks of the Non-Interactive Folding Scheme Verifier
/// described in section 4 of [Nova](https://eprint.iacr.org/2021/370.pdf), where the cmE & cmW checks are
/// delegated to the NIFSCycleFoldGadget.
pub struct NIFSGadget<C: CurveGroup, S: CryptographicSponge, T: TranscriptVar<CF1<C>, S>> {
_c: PhantomData<C>,
_s: PhantomData<S>,
_t: PhantomData<T>,
}
impl<C, S, T> NIFSGadgetTrait<C, S, T> for NIFSGadget<C, S, T>
where
C: CurveGroup,
S: CryptographicSponge,
T: TranscriptVar<CF1<C>, S>,
<C as Group>::ScalarField: Absorb,
{
type CommittedInstance = CommittedInstance<C>;
type CommittedInstanceVar = CommittedInstanceVar<C>;
type Proof = C;
type ProofVar = NonNativeAffineVar<C>;
fn verify(
transcript: &mut T,
pp_hash: FpVar<CF1<C>>,
U_i: Self::CommittedInstanceVar,
// U_i_vec is passed to reuse the already computed U_i_vec from previous methods
U_i_vec: Vec<FpVar<CF1<C>>>,
u_i: Self::CommittedInstanceVar,
cmT: Option<Self::ProofVar>,
) -> Result<(Self::CommittedInstanceVar, Vec<Boolean<CF1<C>>>), SynthesisError> {
let r_bits = ChallengeGadget::<C, CommittedInstance<C>>::get_challenge_gadget(
transcript,
pp_hash.clone(),
U_i_vec,
u_i.clone(),
cmT.clone(),
)?;
let r = Boolean::le_bits_to_fp_var(&r_bits)?;
Ok((
Self::CommittedInstanceVar {
cmE: NonNativeAffineVar::new_constant(ConstraintSystemRef::None, C::zero())?,
cmW: NonNativeAffineVar::new_constant(ConstraintSystemRef::None, C::zero())?,
// ci3.u = U_i.u + r * u_i.u
u: U_i.u + &r * u_i.u,
// ci3.x = U_i.x + r * u_i.x
x: U_i
.x
.iter()
.zip(u_i.x)
.map(|(a, b)| a + &r * &b)
.collect::<Vec<FpVar<CF1<C>>>>(),
},
r_bits,
))
}
}
#[cfg(test)]
pub mod tests {
use super::*;
use ark_crypto_primitives::sponge::poseidon::constraints::PoseidonSpongeVar;
use ark_crypto_primitives::sponge::poseidon::PoseidonSponge;
use ark_pallas::{Fr, Projective};
use ark_r1cs_std::R1CSVar;
use ark_std::UniformRand;
use crate::commitment::pedersen::Pedersen;
use crate::folding::nova::nifs::{
nova::NIFS,
tests::{
test_committed_instance_hash_opt, test_committed_instance_to_sponge_preimage_opt,
test_nifs_gadget_opt,
},
};
#[test]
fn test_nifs_gadget() {
let mut rng = ark_std::test_rng();
// prepare the committed instances to test in-circuit
let ci: Vec<CommittedInstance<Projective>> = (0..2)
.into_iter()
.map(|_| CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
})
.collect();
let cmT = Projective::rand(&mut rng);
let (ci_out, ciVar_out) = test_nifs_gadget_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci, cmT)
.unwrap();
assert_eq!(ciVar_out.u.value().unwrap(), ci_out.u);
assert_eq!(ciVar_out.x.value().unwrap(), ci_out.x);
}
#[test]
fn test_committed_instance_to_sponge_preimage() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
};
test_committed_instance_to_sponge_preimage_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci);
}
#[test]
fn test_committed_instance_hash() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
cmE: Projective::rand(&mut rng),
u: Fr::rand(&mut rng),
cmW: Projective::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
};
test_committed_instance_hash_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci);
}
}

+ 28
- 5
folding-schemes/src/folding/nova/nifs/ova.rs

@ -9,10 +9,12 @@ use ark_std::rand::RngCore;
use ark_std::{One, UniformRand, Zero}; use ark_std::{One, UniformRand, Zero};
use std::marker::PhantomData; use std::marker::PhantomData;
use super::nova::ChallengeGadget;
use super::ova_circuits::CommittedInstanceVar;
use super::NIFSTrait; use super::NIFSTrait;
use crate::arith::r1cs::R1CS; use crate::arith::r1cs::R1CS;
use crate::commitment::CommitmentScheme; use crate::commitment::CommitmentScheme;
use crate::folding::nova::circuits::ChallengeGadget;
use crate::folding::traits::{CommittedInstanceOps, Inputize};
use crate::folding::{circuits::CF1, traits::Dummy}; use crate::folding::{circuits::CF1, traits::Dummy};
use crate::transcript::{AbsorbNonNative, Transcript}; use crate::transcript::{AbsorbNonNative, Transcript};
use crate::utils::vec::{hadamard, mat_vec_mul, vec_scalar_mul, vec_sub}; use crate::utils::vec::{hadamard, mat_vec_mul, vec_scalar_mul, vec_sub};
@ -50,6 +52,24 @@ where
} }
} }
impl<C: CurveGroup> CommittedInstanceOps<C> for CommittedInstance<C> {
type Var = CommittedInstanceVar<C>;
fn get_commitments(&self) -> Vec<C> {
vec![self.cmWE]
}
fn is_incoming(&self) -> bool {
self.u == One::one()
}
}
impl<C: CurveGroup> Inputize<C::ScalarField, CommittedInstanceVar<C>> for CommittedInstance<C> {
fn inputize(&self) -> Vec<C::ScalarField> {
[&[self.u][..], &self.x, &self.cmWE.inputize()].concat()
}
}
/// A Witness in Ova is represented by `w`. It also contains a blinder which can or not be used /// A Witness in Ova is represented by `w`. It also contains a blinder which can or not be used
/// when committing to the witness itself. /// when committing to the witness itself.
#[derive(Debug, Clone, Eq, PartialEq, CanonicalSerialize, CanonicalDeserialize)] #[derive(Debug, Clone, Eq, PartialEq, CanonicalSerialize, CanonicalDeserialize)]
@ -121,7 +141,9 @@ where
type CommittedInstance = CommittedInstance<C>; type CommittedInstance = CommittedInstance<C>;
type Witness = Witness<C>; type Witness = Witness<C>;
type ProverAux = (); type ProverAux = ();
type Proof = ();
// Proof is unused, but set to C::ScalarField so that the NIFSGadgetTrait abstraction can
// define the ProofsVar implementing the AllocVar from Proof
type Proof = C::ScalarField;
fn new_witness(w: Vec<C::ScalarField>, _e_len: usize, rng: impl RngCore) -> Self::Witness { fn new_witness(w: Vec<C::ScalarField>, _e_len: usize, rng: impl RngCore) -> Self::Witness {
Witness::new::<H>(w, rng) Witness::new::<H>(w, rng)
@ -182,11 +204,12 @@ where
let w = Self::fold_witness(r_Fr, W_i, w_i, &())?; let w = Self::fold_witness(r_Fr, W_i, w_i, &())?;
let (ci, _r_bits_v) = Self::verify(&mut transcript_v, pp_hash, U_i, u_i, &())?;
let proof = C::ScalarField::zero();
let (ci, _r_bits_v) = Self::verify(&mut transcript_v, pp_hash, U_i, u_i, &proof)?;
#[cfg(test)] #[cfg(test)]
assert_eq!(_r_bits_v, r_bits); assert_eq!(_r_bits_v, r_bits);
Ok((w, ci, (), r_bits))
Ok((w, ci, proof, r_bits))
} }
fn verify( fn verify(
@ -203,7 +226,7 @@ where
.ok_or(Error::OutOfBounds)?; .ok_or(Error::OutOfBounds)?;
// recall that r=alpha, and u=mu between Nova and Ova respectively // recall that r=alpha, and u=mu between Nova and Ova respectively
let u = U_i.u + r; // u_i.u is always 1 IN ova as we just can do sequential IVC.
let u = U_i.u + r; // u_i.u is always 1 in Ova as we just can do IVC (not PCD).
let cmWE = U_i.cmWE + u_i.cmWE.mul(r); let cmWE = U_i.cmWE + u_i.cmWE.mul(r);
let x = U_i let x = U_i
.x .x

+ 224
- 0
folding-schemes/src/folding/nova/nifs/ova_circuits.rs

@ -0,0 +1,224 @@
/// contains [Ova](https://hackmd.io/V4838nnlRKal9ZiTHiGYzw) NIFS related circuits
use ark_crypto_primitives::sponge::{constraints::AbsorbGadget, Absorb, CryptographicSponge};
use ark_ec::{CurveGroup, Group};
use ark_ff::PrimeField;
use ark_r1cs_std::{
alloc::{AllocVar, AllocationMode},
boolean::Boolean,
eq::EqGadget,
fields::{fp::FpVar, FieldVar},
uint8::UInt8,
ToConstraintFieldGadget,
};
use ark_relations::r1cs::{ConstraintSystemRef, Namespace, SynthesisError};
use ark_std::fmt::Debug;
use core::{borrow::Borrow, marker::PhantomData};
use super::ova::CommittedInstance;
use super::NIFSGadgetTrait;
use crate::folding::circuits::{nonnative::affine::NonNativeAffineVar, CF1};
use crate::folding::traits::CommittedInstanceVarOps;
use crate::transcript::TranscriptVar;
use crate::folding::nova::nifs::nova::ChallengeGadget;
#[derive(Debug, Clone)]
pub struct CommittedInstanceVar<C: CurveGroup> {
pub u: FpVar<C::ScalarField>,
pub x: Vec<FpVar<C::ScalarField>>,
pub cmWE: NonNativeAffineVar<C>,
}
impl<C> AllocVar<CommittedInstance<C>, CF1<C>> for CommittedInstanceVar<C>
where
C: CurveGroup,
{
fn new_variable<T: Borrow<CommittedInstance<C>>>(
cs: impl Into<Namespace<CF1<C>>>,
f: impl FnOnce() -> Result<T, SynthesisError>,
mode: AllocationMode,
) -> Result<Self, SynthesisError> {
f().and_then(|val| {
let cs = cs.into();
let u = FpVar::<C::ScalarField>::new_variable(cs.clone(), || Ok(val.borrow().u), mode)?;
let x: Vec<FpVar<C::ScalarField>> =
Vec::new_variable(cs.clone(), || Ok(val.borrow().x.clone()), mode)?;
let cmWE =
NonNativeAffineVar::<C>::new_variable(cs.clone(), || Ok(val.borrow().cmWE), mode)?;
Ok(Self { u, x, cmWE })
})
}
}
impl<C> AbsorbGadget<C::ScalarField> for CommittedInstanceVar<C>
where
C: CurveGroup,
<C as ark_ec::CurveGroup>::BaseField: ark_ff::PrimeField,
{
fn to_sponge_bytes(&self) -> Result<Vec<UInt8<C::ScalarField>>, SynthesisError> {
FpVar::batch_to_sponge_bytes(&self.to_sponge_field_elements()?)
}
fn to_sponge_field_elements(&self) -> Result<Vec<FpVar<C::ScalarField>>, SynthesisError> {
Ok([
vec![self.u.clone()],
self.x.clone(),
self.cmWE.to_constraint_field()?,
]
.concat())
}
}
impl<C: CurveGroup> CommittedInstanceVarOps<C> for CommittedInstanceVar<C> {
type PointVar = NonNativeAffineVar<C>;
fn get_commitments(&self) -> Vec<Self::PointVar> {
vec![self.cmWE.clone()]
}
fn get_public_inputs(&self) -> &[FpVar<CF1<C>>] {
&self.x
}
fn enforce_incoming(&self) -> Result<(), SynthesisError> {
self.u.enforce_equal(&FpVar::one())
}
fn enforce_partial_equal(&self, other: &Self) -> Result<(), SynthesisError> {
self.u.enforce_equal(&other.u)?;
self.x.enforce_equal(&other.x)
}
}
/// Implements the circuit that does the checks of the Non-Interactive Folding Scheme Verifier
/// described of the Ova variant, where the cmWE check is delegated to the NIFSCycleFoldGadget.
pub struct NIFSGadget<C: CurveGroup, S: CryptographicSponge, T: TranscriptVar<CF1<C>, S>> {
_c: PhantomData<C>,
_s: PhantomData<S>,
_t: PhantomData<T>,
}
impl<C, S, T> NIFSGadgetTrait<C, S, T> for NIFSGadget<C, S, T>
where
C: CurveGroup,
S: CryptographicSponge,
T: TranscriptVar<CF1<C>, S>,
<C as ark_ec::CurveGroup>::BaseField: ark_ff::PrimeField,
<C as Group>::ScalarField: Absorb,
<C as CurveGroup>::BaseField: PrimeField,
{
type CommittedInstance = CommittedInstance<C>;
type CommittedInstanceVar = CommittedInstanceVar<C>;
type Proof = C::ScalarField;
type ProofVar = FpVar<C::ScalarField>; // unused
fn verify(
transcript: &mut T,
pp_hash: FpVar<CF1<C>>,
U_i: Self::CommittedInstanceVar,
// U_i_vec is passed to reuse the already computed U_i_vec from previous methods
U_i_vec: Vec<FpVar<CF1<C>>>,
u_i: Self::CommittedInstanceVar,
_proof: Option<Self::ProofVar>,
) -> Result<(Self::CommittedInstanceVar, Vec<Boolean<CF1<C>>>), SynthesisError> {
let r_bits = ChallengeGadget::<C, CommittedInstance<C>>::get_challenge_gadget(
transcript,
pp_hash.clone(),
U_i_vec,
u_i.clone(),
None,
)?;
let r = Boolean::le_bits_to_fp_var(&r_bits)?;
Ok((
Self::CommittedInstanceVar {
cmWE: NonNativeAffineVar::new_constant(ConstraintSystemRef::None, C::zero())?,
// ci3.u = U_i.u + r * u_i.u (u_i.u is always 1 in Ova)
u: U_i.u + &r,
// ci3.x = U_i.x + r * u_i.x
x: U_i
.x
.iter()
.zip(u_i.x)
.map(|(a, b)| a + &r * &b)
.collect::<Vec<FpVar<CF1<C>>>>(),
},
r_bits,
))
}
}
#[cfg(test)]
pub mod tests {
use super::*;
use ark_crypto_primitives::sponge::poseidon::constraints::PoseidonSpongeVar;
use ark_crypto_primitives::sponge::poseidon::PoseidonSponge;
use ark_pallas::{Fr, Projective};
use ark_r1cs_std::R1CSVar;
use ark_std::UniformRand;
use ark_std::Zero;
use crate::commitment::pedersen::Pedersen;
use crate::folding::nova::nifs::{
ova::NIFS,
tests::{
test_committed_instance_hash_opt, test_committed_instance_to_sponge_preimage_opt,
test_nifs_gadget_opt,
},
};
#[test]
fn test_nifs_gadget() {
let mut rng = ark_std::test_rng();
// prepare the committed instances to test in-circuit
let ci: Vec<CommittedInstance<Projective>> = (0..2)
.into_iter()
.map(|_| CommittedInstance::<Projective> {
u: Fr::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
cmWE: Projective::rand(&mut rng),
})
.collect();
let (ci_out, ciVar_out) = test_nifs_gadget_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci, Fr::zero())
.unwrap();
assert_eq!(ciVar_out.u.value().unwrap(), ci_out.u);
assert_eq!(ciVar_out.x.value().unwrap(), ci_out.x);
}
#[test]
fn test_committed_instance_to_sponge_preimage() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
u: Fr::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
cmWE: Projective::rand(&mut rng),
};
test_committed_instance_to_sponge_preimage_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci);
}
#[test]
fn test_committed_instance_hash() {
let mut rng = ark_std::test_rng();
let ci = CommittedInstance::<Projective> {
u: Fr::rand(&mut rng),
x: vec![Fr::rand(&mut rng); 1],
cmWE: Projective::rand(&mut rng),
};
test_committed_instance_hash_opt::<
NIFS<Projective, Pedersen<Projective>, PoseidonSponge<Fr>>,
NIFSGadget<Projective, PoseidonSponge<Fr>, PoseidonSpongeVar<Fr>>,
>(ci);
}
}

+ 1
- 1
folding-schemes/src/folding/nova/traits.rs

@ -3,8 +3,8 @@ use ark_r1cs_std::fields::fp::FpVar;
use ark_relations::r1cs::SynthesisError; use ark_relations::r1cs::SynthesisError;
use ark_std::{rand::RngCore, UniformRand}; use ark_std::{rand::RngCore, UniformRand};
use super::circuits::CommittedInstanceVar;
use super::decider_eth_circuit::WitnessVar; use super::decider_eth_circuit::WitnessVar;
use super::nifs::nova_circuits::CommittedInstanceVar;
use super::{CommittedInstance, Witness}; use super::{CommittedInstance, Witness};
use crate::arith::{ use crate::arith::{
r1cs::{circuits::R1CSMatricesVar, R1CS}, r1cs::{circuits::R1CSMatricesVar, R1CS},

+ 1
- 1
frontends/src/noir/test_folder/test_mimc/Nargo.toml

@ -5,4 +5,4 @@ authors = [""]
compiler_version = ">=0.30.0" compiler_version = ">=0.30.0"
[dependencies] [dependencies]
mimc = { tag = "v0.1.0", git = "https://github.com/noir-lang/mimc" }

+ 1
- 3
frontends/src/noir/test_folder/test_mimc/src/main.nr

@ -1,6 +1,4 @@
use dep::std;
pub fn main(x: pub [Field; 1]) -> pub Field { pub fn main(x: pub [Field; 1]) -> pub Field {
let hash = std::hash::mimc::mimc_bn254(x);
let hash = mimc::mimc_bn254(x);
hash hash
} }

+ 0
- 3
solidity-verifiers/templates/nova_cyclefold_decider.askama.sol

@ -137,9 +137,6 @@ contract NovaDecider is Groth16Verifier, KZG10Verifier {
public_inputs[{{ z_len * 2 + 2 + num_limbs * 4 }} + 4 + k] = cmT_x_limbs[k]; public_inputs[{{ z_len * 2 + 2 + num_limbs * 4 }} + 4 + k] = cmT_x_limbs[k];
public_inputs[{{ z_len * 2 + 2 + num_limbs * 5 }} + 4 + k] = cmT_y_limbs[k]; public_inputs[{{ z_len * 2 + 2 + num_limbs * 5 }} + 4 + k] = cmT_y_limbs[k];
} }
// last element of the groth16 proof's public inputs is `r`
public_inputs[{{ public_inputs_len - 2 }}] = cmT_r[2];
bool success_g16 = this.verifyProof(pA, pB, pC, public_inputs); bool success_g16 = this.verifyProof(pA, pB, pC, public_inputs);
require(success_g16 == true, "Groth16: verifying proof failed"); require(success_g16 == true, "Groth16: verifying proof failed");

Loading…
Cancel
Save