elastic_elgamal/sharing/mod.rs
1//! [Feldman's verifiable secret sharing][feldman-vss] (VSS) for ElGamal encryption.
2//!
3//! Feldman's VSS is an extension of [Shamir's secret sharing][sss] that provides a degree
4//! of verifiability for the scheme participants and the public. As with other VSS schemes,
5//! the goal is to securely distribute a secret among `n` participants so that the secret can
6//! be recombined by any `t` (but not less) of these participants. Unlike distributed key
7//! generation (DKG), VSS assumes a central authority (a *dealer*) generating the secret
8//! and distributing its shares among participants.
9//!
10//! # Construction
11//!
12//! **Inputs:**
13//!
14//! - Total number of participants `n`
15//! - Minimum number of participants necessary to restore secret `t`
16//! - Prime-order group with discrete log assumption with generator `G`
17//!
18//! **Assumptions:**
19//!
20//! - There is a secure broadcast among participants, which acts as a single source of truth
21//! (e.g., a blockchain). The broadcast is synchronous w.r.t. the protocol steps (in practice,
22//! this means that protocol steps take sufficiently long amount of time).
23//! - Secure synchronous P2P channels can be established between the dealer and participants.
24//! - The adversary is static (corrupts parties before protocol instantiation) and can corrupt
25//! less than a half of participants (including the dealer).
26//!
27//! Feldman's VSS proceeds as follows:
28//!
29//! 1. The dealer generates a secret `x` (a scalar in a group with discrete log assumption).
30//! Along with this scalar, the dealer generates `t` other scalars that are also kept secret.
31//! These scalars form a secret polynomial of degree `t`: `P(z) = x + x_1 * z + x_2 * z^2 + …`.
32//! 2. The dealer publishes coefficients `[x]G`, `[x_1]G`, ..., `[x_t]G` of the *public polynomial*
33//! corresponding to `P`: `Q(z) = [x]G + [z][x_1]G + [z^2][x_2]G + …`. Here, `[x]G` is the shared
34//! public key, and values `Q(i)` at `i = 1..=n` are public key shares of participants.
35//! 3. The dealer distributes secret key shares `s_i = P(i)` among participants `i = 1..=n`
36//! via secure P2P channels. Each participant can verify share validity by calculating
37//! `[s_i]G ?= Q(i)`.
38//!
39//! If a participant receives an incorrect secret share, the participant broadcasts a *complaint*
40//! against the dealer. The dealer responds by broadcasting the participant's share. Either the
41//! share is correct (in which case the complaining participant is at fault), or it is not
42//! (in which case the dealer is at fault).
43//!
44//! To improve auditability, the implemented version of VSS provides zero-knowledge proofs
45//! of possession both for the dealer and participants. The dealer must broadcast the public
46//! polynomial together with the proof; participants should broadcast proof of knowledge of
47//! a secret share once they receive the share from the dealer.
48//!
49//! # Distributed key generation
50//!
51//! Distributed key generation (DKG) differs from the approach implemented in this module
52//! in that there is no centralized dealer trusted by all participants. Instead, the participants
53//! essentially run parallel secret sharing protocol instances where each participant
54//! is a dealer in one of the instances. This approach is implemented
55//! in the [`dkg`](crate::dkg) module of this crate. Beware that it may not protect
56//! from participants biasing the distribution of the shared public key, e.g. by aborting
57//! the protocol; see [Gennaro et al.] for more details.
58//!
59//! [sss]: https://en.wikipedia.org/wiki/Shamir%27s_Secret_Sharing
60//! [feldman-vss]: https://www.cs.umd.edu/~gasarch/TOPICS/secretsharing/feldmanVSS.pdf
61//! [Gennaro et al.]: https://link.springer.com/content/pdf/10.1007/3-540-48910-X_21.pdf
62//!
63//! # Examples
64//!
65//! Threshold encryption scheme requiring 2 of 3 participants.
66//!
67//! ```
68//! # use elastic_elgamal::{
69//! # group::Ristretto, sharing::*, CandidateDecryption, Ciphertext, DiscreteLogTable,
70//! # };
71//! # use std::error::Error as StdError;
72//! # fn main() -> Result<(), Box<dyn StdError>> {
73//! let mut rng = rand::rng();
74//! let params = Params::new(3, 2);
75//!
76//! // Initialize the dealer.
77//! let dealer = Dealer::<Ristretto>::new(params, &mut rng);
78//! let (public_poly, poly_proof) = dealer.public_info();
79//! let key_set = PublicKeySet::new(params, public_poly, poly_proof)?;
80//!
81//! // Initialize participants based on secret shares provided by the dealer.
82//! let participants = (0..3)
83//! .map(|i| ActiveParticipant::new(
84//! key_set.clone(),
85//! i,
86//! dealer.secret_share_for_participant(i),
87//! ))
88//! .collect::<Result<Vec<_>, _>>()?;
89//!
90//! // At last, participants can decrypt messages!
91//! let encrypted_value = 5_u64;
92//! let enc = key_set.shared_key().encrypt(encrypted_value, &mut rng);
93//! let shares_with_proofs = participants
94//! .iter()
95//! .map(|p| p.decrypt_share(enc, &mut rng))
96//! .take(2); // emulate the 3rd participant dropping off
97//!
98//! // Emulate share transfer via untrusted network.
99//! let dec_shares = shares_with_proofs
100//! .enumerate()
101//! .map(|(i, (share, proof))| {
102//! let share = CandidateDecryption::from_bytes(&share.to_bytes()).unwrap();
103//! key_set.verify_share(share, enc, i, &proof).unwrap()
104//! });
105//!
106//! // Combine decryption shares.
107//! let combined = params.combine_shares(dec_shares.enumerate()).unwrap();
108//! // Use a lookup table to decrypt back to scalar.
109//! let lookup_table = DiscreteLogTable::<Ristretto>::new(0..10);
110//! let dec = combined.decrypt(enc, &lookup_table);
111//! assert_eq!(dec, Some(encrypted_value));
112//! # Ok(())
113//! # }
114//! ```
115
116use core::{cmp::Ordering, fmt, ops};
117
118#[cfg(feature = "serde")]
119use serde::{Deserialize, Serialize};
120
121#[cfg(feature = "serde")]
122use crate::serde::{ElementHelper, VecHelper};
123use crate::{VerifiableDecryption, alloc::Vec, group::Group, proofs::VerificationError};
124
125mod key_set;
126mod participant;
127
128pub use self::{
129 key_set::PublicKeySet,
130 participant::{ActiveParticipant, Dealer},
131};
132
133/// Computes multipliers for the Lagrange polynomial interpolation based on the function value
134/// at the given points. The indexes are zero-based, hence points are determined as
135/// `indexes.iter().map(|&i| i + 1)`.
136///
137/// The returned scalars need to be additionally scaled by the common multiplier, equal
138/// to the product of all points, which is returned as the second value.
139fn lagrange_coefficients<G: Group>(indexes: &[usize]) -> (Vec<G::Scalar>, G::Scalar) {
140 // `false` corresponds to positive sign, `true` to negative. This is in order
141 // to make XOR work as expected.
142
143 let mut denominators: Vec<_> = indexes
144 .iter()
145 .map(|&index| {
146 let (sign, denominator) = indexes
147 .iter()
148 .map(|&other_index| match index.cmp(&other_index) {
149 Ordering::Greater => (true, G::Scalar::from((index - other_index) as u64)),
150 Ordering::Less => (false, G::Scalar::from((other_index - index) as u64)),
151 Ordering::Equal => (false, G::Scalar::from(index as u64 + 1)),
152 })
153 .fold(
154 (false, G::Scalar::from(1)),
155 |(sign, magnitude), (elem_sign, elem_magnitude)| {
156 (sign ^ elem_sign, magnitude * elem_magnitude)
157 },
158 );
159
160 if sign { -denominator } else { denominator }
161 })
162 .collect();
163 G::invert_scalars(&mut denominators);
164
165 let scale = indexes
166 .iter()
167 .map(|&index| G::Scalar::from(index as u64 + 1))
168 .fold(G::Scalar::from(1), |acc, value| acc * value);
169 (denominators, scale)
170}
171
172/// Structure representing public polynomial consisting of group elements.
173#[derive(Debug, Clone)]
174#[cfg_attr(feature = "serde", derive(Serialize, Deserialize))]
175#[cfg_attr(feature = "serde", serde(transparent, bound = ""))]
176pub(crate) struct PublicPolynomial<G: Group>(
177 #[cfg_attr(feature = "serde", serde(with = "VecHelper::<ElementHelper<G>, 1>"))]
178 Vec<G::Element>,
179);
180
181impl<G: Group> PublicPolynomial<G> {
182 pub(crate) fn new(values: Vec<G::Element>) -> Self {
183 Self(values)
184 }
185
186 fn value_at_zero(&self) -> G::Element {
187 self.0[0]
188 }
189
190 /// Computes value of this public polynomial at the specified point in variable time.
191 pub(crate) fn value_at(&self, x: G::Scalar) -> G::Element {
192 let mut val = G::Scalar::from(1_u64);
193 let scalars: Vec<_> = (0..self.0.len())
194 .map(|_| {
195 let output = val;
196 val = val * x;
197 output
198 })
199 .collect();
200
201 G::vartime_multi_mul(&scalars, self.0.iter().copied())
202 }
203}
204
205impl<G: Group> ops::AddAssign<&Self> for PublicPolynomial<G> {
206 fn add_assign(&mut self, rhs: &Self) {
207 debug_assert_eq!(
208 self.0.len(),
209 rhs.0.len(),
210 "cannot add polynomials of different degrees"
211 );
212
213 for (val, &rhs_val) in self.0.iter_mut().zip(&rhs.0) {
214 *val = *val + rhs_val;
215 }
216 }
217}
218
219/// Errors that can occur during the secret sharing protocol.
220#[derive(Debug)]
221#[non_exhaustive]
222pub enum Error {
223 /// Public polynomial received from the dealer is malformed.
224 MalformedDealerPolynomial,
225 /// Proof of possession supplied with the dealer's public polynomial is invalid.
226 InvalidDealerProof(VerificationError),
227 /// Secret received from the dealer does not correspond to their commitment via
228 /// the public polynomial.
229 InvalidSecret,
230 /// Number of participants specified in [`Params`] does not match the number
231 /// of provided public keys.
232 ParticipantCountMismatch,
233 /// Participants' public keys do not correspond to a single shared key.
234 MalformedParticipantKeys,
235}
236
237impl fmt::Display for Error {
238 fn fmt(&self, formatter: &mut fmt::Formatter<'_>) -> fmt::Result {
239 match self {
240 Self::MalformedDealerPolynomial => {
241 formatter.write_str("public polynomial received from the dealer is malformed")
242 }
243 Self::InvalidDealerProof(err) => write!(
244 formatter,
245 "proof of possession supplied with the dealer's public polynomial \
246 is invalid: {err}"
247 ),
248 Self::InvalidSecret => formatter.write_str(
249 "secret received from the dealer does not correspond to their commitment via \
250 public polynomial",
251 ),
252 Self::ParticipantCountMismatch => formatter.write_str(
253 "number of participants specified in `Params` does not match the number \
254 of provided public keys",
255 ),
256 Self::MalformedParticipantKeys => formatter
257 .write_str("participants' public keys do not correspond to a single shared key"),
258 }
259 }
260}
261
262#[cfg(feature = "std")]
263impl std::error::Error for Error {
264 fn source(&self) -> Option<&(dyn std::error::Error + 'static)> {
265 match self {
266 Self::InvalidDealerProof(err) => Some(err),
267 _ => None,
268 }
269 }
270}
271
272/// Parameters of a threshold ElGamal encryption scheme.
273#[derive(Debug, Clone, Copy, PartialEq, Eq)]
274#[cfg_attr(feature = "serde", derive(Serialize, Deserialize))]
275pub struct Params {
276 /// Total number of shares / participants.
277 pub shares: usize,
278 /// Number of participants necessary to jointly restore the secret.
279 pub threshold: usize,
280}
281
282impl Params {
283 /// Creates new parameters.
284 ///
285 /// # Panics
286 ///
287 /// Panics if `shares` is equal to zero or if `threshold` is not in `1..=shares`.
288 pub const fn new(shares: usize, threshold: usize) -> Self {
289 assert!(shares > 0);
290 assert!(threshold > 0 && threshold <= shares);
291 Self { shares, threshold }
292 }
293
294 /// Combines shares decrypting the specified `ciphertext`. The shares must be provided
295 /// together with the 0-based indexes of the participants they are coming from.
296 ///
297 /// Returns the combined decryption, or `None` if the number of shares is insufficient.
298 ///
299 /// # Panics
300 ///
301 /// Panics if any index in `shares` exceeds the maximum participant's index as per `params`.
302 pub fn combine_shares<G: Group>(
303 self,
304 shares: impl IntoIterator<Item = (usize, VerifiableDecryption<G>)>,
305 ) -> Option<VerifiableDecryption<G>> {
306 let (indexes, shares): (Vec<_>, Vec<_>) = shares
307 .into_iter()
308 .take(self.threshold)
309 .map(|(index, share)| (index, *share.as_element()))
310 .unzip();
311 if shares.len() < self.threshold {
312 return None;
313 }
314 assert!(
315 indexes.iter().all(|&index| index < self.shares),
316 "Invalid share indexes {:?}; expected values in 0..{}",
317 indexes.iter().copied(),
318 self.shares
319 );
320
321 let (denominators, scale) = lagrange_coefficients::<G>(&indexes);
322 let restored_value = G::vartime_multi_mul(&denominators, shares);
323 let dh_element = restored_value * &scale;
324 Some(VerifiableDecryption::from_element(dh_element))
325 }
326}
327
328#[cfg(test)]
329mod tests {
330 use super::*;
331 use crate::{curve25519::scalar::Scalar as Scalar25519, group::Ristretto};
332
333 #[test]
334 fn lagrange_coeffs_are_computed_correctly() {
335 // d_0 = 2 / (2 - 1) = 2
336 // d_1 = 1 / (1 - 2) = -1
337 let (coeffs, scale) = lagrange_coefficients::<Ristretto>(&[0, 1]);
338 assert_eq!(
339 coeffs,
340 [Scalar25519::from(1_u32), -Scalar25519::from(2_u32).invert()]
341 );
342 assert_eq!(scale, Scalar25519::from(2_u32));
343
344 // d_0 = 3 / (3 - 1) = 3/2
345 // d_1 = 1 / (1 - 3) = -1/2
346 let (coeffs, scale) = lagrange_coefficients::<Ristretto>(&[0, 2]);
347 assert_eq!(
348 coeffs,
349 [
350 Scalar25519::from(2_u32).invert(),
351 -Scalar25519::from(6_u32).invert(),
352 ]
353 );
354 assert_eq!(scale, Scalar25519::from(3_u32));
355
356 // d_0 = 4 * 5 / (4 - 1) * (5 - 1) = 20/12 = 5/3
357 // d_1 = 1 * 5 / (1 - 4) * (5 - 4) = -5/3
358 // d_2 = 1 * 4 / (1 - 5) * (4 - 5) = 4/4 = 1
359 let (coeffs, scale) = lagrange_coefficients::<Ristretto>(&[0, 3, 4]);
360 assert_eq!(
361 coeffs,
362 [
363 Scalar25519::from(12_u32).invert(),
364 -Scalar25519::from(12_u32).invert(),
365 Scalar25519::from(20_u32).invert(),
366 ]
367 );
368 assert_eq!(scale, Scalar25519::from(20_u32));
369 }
370}