Browse Source

ecdsa signature proof (#92)

* ecdsa signature proof

* use the library-provided default circuit

* small reorg

Co-authored-by: Srinath Setty <srinath@microsoft.com>
main
Arthur Greef 2 years ago
committed by GitHub
parent
commit
ed915b2540
No known key found for this signature in database GPG Key ID: 4AEE18F83AFDEB23
3 changed files with 655 additions and 0 deletions
  1. +317
    -0
      examples/ecdsa/circuit.rs
  2. +309
    -0
      examples/ecdsa/main.rs
  3. +29
    -0
      examples/ecdsa/utils.rs

+ 317
- 0
examples/ecdsa/circuit.rs

@ -0,0 +1,317 @@
use bellperson::{
gadgets::{boolean::AllocatedBit, num::AllocatedNum},
ConstraintSystem, SynthesisError,
};
use ff::{PrimeField, PrimeFieldBits};
use generic_array::typenum::U8;
use neptune::{
circuit::poseidon_hash,
poseidon::{Poseidon, PoseidonConstants},
};
use nova_snark::{gadgets::ecc::AllocatedPoint, traits::circuit::StepCircuit};
use subtle::Choice;
// An affine point coordinate that is on the curve.
#[derive(Clone, Copy, Debug)]
pub struct Coordinate<F>
where
F: PrimeField<Repr = [u8; 32]>,
{
pub x: F,
pub y: F,
pub is_infinity: bool,
}
impl<F> Coordinate<F>
where
F: PrimeField<Repr = [u8; 32]>,
{
// New affine point coordiante on the curve so is_infinity = false.
pub fn new(x: F, y: F) -> Self {
Self {
x,
y,
is_infinity: false,
}
}
}
// An ECDSA signature
#[derive(Clone, Debug)]
pub struct EcdsaSignature<Fb, Fs>
where
Fb: PrimeField<Repr = [u8; 32]>,
Fs: PrimeField<Repr = [u8; 32]> + PrimeFieldBits,
{
pk: Coordinate<Fb>, // public key
r: Coordinate<Fb>, // (r, s) is the ECDSA signature
s: Fs,
c: Fs, // hash of the message
g: Coordinate<Fb>, // generator of the group; could be omitted if Nova's traits allow accessing the generator
}
impl<Fb, Fs> EcdsaSignature<Fb, Fs>
where
Fb: PrimeField<Repr = [u8; 32]>,
Fs: PrimeField<Repr = [u8; 32]> + PrimeFieldBits,
{
pub fn new(pk: Coordinate<Fb>, r: Coordinate<Fb>, s: Fs, c: Fs, g: Coordinate<Fb>) -> Self {
Self { pk, r, s, c, g }
}
}
// An ECDSA signature proof that we will use on the primary curve
#[derive(Clone, Debug)]
pub struct EcdsaCircuit<F>
where
F: PrimeField<Repr = [u8; 32]>,
{
pub z_r: Coordinate<F>,
pub z_g: Coordinate<F>,
pub z_pk: Coordinate<F>,
pub z_c: F,
pub z_s: F,
pub r: Coordinate<F>,
pub g: Coordinate<F>,
pub pk: Coordinate<F>,
pub c: F,
pub s: F,
pub c_bits: Vec<Choice>,
pub s_bits: Vec<Choice>,
pub pc: PoseidonConstants<F, U8>,
}
impl<F> EcdsaCircuit<F>
where
F: PrimeField<Repr = [u8; 32]>,
{
// Creates a new [`EcdsaCircuit<Fb, Fs>`]. The base and scalar field elements from the curve
// field used by the signature are converted to scalar field elements from the cyclic curve
// field used by the circuit.
pub fn new<Fb, Fs>(
num_steps: usize,
signatures: &[EcdsaSignature<Fb, Fs>],
pc: &PoseidonConstants<F, U8>,
) -> (F, Vec<Self>)
where
Fb: PrimeField<Repr = [u8; 32]>,
Fs: PrimeField<Repr = [u8; 32]> + PrimeFieldBits,
{
let mut z0 = F::zero();
let mut circuits = Vec::new();
for i in 0..num_steps {
let mut j = i;
if i > 0 {
j = i - 1
};
let z_signature = &signatures[j];
let z_r = Coordinate::new(
F::from_repr(z_signature.r.x.to_repr()).unwrap(),
F::from_repr(z_signature.r.y.to_repr()).unwrap(),
);
let z_g = Coordinate::new(
F::from_repr(z_signature.g.x.to_repr()).unwrap(),
F::from_repr(z_signature.g.y.to_repr()).unwrap(),
);
let z_pk = Coordinate::new(
F::from_repr(z_signature.pk.x.to_repr()).unwrap(),
F::from_repr(z_signature.pk.y.to_repr()).unwrap(),
);
let z_c = F::from_repr(z_signature.c.to_repr()).unwrap();
let z_s = F::from_repr(z_signature.s.to_repr()).unwrap();
let signature = &signatures[i];
let r = Coordinate::new(
F::from_repr(signature.r.x.to_repr()).unwrap(),
F::from_repr(signature.r.y.to_repr()).unwrap(),
);
let g = Coordinate::new(
F::from_repr(signature.g.x.to_repr()).unwrap(),
F::from_repr(signature.g.y.to_repr()).unwrap(),
);
let pk = Coordinate::new(
F::from_repr(signature.pk.x.to_repr()).unwrap(),
F::from_repr(signature.pk.y.to_repr()).unwrap(),
);
let c_bits = Self::to_le_bits(&signature.c);
let s_bits = Self::to_le_bits(&signature.s);
let c = F::from_repr(signature.c.to_repr()).unwrap();
let s = F::from_repr(signature.s.to_repr()).unwrap();
let circuit = EcdsaCircuit {
z_r,
z_g,
z_pk,
z_c,
z_s,
r,
g,
pk,
c,
s,
c_bits,
s_bits,
pc: pc.clone(),
};
circuits.push(circuit);
if i == 0 {
z0 =
Poseidon::<F, U8>::new_with_preimage(&[r.x, r.y, g.x, g.y, pk.x, pk.y, c, s], pc).hash();
}
}
(z0, circuits)
}
// Converts the scalar field element from the curve used by the signature to a bit represenation
// for later use in scalar multiplication using the cyclic curve used by the circuit.
fn to_le_bits<Fs>(fs: &Fs) -> Vec<Choice>
where
Fs: PrimeField<Repr = [u8; 32]> + PrimeFieldBits,
{
let bits = fs
.to_repr()
.iter()
.flat_map(|byte| (0..8).map(move |i| Choice::from((byte >> i) & 1u8)))
.collect::<Vec<Choice>>();
bits
}
// Synthesize a bit representation into circuit gadgets.
fn synthesize_bits<CS: ConstraintSystem<F>>(
cs: &mut CS,
bits: &[Choice],
) -> Result<Vec<AllocatedBit>, SynthesisError> {
let alloc_bits: Vec<AllocatedBit> = bits
.iter()
.enumerate()
.map(|(i, bit)| {
AllocatedBit::alloc(
cs.namespace(|| format!("bit {}", i)),
Some(bit.unwrap_u8() == 1u8),
)
})
.collect::<Result<Vec<AllocatedBit>, SynthesisError>>()
.unwrap();
Ok(alloc_bits)
}
}
impl<F> StepCircuit<F> for EcdsaCircuit<F>
where
F: PrimeField<Repr = [u8; 32]> + PrimeFieldBits,
{
// Prove knowledge of the sk used to generate the Ecdsa signature (R,s)
// with public key PK and message commitment c.
// [s]G == R + [c]PK
fn synthesize<CS: ConstraintSystem<F>>(
&self,
cs: &mut CS,
z: AllocatedNum<F>,
) -> Result<AllocatedNum<F>, SynthesisError> {
let z_rx = AllocatedNum::alloc(cs.namespace(|| "z_rx"), || Ok(self.z_r.x))?;
let z_ry = AllocatedNum::alloc(cs.namespace(|| "z_ry"), || Ok(self.z_r.y))?;
let z_gx = AllocatedNum::alloc(cs.namespace(|| "z_gx"), || Ok(self.z_g.x))?;
let z_gy = AllocatedNum::alloc(cs.namespace(|| "z_gy"), || Ok(self.z_g.y))?;
let z_pkx = AllocatedNum::alloc(cs.namespace(|| "z_pkx"), || Ok(self.z_pk.x))?;
let z_pky = AllocatedNum::alloc(cs.namespace(|| "z_pky"), || Ok(self.z_pk.y))?;
let z_c = AllocatedNum::alloc(cs.namespace(|| "z_c"), || Ok(self.z_c))?;
let z_s = AllocatedNum::alloc(cs.namespace(|| "z_s"), || Ok(self.z_s))?;
let z_hash = poseidon_hash(
cs.namespace(|| "input hash"),
vec![z_rx, z_ry, z_gx, z_gy, z_pkx, z_pky, z_c, z_s],
&self.pc,
)?;
cs.enforce(
|| "z == z1",
|lc| lc + z.get_variable(),
|lc| lc + CS::one(),
|lc| lc + z_hash.get_variable(),
);
let g = AllocatedPoint::alloc(
cs.namespace(|| "G"),
Some((self.g.x, self.g.y, self.g.is_infinity)),
)?;
let s_bits = Self::synthesize_bits(&mut cs.namespace(|| "s_bits"), &self.s_bits)?;
let sg = g.scalar_mul(cs.namespace(|| "[s]G"), s_bits)?;
let r = AllocatedPoint::alloc(
cs.namespace(|| "R"),
Some((self.r.x, self.r.y, self.r.is_infinity)),
)?;
let c_bits = Self::synthesize_bits(&mut cs.namespace(|| "c_bits"), &self.c_bits)?;
let pk = AllocatedPoint::alloc(
cs.namespace(|| "PK"),
Some((self.pk.x, self.pk.y, self.pk.is_infinity)),
)?;
let cpk = pk.scalar_mul(&mut cs.namespace(|| "[c]PK"), c_bits)?;
let rcpk = cpk.add(&mut cs.namespace(|| "R + [c]PK"), &r)?;
let (rcpk_x, rcpk_y, _) = rcpk.get_coordinates();
let (sg_x, sg_y, _) = sg.get_coordinates();
cs.enforce(
|| "sg_x == rcpk_x",
|lc| lc + sg_x.get_variable(),
|lc| lc + CS::one(),
|lc| lc + rcpk_x.get_variable(),
);
cs.enforce(
|| "sg_y == rcpk_y",
|lc| lc + sg_y.get_variable(),
|lc| lc + CS::one(),
|lc| lc + rcpk_y.get_variable(),
);
let rx = AllocatedNum::alloc(cs.namespace(|| "rx"), || Ok(self.r.x))?;
let ry = AllocatedNum::alloc(cs.namespace(|| "ry"), || Ok(self.r.y))?;
let gx = AllocatedNum::alloc(cs.namespace(|| "gx"), || Ok(self.g.x))?;
let gy = AllocatedNum::alloc(cs.namespace(|| "gy"), || Ok(self.g.y))?;
let pkx = AllocatedNum::alloc(cs.namespace(|| "pkx"), || Ok(self.pk.x))?;
let pky = AllocatedNum::alloc(cs.namespace(|| "pky"), || Ok(self.pk.y))?;
let c = AllocatedNum::alloc(cs.namespace(|| "c"), || Ok(self.c))?;
let s = AllocatedNum::alloc(cs.namespace(|| "s"), || Ok(self.s))?;
poseidon_hash(
cs.namespace(|| "output hash"),
vec![rx, ry, gx, gy, pkx, pky, c, s],
&self.pc,
)
}
fn compute(&self, z: &F) -> F {
let z_hash = Poseidon::<F, U8>::new_with_preimage(
&[
self.z_r.x,
self.z_r.y,
self.z_g.x,
self.z_g.y,
self.z_pk.x,
self.z_pk.y,
self.z_c,
self.z_s,
],
&self.pc,
)
.hash();
debug_assert_eq!(z, &z_hash);
Poseidon::<F, U8>::new_with_preimage(
&[
self.r.x, self.r.y, self.g.x, self.g.y, self.pk.x, self.pk.y, self.c, self.s,
],
&self.pc,
)
.hash()
}
}

+ 309
- 0
examples/ecdsa/main.rs

@ -0,0 +1,309 @@
//! Demonstrates how to use Nova to produce a recursive proof of an ECDSA signature.
//! This example proves the knowledge of a sequence of ECDSA signatures with different public keys on different messages,
//! but the example can be adapted to other settings (e.g., proving the validity of the certificate chain with a well-known root public key)
//! Scheme borrowed from https://github.com/filecoin-project/bellperson-gadgets/blob/main/src/eddsa.rs
//! Sign using G1 curve, and prove using G2 curve.
use core::ops::{Add, AddAssign, Mul, MulAssign, Neg};
use ff::{
derive::byteorder::{ByteOrder, LittleEndian},
Field, PrimeField, PrimeFieldBits,
};
use generic_array::typenum::U8;
use neptune::{poseidon::PoseidonConstants, Strength};
use nova_snark::{
traits::{circuit::TrivialTestCircuit, Group as Nova_Group},
CompressedSNARK, PublicParams, RecursiveSNARK,
};
use num_bigint::BigUint;
use pasta_curves::{
arithmetic::CurveAffine,
group::{Curve, Group},
};
use rand::{rngs::OsRng, RngCore};
use sha3::{Digest, Sha3_512};
use subtle::Choice;
mod circuit;
mod utils;
use crate::circuit::{Coordinate, EcdsaCircuit, EcdsaSignature};
use crate::utils::BitIterator;
type G1 = pasta_curves::pallas::Point;
type G2 = pasta_curves::vesta::Point;
type S1 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G2>;
type S2 = nova_snark::spartan_with_ipa_pc::RelaxedR1CSSNARK<G1>;
#[derive(Debug, Clone, Copy)]
pub struct SecretKey(pub <G1 as Group>::Scalar);
impl SecretKey {
pub fn random(mut rng: impl RngCore) -> Self {
let secret = <G1 as Group>::Scalar::random(&mut rng);
Self(secret)
}
}
#[derive(Debug, Clone, Copy)]
pub struct PublicKey(pub G1);
impl PublicKey {
pub fn from_secret_key(s: &SecretKey) -> Self {
let point = G1::generator() * s.0;
Self(point)
}
}
#[derive(Clone)]
pub struct Signature {
pub r: G1,
pub s: <G1 as Group>::Scalar,
}
impl SecretKey {
pub fn sign(self, c: <G1 as Group>::Scalar, mut rng: impl RngCore) -> Signature {
// T
let mut t = [0u8; 80];
rng.fill_bytes(&mut t[..]);
// h = H*(T || M)
let h = Self::hash_to_scalar(b"Nova_Ecdsa_Hash", &t[..], &c.to_repr());
// R = [h]G
let r = G1::generator().mul(h);
// s = h + c * sk
let mut s = c;
s.mul_assign(&self.0);
s.add_assign(&h);
Signature { r, s }
}
fn mul_bits<B: AsRef<[u64]>>(
s: &<G1 as Group>::Scalar,
bits: BitIterator<B>,
) -> <G1 as Group>::Scalar {
let mut x = <G1 as Group>::Scalar::zero();
for bit in bits {
x.double();
if bit {
x.add_assign(s)
}
}
x
}
fn to_uniform(digest: &[u8]) -> <G1 as Group>::Scalar {
assert_eq!(digest.len(), 64);
let mut bits: [u64; 8] = [0; 8];
LittleEndian::read_u64_into(digest, &mut bits);
Self::mul_bits(&<G1 as Group>::Scalar::one(), BitIterator::new(bits))
}
pub fn to_uniform_32(digest: &[u8]) -> <G1 as Group>::Scalar {
assert_eq!(digest.len(), 32);
let mut bits: [u64; 4] = [0; 4];
LittleEndian::read_u64_into(digest, &mut bits);
Self::mul_bits(&<G1 as Group>::Scalar::one(), BitIterator::new(bits))
}
pub fn hash_to_scalar(persona: &[u8], a: &[u8], b: &[u8]) -> <G1 as Group>::Scalar {
let mut hasher = Sha3_512::new();
hasher.input(persona);
hasher.input(a);
hasher.input(b);
let digest = hasher.result();
Self::to_uniform(digest.as_ref())
}
}
impl PublicKey {
pub fn verify(&self, c: <G1 as Group>::Scalar, signature: &Signature) -> bool {
let modulus = Self::modulus_as_scalar();
let order_check_pk = self.0.mul(modulus);
if !order_check_pk.eq(&G1::identity()) {
return false;
}
let order_check_r = signature.r.mul(modulus);
if !order_check_r.eq(&G1::identity()) {
return false;
}
// 0 = [-s]G + R + [c]PK
self
.0
.mul(c)
.add(&signature.r)
.add(G1::generator().mul(signature.s).neg())
.eq(&G1::identity())
}
fn modulus_as_scalar() -> <G1 as Group>::Scalar {
let mut bits = <G1 as Group>::Scalar::char_le_bits().to_bitvec();
let mut acc = BigUint::new(Vec::<u32>::new());
while let Some(b) = bits.pop() {
acc <<= 1_i32;
acc += b as u8;
}
let modulus = acc.to_str_radix(10);
<G1 as Group>::Scalar::from_str_vartime(&modulus).unwrap()
}
}
fn main() {
// In a VERY LIMITED case of messages known to be unique due to application level
// and being less than the group order when interpreted as integer, one can sign
// the message directly without hashing
pub const MAX_MESSAGE_LEN: usize = 16;
assert!(MAX_MESSAGE_LEN * 8 <= <G1 as Group>::Scalar::CAPACITY as usize);
// produce public parameters
println!("Generating public parameters...");
let pc = PoseidonConstants::<<G2 as Group>::Scalar, U8>::new_with_strength(Strength::Standard);
let circuit_primary = EcdsaCircuit::<<G2 as Nova_Group>::Scalar> {
z_r: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
z_g: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
z_pk: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
z_c: <G2 as Nova_Group>::Scalar::zero(),
z_s: <G2 as Nova_Group>::Scalar::zero(),
r: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
g: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
pk: Coordinate::new(
<G2 as Nova_Group>::Scalar::zero(),
<G2 as Nova_Group>::Scalar::zero(),
),
c: <G2 as Nova_Group>::Scalar::zero(),
s: <G2 as Nova_Group>::Scalar::zero(),
c_bits: vec![Choice::from(0u8); 256],
s_bits: vec![Choice::from(0u8); 256],
pc: pc.clone(),
};
let circuit_secondary = TrivialTestCircuit::default();
let pp = PublicParams::<
G2,
G1,
EcdsaCircuit<<G2 as Group>::Scalar>,
TrivialTestCircuit<<G1 as Group>::Scalar>,
>::setup(circuit_primary, circuit_secondary.clone());
// produce non-deterministic advice
println!("Generating non-deterministic advice...");
let num_steps = 3;
let signatures = || {
let mut signatures = Vec::new();
for i in 0..num_steps {
let sk = SecretKey::random(&mut OsRng);
let pk = PublicKey::from_secret_key(&sk);
let message = format!("MESSAGE{}", i).as_bytes().to_owned();
assert!(message.len() <= MAX_MESSAGE_LEN);
let mut digest: Vec<u8> = message.to_vec();
for _ in 0..(32 - message.len() as u32) {
digest.extend(&[0u8; 1]);
}
let c = SecretKey::to_uniform_32(digest.as_ref());
let signature_primary = sk.sign(c, &mut OsRng);
let result = pk.verify(c, &signature_primary);
assert!(result);
// Affine coordinates guaranteed to be on the curve
let rxy = signature_primary.r.to_affine().coordinates().unwrap();
let gxy = G1::generator().to_affine().coordinates().unwrap();
let pkxy = pk.0.to_affine().coordinates().unwrap();
let s = signature_primary.s;
signatures.push(EcdsaSignature::<
<G1 as Nova_Group>::Base,
<G1 as Nova_Group>::Scalar,
>::new(
Coordinate::<<G1 as Nova_Group>::Base>::new(*pkxy.x(), *pkxy.y()),
Coordinate::<<G1 as Nova_Group>::Base>::new(*rxy.x(), *rxy.y()),
s,
c,
Coordinate::<<G1 as Nova_Group>::Base>::new(*gxy.x(), *gxy.y()),
));
}
signatures
};
let (z0_primary, circuits_primary) = EcdsaCircuit::<<G2 as Nova_Group>::Scalar>::new::<
<G1 as Nova_Group>::Base,
<G1 as Nova_Group>::Scalar,
>(num_steps, &signatures(), &pc);
// Secondary circuit
let z0_secondary = <G1 as Group>::Scalar::zero();
// produce a recursive SNARK
println!("Generating a RecursiveSNARK...");
type C1 = EcdsaCircuit<<G2 as Nova_Group>::Scalar>;
type C2 = TrivialTestCircuit<<G1 as Nova_Group>::Scalar>;
let mut recursive_snark: Option<RecursiveSNARK<G2, G1, C1, C2>> = None;
for (i, circuit_primary) in circuits_primary.iter().take(num_steps).enumerate() {
let result = RecursiveSNARK::prove_step(
&pp,
recursive_snark,
circuit_primary.clone(),
circuit_secondary.clone(),
z0_primary,
z0_secondary,
);
assert!(result.is_ok());
println!("RecursiveSNARK::prove_step {}: {:?}", i, result.is_ok());
recursive_snark = Some(result.unwrap());
}
assert!(recursive_snark.is_some());
let recursive_snark = recursive_snark.unwrap();
// verify the recursive SNARK
println!("Verifying the RecursiveSNARK...");
let res = recursive_snark.verify(&pp, num_steps, z0_primary, z0_secondary);
println!("RecursiveSNARK::verify: {:?}", res.is_ok());
assert!(res.is_ok());
// produce a compressed SNARK
println!("Generating a CompressedSNARK...");
let res = CompressedSNARK::<_, _, _, _, S1, S2>::prove(&pp, &recursive_snark);
println!("CompressedSNARK::prove: {:?}", res.is_ok());
assert!(res.is_ok());
let compressed_snark = res.unwrap();
// verify the compressed SNARK
println!("Verifying a CompressedSNARK...");
let res = compressed_snark.verify(&pp, num_steps, z0_primary, z0_secondary);
println!("CompressedSNARK::verify: {:?}", res.is_ok());
assert!(res.is_ok());
}

+ 29
- 0
examples/ecdsa/utils.rs

@ -0,0 +1,29 @@
#[derive(Debug)]
pub struct BitIterator<E> {
t: E,
n: usize,
}
impl<E: AsRef<[u64]>> BitIterator<E> {
pub fn new(t: E) -> Self {
let n = t.as_ref().len() * 64;
BitIterator { t, n }
}
}
impl<E: AsRef<[u64]>> Iterator for BitIterator<E> {
type Item = bool;
fn next(&mut self) -> Option<bool> {
if self.n == 0 {
None
} else {
self.n -= 1;
let part = self.n / 64;
let bit = self.n - (64 * part);
Some(self.t.as_ref()[part] & (1 << bit) > 0)
}
}
}

Loading…
Cancel
Save