Accept more types in threshold crypto API.

This removes some unnecessary allocation and conversion by accepting
more primitive types and references as the index in threshold decryption
and signing, and as the argument to a polynomial.
master
Andreas Fackler 7 years ago committed by Vladimir Komendantskiy
parent a78a14fa9d
commit 23e8ebe967
  1. 55
      into_fr.rs
  2. 44
      mod.rs
  3. 81
      poly.rs

@ -0,0 +1,55 @@
use pairing::bls12_381::Fr;
use pairing::{Field, PrimeField};
/// A conversion into an element of the field `Fr`.
pub trait IntoFr: Copy {
fn into_fr(self) -> Fr;
}
impl IntoFr for Fr {
fn into_fr(self) -> Fr {
self
}
}
impl IntoFr for u64 {
fn into_fr(self) -> Fr {
Fr::from_repr(self.into()).expect("modulus is greater than u64::MAX")
}
}
impl IntoFr for usize {
fn into_fr(self) -> Fr {
(self as u64).into_fr()
}
}
impl IntoFr for i32 {
fn into_fr(self) -> Fr {
if self >= 0 {
(self as u64).into_fr()
} else {
let mut result = ((-self) as u64).into_fr();
result.negate();
result
}
}
}
impl IntoFr for i64 {
fn into_fr(self) -> Fr {
if self >= 0 {
(self as u64).into_fr()
} else {
let mut result = ((-self) as u64).into_fr();
result.negate();
result
}
}
}
impl<'a, T: IntoFr> IntoFr for &'a T {
fn into_fr(self) -> Fr {
(*self).into_fr()
}
}

@ -3,6 +3,7 @@
#![cfg_attr(feature = "cargo-clippy", allow(derive_hash_xor_eq))] #![cfg_attr(feature = "cargo-clippy", allow(derive_hash_xor_eq))]
pub mod error; pub mod error;
mod into_fr;
pub mod poly; pub mod poly;
#[cfg(feature = "serialization-protobuf")] #[cfg(feature = "serialization-protobuf")]
pub mod protobuf_impl; pub mod protobuf_impl;
@ -14,12 +15,13 @@ use std::ptr::write_volatile;
use byteorder::{BigEndian, ByteOrder}; use byteorder::{BigEndian, ByteOrder};
use init_with::InitWith; use init_with::InitWith;
use pairing::bls12_381::{Bls12, Fr, FrRepr, G1, G1Affine, G2, G2Affine}; use pairing::bls12_381::{Bls12, Fr, G1, G1Affine, G2, G2Affine};
use pairing::{CurveAffine, CurveProjective, Engine, Field, PrimeField}; use pairing::{CurveAffine, CurveProjective, Engine, Field};
use rand::{ChaChaRng, OsRng, Rng, SeedableRng}; use rand::{ChaChaRng, OsRng, Rng, SeedableRng};
use ring::digest; use ring::digest;
use self::error::{ErrorKind, Result}; use self::error::{ErrorKind, Result};
use self::into_fr::IntoFr;
use self::poly::{Commitment, Poly}; use self::poly::{Commitment, Poly};
use fmt::HexBytes; use fmt::HexBytes;
@ -325,26 +327,26 @@ impl PublicKeySet {
} }
/// Returns the `i`-th public key share. /// Returns the `i`-th public key share.
pub fn public_key_share<T: Into<FrRepr>>(&self, i: T) -> PublicKeyShare { pub fn public_key_share<T: IntoFr>(&self, i: T) -> PublicKeyShare {
let value = self.commit.evaluate(from_repr_plus_1::<Fr>(i.into())); let value = self.commit.evaluate(into_fr_plus_1(i));
PublicKeyShare(PublicKey(value)) PublicKeyShare(PublicKey(value))
} }
/// Combines the shares into a signature that can be verified with the main public key. /// Combines the shares into a signature that can be verified with the main public key.
pub fn combine_signatures<'a, ITR, IND>(&self, shares: ITR) -> Result<Signature> pub fn combine_signatures<'a, T, I>(&self, shares: I) -> Result<Signature>
where where
ITR: IntoIterator<Item = (&'a IND, &'a SignatureShare)>, I: IntoIterator<Item = (T, &'a SignatureShare)>,
IND: Into<FrRepr> + Clone + 'a, T: IntoFr,
{ {
let samples = shares.into_iter().map(|(i, share)| (i, &(share.0).0)); let samples = shares.into_iter().map(|(i, share)| (i, &(share.0).0));
Ok(Signature(interpolate(self.commit.degree() + 1, samples)?)) Ok(Signature(interpolate(self.commit.degree() + 1, samples)?))
} }
/// Combines the shares to decrypt the ciphertext. /// Combines the shares to decrypt the ciphertext.
pub fn decrypt<'a, ITR, IND>(&self, shares: ITR, ct: &Ciphertext) -> Result<Vec<u8>> pub fn decrypt<'a, T, I>(&self, shares: I, ct: &Ciphertext) -> Result<Vec<u8>>
where where
ITR: IntoIterator<Item = (&'a IND, &'a DecryptionShare)>, I: IntoIterator<Item = (T, &'a DecryptionShare)>,
IND: Into<FrRepr> + Clone + 'a, T: IntoFr,
{ {
let samples = shares.into_iter().map(|(i, share)| (i, &share.0)); let samples = shares.into_iter().map(|(i, share)| (i, &share.0));
let g = interpolate(self.commit.degree() + 1, samples)?; let g = interpolate(self.commit.degree() + 1, samples)?;
@ -381,8 +383,8 @@ impl SecretKeySet {
} }
/// Returns the `i`-th secret key share. /// Returns the `i`-th secret key share.
pub fn secret_key_share<T: Into<FrRepr>>(&self, i: T) -> SecretKeyShare { pub fn secret_key_share<T: IntoFr>(&self, i: T) -> SecretKeyShare {
let value = self.poly.evaluate(from_repr_plus_1::<Fr>(i.into())); let value = self.poly.evaluate(into_fr_plus_1(i));
SecretKeyShare(SecretKey(value)) SecretKeyShare(SecretKey(value))
} }
@ -441,15 +443,15 @@ fn xor_vec(x: &[u8], y: &[u8]) -> Vec<u8> {
/// Given a list of `t` samples `(i - 1, f(i) * g)` for a polynomial `f` of degree `t - 1`, and a /// Given a list of `t` samples `(i - 1, f(i) * g)` for a polynomial `f` of degree `t - 1`, and a
/// group generator `g`, returns `f(0) * g`. /// group generator `g`, returns `f(0) * g`.
fn interpolate<'a, C, ITR, IND>(t: usize, items: ITR) -> Result<C> fn interpolate<'a, C, T, I>(t: usize, items: I) -> Result<C>
where where
C: CurveProjective, C: CurveProjective<Scalar = Fr>,
ITR: IntoIterator<Item = (&'a IND, &'a C)>, I: IntoIterator<Item = (T, &'a C)>,
IND: Into<<C::Scalar as PrimeField>::Repr> + Clone + 'a, T: IntoFr,
{ {
let samples: Vec<_> = items let samples: Vec<_> = items
.into_iter() .into_iter()
.map(|(i, sample)| (from_repr_plus_1::<C::Scalar>(i.clone().into()), sample)) .map(|(i, sample)| (into_fr_plus_1(i), sample))
.collect(); .collect();
if samples.len() < t { if samples.len() < t {
return Err(ErrorKind::NotEnoughShares.into()); return Err(ErrorKind::NotEnoughShares.into());
@ -475,10 +477,10 @@ where
Ok(result) Ok(result)
} }
fn from_repr_plus_1<F: PrimeField>(repr: F::Repr) -> F { fn into_fr_plus_1<I: IntoFr>(x: I) -> Fr {
let mut x = F::one(); let mut result = Fr::one();
x.add_assign(&F::from_repr(repr).expect("invalid index")); result.add_assign(&x.into_fr());
x result
} }
#[cfg(test)] #[cfg(test)]

@ -21,10 +21,12 @@ use std::hash::{Hash, Hasher};
use std::ptr::write_volatile; use std::ptr::write_volatile;
use std::{cmp, iter, ops}; use std::{cmp, iter, ops};
use pairing::bls12_381::{Fr, FrRepr, G1, G1Affine}; use pairing::bls12_381::{Fr, G1, G1Affine};
use pairing::{CurveAffine, CurveProjective, Field, PrimeField}; use pairing::{CurveAffine, CurveProjective, Field};
use rand::Rng; use rand::Rng;
use super::IntoFr;
/// A univariate polynomial in the prime field. /// A univariate polynomial in the prime field.
#[derive(Clone, Debug, Serialize, Deserialize, PartialEq, Eq)] #[derive(Clone, Debug, Serialize, Deserialize, PartialEq, Eq)]
pub struct Poly { pub struct Poly {
@ -176,15 +178,13 @@ impl Poly {
/// Returns the unique polynomial `f` of degree `samples.len() - 1` with the given values /// Returns the unique polynomial `f` of degree `samples.len() - 1` with the given values
/// `(x, f(x))`. /// `(x, f(x))`.
pub fn interpolate<'a, T, I>(samples_repr: I) -> Self pub fn interpolate<T, U, I>(samples_repr: I) -> Self
where where
I: IntoIterator<Item = (&'a T, &'a Fr)>, I: IntoIterator<Item = (T, U)>,
T: Into<FrRepr> + Clone + 'a, T: IntoFr,
U: IntoFr,
{ {
let convert = |(x_repr, y): (&T, &Fr)| { let convert = |(x, y): (T, U)| (x.into_fr(), y.into_fr());
let x = Fr::from_repr(x_repr.clone().into()).expect("invalid index");
(x, *y)
};
let samples: Vec<(Fr, Fr)> = samples_repr.into_iter().map(convert).collect(); let samples: Vec<(Fr, Fr)> = samples_repr.into_iter().map(convert).collect();
Self::compute_interpolation(&samples) Self::compute_interpolation(&samples)
} }
@ -195,12 +195,12 @@ impl Poly {
} }
/// Returns the value at the point `i`. /// Returns the value at the point `i`.
pub fn evaluate<T: Into<FrRepr>>(&self, i: T) -> Fr { pub fn evaluate<T: IntoFr>(&self, i: T) -> Fr {
let mut result = match self.coeff.last() { let mut result = match self.coeff.last() {
None => return Fr::zero(), None => return Fr::zero(),
Some(c) => *c, Some(c) => *c,
}; };
let x = Fr::from_repr(i.into()).expect("invalid index"); let x = i.into_fr();
for c in self.coeff.iter().rev().skip(1) { for c in self.coeff.iter().rev().skip(1) {
result.mul_assign(&x); result.mul_assign(&x);
result.add_assign(c); result.add_assign(c);
@ -306,12 +306,12 @@ impl Commitment {
} }
/// Returns the `i`-th public key share. /// Returns the `i`-th public key share.
pub fn evaluate<T: Into<FrRepr>>(&self, i: T) -> G1 { pub fn evaluate<T: IntoFr>(&self, i: T) -> G1 {
let mut result = match self.coeff.last() { let mut result = match self.coeff.last() {
None => return G1::zero(), None => return G1::zero(),
Some(c) => *c, Some(c) => *c,
}; };
let x = Fr::from_repr(i.into()).expect("invalid index"); let x = i.into_fr();
for c in self.coeff.iter().rev().skip(1) { for c in self.coeff.iter().rev().skip(1) {
result.mul_assign(x); result.mul_assign(x);
result.add_assign(c); result.add_assign(c);
@ -367,7 +367,7 @@ impl BivarPoly {
} }
/// Returns the polynomial's value at the point `(x, y)`. /// Returns the polynomial's value at the point `(x, y)`.
pub fn evaluate<T: Into<FrRepr>>(&self, x: T, y: T) -> Fr { pub fn evaluate<T: IntoFr>(&self, x: T, y: T) -> Fr {
let x_pow = self.powers(x); let x_pow = self.powers(x);
let y_pow = self.powers(y); let y_pow = self.powers(y);
// TODO: Can we save a few multiplication steps here due to the symmetry? // TODO: Can we save a few multiplication steps here due to the symmetry?
@ -384,7 +384,7 @@ impl BivarPoly {
} }
/// Returns the `x`-th row, as a univariate polynomial. /// Returns the `x`-th row, as a univariate polynomial.
pub fn row<T: Into<FrRepr>>(&self, x: T) -> Poly { pub fn row<T: IntoFr>(&self, x: T) -> Poly {
let x_pow = self.powers(x); let x_pow = self.powers(x);
let coeff: Vec<Fr> = (0..=self.degree) let coeff: Vec<Fr> = (0..=self.degree)
.map(|i| { .map(|i| {
@ -410,8 +410,8 @@ impl BivarPoly {
} }
/// Returns the `0`-th to `degree`-th power of `x`. /// Returns the `0`-th to `degree`-th power of `x`.
fn powers<T: Into<FrRepr>>(&self, x_repr: T) -> Vec<Fr> { fn powers<T: IntoFr>(&self, x: T) -> Vec<Fr> {
powers(x_repr, self.degree) powers(x, self.degree)
} }
} }
@ -441,7 +441,7 @@ impl BivarCommitment {
} }
/// Returns the commitment's value at the point `(x, y)`. /// Returns the commitment's value at the point `(x, y)`.
pub fn evaluate<T: Into<FrRepr>>(&self, x: T, y: T) -> G1 { pub fn evaluate<T: IntoFr>(&self, x: T, y: T) -> G1 {
let x_pow = self.powers(x); let x_pow = self.powers(x);
let y_pow = self.powers(y); let y_pow = self.powers(y);
// TODO: Can we save a few multiplication steps here due to the symmetry? // TODO: Can we save a few multiplication steps here due to the symmetry?
@ -458,7 +458,7 @@ impl BivarCommitment {
} }
/// Returns the `x`-th row, as a commitment to a univariate polynomial. /// Returns the `x`-th row, as a commitment to a univariate polynomial.
pub fn row<T: Into<FrRepr>>(&self, x: T) -> Commitment { pub fn row<T: IntoFr>(&self, x: T) -> Commitment {
let x_pow = self.powers(x); let x_pow = self.powers(x);
let coeff: Vec<G1> = (0..=self.degree) let coeff: Vec<G1> = (0..=self.degree)
.map(|i| { .map(|i| {
@ -475,18 +475,18 @@ impl BivarCommitment {
} }
/// Returns the `0`-th to `degree`-th power of `x`. /// Returns the `0`-th to `degree`-th power of `x`.
fn powers<T: Into<FrRepr>>(&self, x_repr: T) -> Vec<Fr> { fn powers<T: IntoFr>(&self, x: T) -> Vec<Fr> {
powers(x_repr, self.degree) powers(x, self.degree)
} }
} }
/// Returns the `0`-th to `degree`-th power of `x`. /// Returns the `0`-th to `degree`-th power of `x`.
fn powers<P: PrimeField, T: Into<P::Repr>>(x_repr: T, degree: usize) -> Vec<P> { fn powers<T: IntoFr>(into_x: T, degree: usize) -> Vec<Fr> {
let x = &P::from_repr(x_repr.into()).expect("invalid index"); let x = into_x.into_fr();
let mut x_pow_i = P::one(); let mut x_pow_i = Fr::one();
iter::once(x_pow_i) iter::once(x_pow_i)
.chain((0..degree).map(|_| { .chain((0..degree).map(|_| {
x_pow_i.mul_assign(x); x_pow_i.mul_assign(&x);
x_pow_i x_pow_i
})) }))
.collect() .collect()
@ -507,18 +507,14 @@ fn coeff_pos(i: usize, j: usize) -> usize {
mod tests { mod tests {
use std::collections::BTreeMap; use std::collections::BTreeMap;
use super::{coeff_pos, BivarPoly, Poly}; use super::{coeff_pos, BivarPoly, IntoFr, Poly};
use pairing::bls12_381::{Fr, G1Affine}; use pairing::bls12_381::{Fr, G1Affine};
use pairing::{CurveAffine, Field, PrimeField}; use pairing::{CurveAffine, Field};
use rand; use rand;
fn fr(x: i64) -> Fr { fn fr(x: i64) -> Fr {
let mut result = Fr::from_repr((x.abs() as u64).into()).unwrap(); x.into_fr()
if x < 0 {
result.negate();
}
result
} }
#[test] #[test]
@ -552,8 +548,7 @@ mod tests {
for &(x, y) in &samples { for &(x, y) in &samples {
assert_eq!(y, poly.evaluate(x)); assert_eq!(y, poly.evaluate(x));
} }
let sample_iter = samples.iter().map(|&(ref x, ref y)| (x, y)); assert_eq!(Poly::interpolate(samples), poly);
assert_eq!(Poly::interpolate(sample_iter), poly);
} }
#[test] #[test]
@ -579,16 +574,16 @@ mod tests {
for (bi_poly, bi_commit) in bi_polys.iter().zip(&pub_bi_commits) { for (bi_poly, bi_commit) in bi_polys.iter().zip(&pub_bi_commits) {
for m in 1..=node_num { for m in 1..=node_num {
// Node `m` receives its row and verifies it. // Node `m` receives its row and verifies it.
let row_poly = bi_poly.row(m as u64); let row_poly = bi_poly.row(m);
let row_commit = bi_commit.row(m as u64); let row_commit = bi_commit.row(m);
assert_eq!(row_poly.commitment(), row_commit); assert_eq!(row_poly.commitment(), row_commit);
// Node `s` receives the `s`-th value and verifies it. // Node `s` receives the `s`-th value and verifies it.
for s in 1..=node_num { for s in 1..=node_num {
let val = row_poly.evaluate(s as u64); let val = row_poly.evaluate(s);
let val_g1 = G1Affine::one().mul(val); let val_g1 = G1Affine::one().mul(val);
assert_eq!(bi_commit.evaluate(m as u64, s as u64), val_g1); assert_eq!(bi_commit.evaluate(m, s), val_g1);
// The node can't verify this directly, but it should have the correct value: // The node can't verify this directly, but it should have the correct value:
assert_eq!(bi_poly.evaluate(m as u64, s as u64), val); assert_eq!(bi_poly.evaluate(m, s), val);
} }
// A cheating dealer who modified the polynomial would be detected. // A cheating dealer who modified the polynomial would be detected.
@ -604,10 +599,10 @@ mod tests {
// `m` received three correct entries from that row: // `m` received three correct entries from that row:
let received: BTreeMap<_, _> = [1, 2, 4] let received: BTreeMap<_, _> = [1, 2, 4]
.iter() .iter()
.map(|&i| (i, bi_poly.evaluate(m as u64, i as u64))) .map(|&i| (i, bi_poly.evaluate(m, i)))
.collect(); .collect();
let my_row = Poly::interpolate(&received); let my_row = Poly::interpolate(received);
assert_eq!(bi_poly.evaluate(m as u64, 0), my_row.evaluate(0)); assert_eq!(bi_poly.evaluate(m, 0), my_row.evaluate(0));
assert_eq!(row_poly, my_row); assert_eq!(row_poly, my_row);
// The node sums up all values number `0` it received from the different dealer. No // The node sums up all values number `0` it received from the different dealer. No
@ -626,7 +621,7 @@ mod tests {
sec_key_set += bi_poly.row(0); sec_key_set += bi_poly.row(0);
} }
for m in 1..=node_num { for m in 1..=node_num {
assert_eq!(sec_key_set.evaluate(m as u64), sec_keys[m - 1]); assert_eq!(sec_key_set.evaluate(m), sec_keys[m - 1]);
} }
// The sum of the first rows of the public commitments is the commitment to the secret key // The sum of the first rows of the public commitments is the commitment to the secret key

Loading…
Cancel
Save