Browse Source

MinRoot example improvements (#88)

* support multiple iterations of MinRoot per Nova step

* small edits to println

* fix declaration
main
Srinath Setty 2 years ago
committed by GitHub
parent
commit
a04566bb81
No known key found for this signature in database GPG Key ID: 4AEE18F83AFDEB23
2 changed files with 184 additions and 84 deletions
  1. +176
    -84
      examples/minroot.rs
  2. +8
    -0
      src/lib.rs

+ 176
- 84
examples/minroot.rs

@ -1,10 +1,8 @@
//! Demonstrates how to use Nova to produce a recursive proof of the correct execution of //! Demonstrates how to use Nova to produce a recursive proof of the correct execution of
//! iterations of the MinRoot function, thereby realizing a Nova-based verifiable delay function (VDF). //! iterations of the MinRoot function, thereby realizing a Nova-based verifiable delay function (VDF).
//! We currently execute a single iteration of the MinRoot function per step of Nova's recursion.
//! We execute a configurable number of iterations of the MinRoot function per step of Nova's recursion.
type G1 = pasta_curves::pallas::Point; type G1 = pasta_curves::pallas::Point;
type G2 = pasta_curves::vesta::Point; type G2 = pasta_curves::vesta::Point;
type S1 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G1>;
type S2 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G2>;
use ::bellperson::{gadgets::num::AllocatedNum, ConstraintSystem, SynthesisError}; use ::bellperson::{gadgets::num::AllocatedNum, ConstraintSystem, SynthesisError};
use ff::PrimeField; use ff::PrimeField;
use generic_array::typenum::U2; use generic_array::typenum::U2;
@ -19,42 +17,19 @@ use nova_snark::{
}; };
use num_bigint::BigUint; use num_bigint::BigUint;
use std::marker::PhantomData; use std::marker::PhantomData;
use std::time::Instant;
// A trivial test circuit that we will use on the secondary curve
#[derive(Clone, Debug)] #[derive(Clone, Debug)]
struct TrivialTestCircuit<F: PrimeField> {
_p: PhantomData<F>,
}
impl<F> StepCircuit<F> for TrivialTestCircuit<F>
where
F: PrimeField,
{
fn synthesize<CS: ConstraintSystem<F>>(
&self,
_cs: &mut CS,
z: AllocatedNum<F>,
) -> Result<AllocatedNum<F>, SynthesisError> {
Ok(z)
}
fn compute(&self, z: &F) -> F {
*z
}
}
#[derive(Clone, Debug)]
struct MinRootCircuit<F: PrimeField> {
struct MinRootIteration<F: PrimeField> {
x_i: F, x_i: F,
y_i: F, y_i: F,
x_i_plus_1: F, x_i_plus_1: F,
y_i_plus_1: F, y_i_plus_1: F,
pc: PoseidonConstants<F, U2>,
} }
impl<F: PrimeField> MinRootCircuit<F> {
impl<F: PrimeField> MinRootIteration<F> {
// produces a sample non-deterministic advice, executing one invocation of MinRoot per step // produces a sample non-deterministic advice, executing one invocation of MinRoot per step
fn new(num_steps: usize, x_0: &F, y_0: &F, pc: &PoseidonConstants<F, U2>) -> (F, Vec<Self>) {
fn new(num_iters: usize, x_0: &F, y_0: &F, pc: &PoseidonConstants<F, U2>) -> (F, Vec<Self>) {
// although this code is written generically, it is tailored to Pallas' scalar field // although this code is written generically, it is tailored to Pallas' scalar field
// (p - 3 / 5) // (p - 3 / 5)
let exp = BigUint::parse_bytes( let exp = BigUint::parse_bytes(
@ -66,7 +41,7 @@ impl MinRootCircuit {
let mut res = Vec::new(); let mut res = Vec::new();
let mut x_i = *x_0; let mut x_i = *x_0;
let mut y_i = *y_0; let mut y_i = *y_0;
for _i in 0..num_steps {
for _i in 0..num_iters {
let x_i_plus_1 = (x_i + y_i).pow_vartime(exp.to_u64_digits()); // computes the fifth root of x_i + y_i let x_i_plus_1 = (x_i + y_i).pow_vartime(exp.to_u64_digits()); // computes the fifth root of x_i + y_i
// sanity check // sanity check
@ -82,7 +57,6 @@ impl MinRootCircuit {
y_i, y_i,
x_i_plus_1, x_i_plus_1,
y_i_plus_1, y_i_plus_1,
pc: pc.clone(),
}); });
x_i = x_i_plus_1; x_i = x_i_plus_1;
@ -95,6 +69,12 @@ impl MinRootCircuit {
} }
} }
#[derive(Clone, Debug)]
struct MinRootCircuit<F: PrimeField> {
seq: Vec<MinRootIteration<F>>,
pc: PoseidonConstants<F, U2>,
}
impl<F> StepCircuit<F> for MinRootCircuit<F> impl<F> StepCircuit<F> for MinRootCircuit<F>
where where
F: PrimeField, F: PrimeField,
@ -104,65 +84,102 @@ where
cs: &mut CS, cs: &mut CS,
z: AllocatedNum<F>, z: AllocatedNum<F>,
) -> Result<AllocatedNum<F>, SynthesisError> { ) -> Result<AllocatedNum<F>, SynthesisError> {
// Allocate four variables for holding non-deterministic advice: x_i, y_i, x_i_plus_1, y_i_plus_1
let x_i = AllocatedNum::alloc(cs.namespace(|| "x_i"), || Ok(self.x_i))?;
let y_i = AllocatedNum::alloc(cs.namespace(|| "y_i"), || Ok(self.y_i))?;
let x_i_plus_1 = AllocatedNum::alloc(cs.namespace(|| "x_i_plus_1"), || Ok(self.x_i_plus_1))?;
// check that z = hash(x_i, y_i), where z is an output from the prior step
let z_hash = poseidon_hash(
cs.namespace(|| "input hash"),
vec![x_i.clone(), y_i.clone()],
&self.pc,
)?;
cs.enforce(
|| "z =? z_hash",
|lc| lc + z_hash.get_variable(),
|lc| lc + CS::one(),
|lc| lc + z.get_variable(),
);
let mut z_out: Result<AllocatedNum<F>, SynthesisError> = Err(SynthesisError::AssignmentMissing);
for i in 0..self.seq.len() {
// Allocate four variables for holding non-deterministic advice: x_i, y_i, x_i_plus_1, y_i_plus_1
let x_i = AllocatedNum::alloc(cs.namespace(|| format!("x_i_iter_{}", i)), || {
Ok(self.seq[i].x_i)
})?;
let y_i = AllocatedNum::alloc(cs.namespace(|| format!("y_i_iter_{}", i)), || {
Ok(self.seq[i].y_i)
})?;
let x_i_plus_1 =
AllocatedNum::alloc(cs.namespace(|| format!("x_i_plus_1_iter_{}", i)), || {
Ok(self.seq[i].x_i_plus_1)
})?;
// check the following conditions hold:
// (i) x_i_plus_1 = (x_i + y_i)^{1/5}, which can be more easily checked with x_i_plus_1^5 = x_i + y_i
// (ii) y_i_plus_1 = x_i
// (1) constraints for condition (i) are below
// (2) constraints for condition (ii) is avoided because we just used x_i wherever y_i_plus_1 is used
let x_i_plus_1_sq = x_i_plus_1.square(cs.namespace(|| "x_i_plus_1_sq"))?;
let x_i_plus_1_quad = x_i_plus_1_sq.square(cs.namespace(|| "x_i_plus_1_quad"))?;
let x_i_plus_1_pow_5 = x_i_plus_1_quad.mul(cs.namespace(|| "x_i_plus_1_pow_5"), &x_i_plus_1)?;
cs.enforce(
|| "x_i_plus_1_pow_5 = x_i + y_i",
|lc| lc + x_i_plus_1_pow_5.get_variable(),
|lc| lc + CS::one(),
|lc| lc + x_i.get_variable() + y_i.get_variable(),
);
// check that z = hash(x_i, y_i), where z is an output from the prior step
if i == 0 {
let z_hash = poseidon_hash(
cs.namespace(|| "input hash"),
vec![x_i.clone(), y_i.clone()],
&self.pc,
)?;
cs.enforce(
|| "z =? z_hash",
|lc| lc + z_hash.get_variable(),
|lc| lc + CS::one(),
|lc| lc + z.get_variable(),
);
}
// return hash(x_i_plus_1, y_i_plus_1) since Nova circuits expect a single output
poseidon_hash(
cs.namespace(|| "output hash"),
vec![x_i_plus_1, x_i.clone()],
&self.pc,
)
// check the following conditions hold:
// (i) x_i_plus_1 = (x_i + y_i)^{1/5}, which can be more easily checked with x_i_plus_1^5 = x_i + y_i
// (ii) y_i_plus_1 = x_i
// (1) constraints for condition (i) are below
// (2) constraints for condition (ii) is avoided because we just used x_i wherever y_i_plus_1 is used
let x_i_plus_1_sq =
x_i_plus_1.square(cs.namespace(|| format!("x_i_plus_1_sq_iter_{}", i)))?;
let x_i_plus_1_quad =
x_i_plus_1_sq.square(cs.namespace(|| format!("x_i_plus_1_quad_{}", i)))?;
let x_i_plus_1_pow_5 = x_i_plus_1_quad.mul(
cs.namespace(|| format!("x_i_plus_1_pow_5_{}", i)),
&x_i_plus_1,
)?;
cs.enforce(
|| format!("x_i_plus_1_pow_5 = x_i + y_i_iter_{}", i),
|lc| lc + x_i_plus_1_pow_5.get_variable(),
|lc| lc + CS::one(),
|lc| lc + x_i.get_variable() + y_i.get_variable(),
);
// return hash(x_i_plus_1, y_i_plus_1) since Nova circuits expect a single output
if i == self.seq.len() - 1 {
z_out = poseidon_hash(
cs.namespace(|| "output hash"),
vec![x_i_plus_1, x_i.clone()],
&self.pc,
);
}
}
z_out
} }
fn compute(&self, z: &F) -> F { fn compute(&self, z: &F) -> F {
// sanity check // sanity check
let z_hash = Poseidon::<F, U2>::new_with_preimage(&[self.x_i, self.y_i], &self.pc).hash();
let z_hash =
Poseidon::<F, U2>::new_with_preimage(&[self.seq[0].x_i, self.seq[0].y_i], &self.pc).hash();
debug_assert_eq!(z, &z_hash); debug_assert_eq!(z, &z_hash);
// compute output hash using advice // compute output hash using advice
Poseidon::<F, U2>::new_with_preimage(&[self.x_i_plus_1, self.y_i_plus_1], &self.pc).hash()
let iters = self.seq.len();
Poseidon::<F, U2>::new_with_preimage(
&[
self.seq[iters - 1].x_i_plus_1,
self.seq[iters - 1].y_i_plus_1,
],
&self.pc,
)
.hash()
} }
} }
fn main() { fn main() {
let pc = PoseidonConstants::<<G1 as Group>::Scalar, U2>::new_with_strength(Strength::Standard);
let num_steps = 10;
let num_iters_per_step = 10; // number of iterations of MinRoot per Nova's recursive step
let pc = PoseidonConstants::<<G1 as Group>::Scalar, U2>::new_with_strength(Strength::Standard);
let circuit_primary = MinRootCircuit { let circuit_primary = MinRootCircuit {
x_i: <G1 as Group>::Scalar::zero(),
y_i: <G1 as Group>::Scalar::zero(),
x_i_plus_1: <G1 as Group>::Scalar::zero(),
y_i_plus_1: <G1 as Group>::Scalar::zero(),
seq: vec![
MinRootIteration {
x_i: <G1 as Group>::Scalar::zero(),
y_i: <G1 as Group>::Scalar::zero(),
x_i_plus_1: <G1 as Group>::Scalar::zero(),
y_i_plus_1: <G1 as Group>::Scalar::zero(),
};
num_iters_per_step
],
pc: pc.clone(), pc: pc.clone(),
}; };
@ -170,22 +187,51 @@ fn main() {
_p: Default::default(), _p: Default::default(),
}; };
println!("Nova-based VDF with MinRoot delay function");
println!("==========================================");
println!(
"Proving {} iterations of MinRoot per step",
num_iters_per_step
);
// produce public parameters // produce public parameters
println!("Producing public parameters...");
let pp = PublicParams::< let pp = PublicParams::<
G1, G1,
G2, G2,
MinRootCircuit<<G1 as Group>::Scalar>, MinRootCircuit<<G1 as Group>::Scalar>,
TrivialTestCircuit<<G2 as Group>::Scalar>, TrivialTestCircuit<<G2 as Group>::Scalar>,
>::setup(circuit_primary, circuit_secondary.clone()); >::setup(circuit_primary, circuit_secondary.clone());
println!(
"Number of constraints per step (primary circuit): {}",
pp.num_constraints().0
);
println!(
"Number of constraints per step (secondary circuit): {}",
pp.num_constraints().1
);
// produce non-deterministic advice // produce non-deterministic advice
let num_steps = 3;
let (z0_primary, minroot_circuits) = MinRootCircuit::new(
num_steps,
let (z0_primary, minroot_iterations) = MinRootIteration::new(
num_iters_per_step * num_steps,
&<G1 as Group>::Scalar::zero(), &<G1 as Group>::Scalar::zero(),
&<G1 as Group>::Scalar::one(), &<G1 as Group>::Scalar::one(),
&pc, &pc,
); );
let minroot_circuits = (0..num_steps)
.map(|i| MinRootCircuit {
seq: (0..num_iters_per_step)
.map(|j| MinRootIteration {
x_i: minroot_iterations[i * num_iters_per_step + j].x_i,
y_i: minroot_iterations[i * num_iters_per_step + j].y_i,
x_i_plus_1: minroot_iterations[i * num_iters_per_step + j].x_i_plus_1,
y_i_plus_1: minroot_iterations[i * num_iters_per_step + j].y_i_plus_1,
})
.collect::<Vec<_>>(),
pc: pc.clone(),
})
.collect::<Vec<_>>();
let z0_secondary = <G2 as Group>::Scalar::zero(); let z0_secondary = <G2 as Group>::Scalar::zero();
type C1 = MinRootCircuit<<G1 as Group>::Scalar>; type C1 = MinRootCircuit<<G1 as Group>::Scalar>;
@ -195,6 +241,7 @@ fn main() {
let mut recursive_snark: Option<RecursiveSNARK<G1, G2, C1, C2>> = None; let mut recursive_snark: Option<RecursiveSNARK<G1, G2, C1, C2>> = None;
for (i, circuit_primary) in minroot_circuits.iter().take(num_steps).enumerate() { for (i, circuit_primary) in minroot_circuits.iter().take(num_steps).enumerate() {
let start = Instant::now();
let res = RecursiveSNARK::prove_step( let res = RecursiveSNARK::prove_step(
&pp, &pp,
recursive_snark, recursive_snark,
@ -204,7 +251,12 @@ fn main() {
z0_secondary, z0_secondary,
); );
assert!(res.is_ok()); assert!(res.is_ok());
println!("RecursiveSNARK::prove_step {}: {:?}", i, res.is_ok());
println!(
"RecursiveSNARK::prove_step {}: {:?}, took {:?} ",
i,
res.is_ok(),
start.elapsed()
);
recursive_snark = Some(res.unwrap()); recursive_snark = Some(res.unwrap());
} }
@ -213,20 +265,60 @@ fn main() {
// verify the recursive SNARK // verify the recursive SNARK
println!("Verifying a RecursiveSNARK..."); println!("Verifying a RecursiveSNARK...");
let start = Instant::now();
let res = recursive_snark.verify(&pp, num_steps, z0_primary, z0_secondary); let res = recursive_snark.verify(&pp, num_steps, z0_primary, z0_secondary);
println!("RecursiveSNARK::verify: {:?}", res.is_ok());
println!(
"RecursiveSNARK::verify: {:?}, took {:?}",
res.is_ok(),
start.elapsed()
);
assert!(res.is_ok()); assert!(res.is_ok());
// produce a compressed SNARK // produce a compressed SNARK
println!("Generating a CompressedSNARK...");
println!("Generating a CompressedSNARK using Spartan with IPA-PC...");
let start = Instant::now();
type S1 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G1>;
type S2 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G2>;
let res = CompressedSNARK::<_, _, _, _, S1, S2>::prove(&pp, &recursive_snark); let res = CompressedSNARK::<_, _, _, _, S1, S2>::prove(&pp, &recursive_snark);
println!("CompressedSNARK::prove: {:?}", res.is_ok());
println!(
"CompressedSNARK::prove: {:?}, took {:?}",
res.is_ok(),
start.elapsed()
);
assert!(res.is_ok()); assert!(res.is_ok());
let compressed_snark = res.unwrap(); let compressed_snark = res.unwrap();
// verify the compressed SNARK // verify the compressed SNARK
println!("Verifying a CompressedSNARK..."); println!("Verifying a CompressedSNARK...");
let start = Instant::now();
let res = compressed_snark.verify(&pp, num_steps, z0_primary, z0_secondary); let res = compressed_snark.verify(&pp, num_steps, z0_primary, z0_secondary);
println!("CompressedSNARK::verify: {:?}", res.is_ok());
println!(
"CompressedSNARK::verify: {:?}, took {:?}",
res.is_ok(),
start.elapsed()
);
assert!(res.is_ok()); assert!(res.is_ok());
} }
// A trivial test circuit that we use on the secondary curve
#[derive(Clone, Debug)]
struct TrivialTestCircuit<F: PrimeField> {
_p: PhantomData<F>,
}
impl<F> StepCircuit<F> for TrivialTestCircuit<F>
where
F: PrimeField,
{
fn synthesize<CS: ConstraintSystem<F>>(
&self,
_cs: &mut CS,
z: AllocatedNum<F>,
) -> Result<AllocatedNum<F>, SynthesisError> {
Ok(z)
}
fn compute(&self, z: &F) -> F {
*z
}
}

+ 8
- 0
src/lib.rs

@ -128,6 +128,14 @@ where
_p_c2: Default::default(), _p_c2: Default::default(),
} }
} }
/// Returns the number of constraints in the primary and secondary circuits
pub fn num_constraints(&self) -> (usize, usize) {
(
self.r1cs_shape_primary.num_cons,
self.r1cs_shape_secondary.num_cons,
)
}
} }
/// A SNARK that proves the correct execution of an incremental computation /// A SNARK that proves the correct execution of an incremental computation

Loading…
Cancel
Save