elastic_elgamal/proofs/
ring.rs

1//! Ring proofs.
2
3use core::{fmt, mem};
4
5use elliptic_curve::rand_core::{CryptoRng, RngCore};
6use merlin::Transcript;
7#[cfg(feature = "serde")]
8use serde::{Deserialize, Serialize};
9
10#[cfg(feature = "serde")]
11use crate::serde::{ScalarHelper, VecHelper};
12use crate::{
13    Ciphertext, PublicKey, SecretKey,
14    alloc::{Vec, vec},
15    encryption::ExtendedCiphertext,
16    group::Group,
17    proofs::{TranscriptForGroup, VerificationError},
18};
19
20/// An incomplete ring proving that the encrypted value is in the a priori known set of
21/// admissible values.
22struct Ring<'a, G: Group> {
23    // Public parameters of the ring.
24    index: usize,
25    admissible_values: &'a [G::Element],
26    ciphertext: Ciphertext<G>,
27
28    // ZKP-related public values.
29    transcript: Transcript,
30    responses: &'a mut [G::Scalar],
31    terminal_commitments: (G::Element, G::Element),
32
33    // Secret values.
34    value_index: usize,
35    discrete_log: SecretKey<G>,
36    random_scalar: SecretKey<G>,
37}
38
39impl<G: Group> fmt::Debug for Ring<'_, G> {
40    fn fmt(&self, formatter: &mut fmt::Formatter<'_>) -> fmt::Result {
41        formatter
42            .debug_struct("Ring")
43            .field("index", &self.index)
44            .field("admissible_values", &self.admissible_values)
45            .field("ciphertext", &self.ciphertext)
46            .field("responses", &self.responses)
47            .field("terminal_commitments", &self.terminal_commitments)
48            .finish()
49    }
50}
51
52impl<'a, G: Group> Ring<'a, G> {
53    #[allow(clippy::too_many_arguments)] // fine for a private function
54    fn new<R: CryptoRng + RngCore>(
55        index: usize,
56        log_base: G::Element,
57        ciphertext: ExtendedCiphertext<G>,
58        admissible_values: &'a [G::Element],
59        value_index: usize,
60        transcript: &Transcript,
61        responses: &'a mut [G::Scalar],
62        rng: &mut R,
63    ) -> Self {
64        assert!(
65            !admissible_values.is_empty(),
66            "No admissible values supplied"
67        );
68        assert!(
69            value_index < admissible_values.len(),
70            "Specified value index is out of bounds"
71        );
72        debug_assert_eq!(
73            responses.len(),
74            admissible_values.len(),
75            "Number of responses doesn't match number of admissible values"
76        );
77
78        let random_element = ciphertext.inner.random_element;
79        let blinded_value = ciphertext.inner.blinded_element;
80        debug_assert!(
81            {
82                let expected_blinded_value = log_base * ciphertext.random_scalar.expose_scalar()
83                    + admissible_values[value_index];
84                expected_blinded_value == blinded_value
85            },
86            "Specified ciphertext does not match the specified `value_index`"
87        );
88
89        let mut transcript = transcript.clone();
90        transcript.start_proof(b"ring_enc");
91        transcript.append_message(b"enc", &ciphertext.inner.to_bytes());
92        // NB: we don't add `admissible_values` to the transcript since we assume that
93        // they are fixed in the higher-level protocol.
94        transcript.append_u64(b"i", index as u64);
95
96        // Choose a random scalar to use in the equation matching the known discrete log.
97        let random_scalar = SecretKey::<G>::generate(rng);
98        let mut commitments = (
99            G::mul_generator(random_scalar.expose_scalar()),
100            log_base * random_scalar.expose_scalar(),
101        );
102
103        let it = admissible_values.iter().enumerate().skip(value_index + 1);
104        for (eq_index, &admissible_value) in it {
105            let mut eq_transcript = transcript.clone();
106            eq_transcript.append_u64(b"j", eq_index as u64 - 1);
107            eq_transcript.append_element::<G>(b"R_G", &commitments.0);
108            eq_transcript.append_element::<G>(b"R_K", &commitments.1);
109            let challenge = eq_transcript.challenge_scalar::<G>(b"c");
110
111            let response = G::generate_scalar(rng);
112            responses[eq_index] = response;
113            let dh_element = blinded_value - admissible_value;
114            commitments = (
115                G::mul_generator(&response) - random_element * &challenge,
116                G::multi_mul([&response, &-challenge], [log_base, dh_element]),
117            );
118        }
119
120        Self {
121            index,
122            value_index,
123            admissible_values,
124            ciphertext: ciphertext.inner,
125            transcript,
126            responses,
127            terminal_commitments: commitments,
128            discrete_log: ciphertext.random_scalar,
129            random_scalar,
130        }
131    }
132
133    /// Completes the ring by calculating the common challenge and closing all rings using it.
134    ///
135    /// # Return value
136    ///
137    /// Returns the common challenge.
138    fn aggregate<R: CryptoRng + RngCore>(
139        rings: Vec<Self>,
140        log_base: G::Element,
141        transcript: &mut Transcript,
142        rng: &mut R,
143    ) -> G::Scalar {
144        debug_assert!(
145            rings.iter().enumerate().all(|(i, ring)| i == ring.index),
146            "Rings have bogus indexes"
147        );
148
149        for ring in &rings {
150            let commitments = &ring.terminal_commitments;
151            transcript.append_element::<G>(b"R_G", &commitments.0);
152            transcript.append_element::<G>(b"R_K", &commitments.1);
153        }
154
155        let common_challenge = transcript.challenge_scalar::<G>(b"c");
156        for ring in rings {
157            ring.finalize(log_base, common_challenge, rng);
158        }
159        common_challenge
160    }
161
162    fn finalize<R: CryptoRng + RngCore>(
163        self,
164        log_base: G::Element,
165        common_challenge: G::Scalar,
166        rng: &mut R,
167    ) {
168        // Compute remaining responses for non-reversible equations.
169        let mut challenge = common_challenge;
170        let it = self.admissible_values[..self.value_index]
171            .iter()
172            .enumerate();
173        for (eq_index, &admissible_value) in it {
174            let response = G::generate_scalar(rng);
175            self.responses[eq_index] = response;
176            let dh_element = self.ciphertext.blinded_element - admissible_value;
177            let commitments = (
178                G::mul_generator(&response) - self.ciphertext.random_element * &challenge,
179                G::multi_mul([&response, &-challenge], [log_base, dh_element]),
180            );
181
182            let mut eq_transcript = self.transcript.clone();
183            eq_transcript.append_u64(b"j", eq_index as u64);
184            eq_transcript.append_element::<G>(b"R_G", &commitments.0);
185            eq_transcript.append_element::<G>(b"R_K", &commitments.1);
186            challenge = eq_transcript.challenge_scalar::<G>(b"c");
187        }
188
189        // Finally, compute the response for equation #`value_index`, using our knowledge
190        // of the trapdoor.
191        debug_assert_eq!(self.responses[self.value_index], G::Scalar::from(0_u64));
192        self.responses[self.value_index] =
193            challenge * self.discrete_log.expose_scalar() + self.random_scalar.expose_scalar();
194    }
195}
196
197/// Zero-knowledge proof that the one or more encrypted values is each in the a priori known set of
198/// admissible values. (Admissible values may differ among encrypted values.)
199///
200/// # Construction
201///
202/// In short, a proof is constructed almost identically to [Borromean ring signatures] by
203/// Maxwell and Poelstra, with the only major difference being that we work on ElGamal ciphertexts
204/// instead of group elements (= public keys).
205///
206/// A proof consists of one or more *rings*. Each ring proves than a certain
207/// ElGamal ciphertext `E = (R, B)` for public key `K` in a group with generator `G`
208/// encrypts one of distinct admissible values `x_0`, `x_1`, ..., `x_n`.
209/// `K` and `G` are shared among rings, admissible values are generally not.
210/// Different rings may have different number of admissible values.
211///
212/// ## Single ring
213///
214/// A ring is a challenge `e_0` and a set of responses `s_0`, `s_1`, ..., `s_n`, which
215/// must satisfy the following verification procedure:
216///
217/// For each `j` in `0..=n`, compute
218///
219/// ```text
220/// R_G(j) = [s_j]G - [e_j]R;
221/// R_K(j) = [s_j]K - [e_j](B - [x_j]G);
222/// e_{j+1} = H(j, R_G(j), R_K(j));
223/// ```
224///
225/// Here, `H` is a cryptographic hash function. The ring is valid if `e_0 = e_{n+1}`.
226///
227/// This construction is almost identical to [Abe–Ohkubo–Suzuki ring signatures][ring],
228/// with the only difference that two group elements are hashed on each iteration instead of one.
229/// If admissible values consist of a single value, this protocol reduces to
230/// [`LogEqualityProof`] / Chaum–Pedersen protocol.
231///
232/// As with "ordinary" ring signatures, constructing a ring is only feasible when knowing
233/// additional *trapdoor information*. Namely, the prover must know
234///
235/// ```text
236/// r = dlog_G(R) = dlog_K(B - [x_j]G)
237/// ```
238///
239/// for a certain `j`. (This discrete log `r` is the random scalar used in ElGamal encryption.)
240/// With this info, the prover constructs the ring as follows:
241///
242/// 1. Select random scalar `x` and compute `R_G(j) = [x]G`, `R_K(j) = [x]K`.
243/// 2. Compute `e_{j+1}`, ... `e_n`, ..., `e_j` ("wrapping" around `e_0 = e_{n+1}`)
244///    as per verification formulas. `s_*` scalars are selected uniformly at random.
245/// 3. Compute `s_j` using the trapdoor information: `s_j = x + e_j * r`.
246///
247/// ## Multiple rings
248///
249/// Transformation to multiple rings is analogous to one in [Borromean ring signatures].
250/// Namely, challenge `e_0` is shared among all rings and is computed by hashing
251/// values of `R_G` and `R_K` with the maximum index for each of the rings.
252///
253/// # Applications
254///
255/// ## Voting protocols
256///
257/// [`EncryptedChoice`](crate::app::EncryptedChoice) uses `RingProof` to prove that all encrypted
258/// values are Boolean (0 or 1). Using a common challenge allows to reduce proof size by ~33%.
259///
260/// ## Range proofs
261///
262/// See [`RangeProof`](crate::RangeProof).
263///
264/// # Implementation details
265///
266/// - The proof is serialized as the common challenge `e_0` followed by `s_i` scalars for
267///   all the rings.
268/// - Standalone proof generation and verification are not exposed in public crate APIs.
269///   Rather, proofs are part of large protocols, such as [`PublicKey::encrypt_bool()`] /
270///   [`PublicKey::verify_bool()`].
271/// - The context of the proof is set using [`Transcript`] APIs, which provides hash functions
272///   in the protocol described above. Importantly, the proof itself commits to encrypted values
273///   and ring indexes, but not to the admissible values across the rings. This must be taken
274///   care of in a higher-level protocol, and this is the case for protocols exposed by the crate.
275///
276/// [`LogEqualityProof`]: crate::LogEqualityProof
277/// [Borromean ring signatures]: https://raw.githubusercontent.com/Blockstream/borromean_paper/master/borromean_draft_0.01_34241bb.pdf
278/// [ring]: https://link.springer.com/content/pdf/10.1007/3-540-36178-2_26.pdf
279#[derive(Debug, Clone)]
280#[cfg_attr(feature = "serde", derive(Serialize, Deserialize))]
281#[cfg_attr(feature = "serde", serde(bound = ""))]
282pub struct RingProof<G: Group> {
283    #[cfg_attr(feature = "serde", serde(with = "ScalarHelper::<G>"))]
284    common_challenge: G::Scalar,
285    #[cfg_attr(feature = "serde", serde(with = "VecHelper::<ScalarHelper<G>, 2>"))]
286    ring_responses: Vec<G::Scalar>,
287}
288
289impl<G: Group> RingProof<G> {
290    fn initialize_transcript(transcript: &mut Transcript, receiver: &PublicKey<G>) {
291        transcript.start_proof(b"multi_ring_enc");
292        transcript.append_element_bytes(b"K", receiver.as_bytes());
293    }
294
295    pub(crate) fn new(common_challenge: G::Scalar, ring_responses: Vec<G::Scalar>) -> Self {
296        Self {
297            common_challenge,
298            ring_responses,
299        }
300    }
301
302    pub(crate) fn verify<'a>(
303        &self,
304        receiver: &PublicKey<G>,
305        admissible_values: impl Iterator<Item = &'a [G::Element]> + Clone,
306        ciphertexts: impl Iterator<Item = Ciphertext<G>>,
307        transcript: &mut Transcript,
308    ) -> Result<(), VerificationError> {
309        // Do quick preliminary checks.
310        let total_rings_size: usize = admissible_values.clone().map(<[_]>::len).sum();
311        VerificationError::check_lengths(
312            "items in all rings",
313            self.total_rings_size(),
314            total_rings_size,
315        )?;
316
317        Self::initialize_transcript(transcript, receiver);
318        // We add common commitments to the `transcript` as we cycle through rings,
319        // so we need a separate transcript copy to initialize ring transcripts.
320        let initial_ring_transcript = transcript.clone();
321
322        let it = admissible_values.zip(ciphertexts).enumerate();
323        let mut starting_response = 0;
324        for (ring_index, (values, ciphertext)) in it {
325            let mut challenge = self.common_challenge;
326            let mut commitments = (G::generator(), G::generator());
327
328            let mut ring_transcript = initial_ring_transcript.clone();
329            ring_transcript.start_proof(b"ring_enc");
330            ring_transcript.append_message(b"enc", &ciphertext.to_bytes());
331            ring_transcript.append_u64(b"i", ring_index as u64);
332
333            for (eq_index, (&admissible_value, response)) in values
334                .iter()
335                .zip(&self.ring_responses[starting_response..])
336                .enumerate()
337            {
338                let dh_element = ciphertext.blinded_element - admissible_value;
339                let neg_challenge = -challenge;
340
341                commitments = (
342                    G::vartime_double_mul_generator(
343                        &neg_challenge,
344                        ciphertext.random_element,
345                        response,
346                    ),
347                    G::vartime_multi_mul(
348                        [response, &neg_challenge],
349                        [receiver.as_element(), dh_element],
350                    ),
351                );
352
353                // We can skip deriving the challenge for the last equation; it's not used anyway.
354                if eq_index + 1 < values.len() {
355                    let mut eq_transcript = ring_transcript.clone();
356                    eq_transcript.append_u64(b"j", eq_index as u64);
357                    eq_transcript.append_element::<G>(b"R_G", &commitments.0);
358                    eq_transcript.append_element::<G>(b"R_K", &commitments.1);
359                    challenge = eq_transcript.challenge_scalar::<G>(b"c");
360                }
361            }
362
363            starting_response += values.len();
364            transcript.append_element::<G>(b"R_G", &commitments.0);
365            transcript.append_element::<G>(b"R_K", &commitments.1);
366        }
367
368        let expected_challenge = transcript.challenge_scalar::<G>(b"c");
369        if expected_challenge == self.common_challenge {
370            Ok(())
371        } else {
372            Err(VerificationError::ChallengeMismatch)
373        }
374    }
375
376    pub(crate) fn total_rings_size(&self) -> usize {
377        self.ring_responses.len()
378    }
379
380    /// Serializes this proof into bytes. As described [above](#implementation-details),
381    /// the proof is serialized as the common challenge `e_0` followed by response scalars `s_*`
382    /// corresponding successively to each admissible value in each ring.
383    pub fn to_bytes(&self) -> Vec<u8> {
384        let mut bytes = vec![0_u8; G::SCALAR_SIZE * (1 + self.total_rings_size())];
385        G::serialize_scalar(&self.common_challenge, &mut bytes[..G::SCALAR_SIZE]);
386
387        let chunks = bytes[G::SCALAR_SIZE..].chunks_mut(G::SCALAR_SIZE);
388        for (response, buffer) in self.ring_responses.iter().zip(chunks) {
389            G::serialize_scalar(response, buffer);
390        }
391        bytes
392    }
393
394    /// Attempts to deserialize a proof from bytes. Returns `None` if `bytes` do not represent
395    /// a well-formed proof.
396    #[allow(clippy::missing_panics_doc)] // triggered by `debug_assert`
397    pub fn from_bytes(bytes: &[u8]) -> Option<Self> {
398        if bytes.len() % G::SCALAR_SIZE != 0 || bytes.len() < 3 * G::SCALAR_SIZE {
399            return None;
400        }
401        let common_challenge = G::deserialize_scalar(&bytes[..G::SCALAR_SIZE])?;
402
403        let ring_responses: Option<Vec<_>> = bytes[G::SCALAR_SIZE..]
404            .chunks(G::SCALAR_SIZE)
405            .map(G::deserialize_scalar)
406            .collect();
407        let ring_responses = ring_responses?;
408        debug_assert!(ring_responses.len() >= 2);
409
410        Some(Self {
411            common_challenge,
412            ring_responses,
413        })
414    }
415}
416
417/// **NB.** Separate method calls of the builder depend on the position of the encrypted values
418/// within admissible ones. This means that if a proof is constructed with interruptions between
419/// method calls, there is a chance for an adversary to perform a timing attack.
420#[doc(hidden)] // only public for benchmarking
421pub struct RingProofBuilder<'a, G: Group, R> {
422    receiver: &'a PublicKey<G>,
423    transcript: &'a mut Transcript,
424    rings: Vec<Ring<'a, G>>,
425    ring_responses: &'a mut [G::Scalar],
426    rng: &'a mut R,
427}
428
429impl<G: Group, R: fmt::Debug> fmt::Debug for RingProofBuilder<'_, G, R> {
430    fn fmt(&self, formatter: &mut fmt::Formatter<'_>) -> fmt::Result {
431        formatter
432            .debug_struct("RingProofBuilder")
433            .field("receiver", self.receiver)
434            .field("rings", &self.rings)
435            .field("rng", self.rng)
436            .finish()
437    }
438}
439
440impl<'a, G: Group, R: RngCore + CryptoRng> RingProofBuilder<'a, G, R> {
441    /// Starts building a [`RingProof`].
442    pub fn new(
443        receiver: &'a PublicKey<G>,
444        ring_count: usize,
445        ring_responses: &'a mut [G::Scalar],
446        transcript: &'a mut Transcript,
447        rng: &'a mut R,
448    ) -> Self {
449        RingProof::<G>::initialize_transcript(transcript, receiver);
450        Self {
451            receiver,
452            transcript,
453            rings: Vec::with_capacity(ring_count),
454            ring_responses,
455            rng,
456        }
457    }
458
459    /// Adds a value among `admissible_values` as a new ring to this proof.
460    pub fn add_value(
461        &mut self,
462        admissible_values: &'a [G::Element],
463        value_index: usize,
464    ) -> ExtendedCiphertext<G> {
465        let ext_ciphertext =
466            ExtendedCiphertext::new(admissible_values[value_index], self.receiver, self.rng);
467        self.add_precomputed_value(ext_ciphertext.clone(), admissible_values, value_index);
468        ext_ciphertext
469    }
470
471    pub(crate) fn add_precomputed_value(
472        &mut self,
473        ciphertext: ExtendedCiphertext<G>,
474        admissible_values: &'a [G::Element],
475        value_index: usize,
476    ) {
477        let ring_responses = mem::take(&mut self.ring_responses);
478        let (responses_for_ring, rest) = ring_responses.split_at_mut(admissible_values.len());
479        self.ring_responses = rest;
480
481        let ring = Ring::new(
482            self.rings.len(),
483            self.receiver.as_element(),
484            ciphertext,
485            admissible_values,
486            value_index,
487            &*self.transcript,
488            responses_for_ring,
489            self.rng,
490        );
491        self.rings.push(ring);
492    }
493
494    /// Finishes building all rings and returns a common challenge.
495    pub fn build(self) -> G::Scalar {
496        debug_assert!(
497            self.ring_responses.is_empty(),
498            "Not all ring_responses were used"
499        );
500        Ring::aggregate(
501            self.rings,
502            self.receiver.as_element(),
503            self.transcript,
504            self.rng,
505        )
506    }
507}
508
509#[cfg(test)]
510mod tests {
511    use core::iter;
512
513    use rand::Rng;
514    use test_casing::test_casing;
515
516    use super::*;
517    use crate::{
518        curve25519::{ristretto::RistrettoPoint, scalar::Scalar as Scalar25519, traits::Identity},
519        group::{ElementOps, Ristretto},
520    };
521
522    type Keypair = crate::Keypair<Ristretto>;
523
524    #[test]
525    fn single_ring_with_2_elements_works() {
526        let mut rng = rand::rng();
527        let keypair = Keypair::generate(&mut rng);
528        let log_base = keypair.public().as_element();
529        let admissible_values = [RistrettoPoint::identity(), Ristretto::generator()];
530
531        let value = RistrettoPoint::identity();
532        let ext_ciphertext = ExtendedCiphertext::new(value, keypair.public(), &mut rng);
533        let ciphertext = ext_ciphertext.inner;
534
535        let mut transcript = Transcript::new(b"test_ring_encryption");
536        RingProof::initialize_transcript(&mut transcript, keypair.public());
537
538        let mut ring_responses = vec![Scalar25519::default(); 2];
539        let signature_ring = Ring::new(
540            0,
541            log_base,
542            ext_ciphertext,
543            &admissible_values,
544            0,
545            &transcript,
546            &mut ring_responses,
547            &mut rng,
548        );
549        let common_challenge =
550            Ring::aggregate(vec![signature_ring], log_base, &mut transcript, &mut rng);
551
552        RingProof::new(common_challenge, ring_responses)
553            .verify(
554                keypair.public(),
555                iter::once(&admissible_values as &[_]),
556                iter::once(ciphertext),
557                &mut Transcript::new(b"test_ring_encryption"),
558            )
559            .unwrap();
560
561        // Check a proof for encryption of 1.
562        let value = Ristretto::generator();
563        let ext_ciphertext = ExtendedCiphertext::new(value, keypair.public(), &mut rng);
564        let ciphertext = ext_ciphertext.inner;
565
566        let mut transcript = Transcript::new(b"test_ring_encryption");
567        RingProof::initialize_transcript(&mut transcript, keypair.public());
568        let mut ring_responses = vec![Scalar25519::default(); 2];
569        let signature_ring = Ring::new(
570            0,
571            log_base,
572            ext_ciphertext,
573            &admissible_values,
574            1,
575            &transcript,
576            &mut ring_responses,
577            &mut rng,
578        );
579        let common_challenge =
580            Ring::aggregate(vec![signature_ring], log_base, &mut transcript, &mut rng);
581
582        RingProof::new(common_challenge, ring_responses)
583            .verify(
584                keypair.public(),
585                iter::once(&admissible_values as &[_]),
586                iter::once(ciphertext),
587                &mut Transcript::new(b"test_ring_encryption"),
588            )
589            .unwrap();
590    }
591
592    #[test]
593    fn single_ring_with_4_elements_works() {
594        let mut rng = rand::rng();
595        let keypair = Keypair::generate(&mut rng);
596        let log_base = keypair.public().as_element();
597        let admissible_values: Vec<_> = (0_u32..4)
598            .map(|i| Ristretto::mul_generator(&Scalar25519::from(i)))
599            .collect();
600
601        for _ in 0..100 {
602            let val: u32 = rng.random_range(0..4);
603            let element_val = Ristretto::mul_generator(&Scalar25519::from(val));
604            let ext_ciphertext = ExtendedCiphertext::new(element_val, keypair.public(), &mut rng);
605            let ciphertext = ext_ciphertext.inner;
606
607            let mut transcript = Transcript::new(b"test_ring_encryption");
608            RingProof::initialize_transcript(&mut transcript, keypair.public());
609
610            let mut ring_responses = vec![Scalar25519::default(); 4];
611            let signature_ring = Ring::new(
612                0,
613                log_base,
614                ext_ciphertext,
615                &admissible_values,
616                val as usize,
617                &transcript,
618                &mut ring_responses,
619                &mut rng,
620            );
621            let common_challenge =
622                Ring::aggregate(vec![signature_ring], log_base, &mut transcript, &mut rng);
623
624            RingProof::new(common_challenge, ring_responses)
625                .verify(
626                    keypair.public(),
627                    iter::once(admissible_values.as_slice()),
628                    iter::once(ciphertext),
629                    &mut Transcript::new(b"test_ring_encryption"),
630                )
631                .unwrap();
632        }
633    }
634
635    #[test_casing(5, 3..=7)]
636    fn multiple_rings_with_boolean_flags_work(ring_count: usize) {
637        let mut rng = rand::rng();
638        let keypair = Keypair::generate(&mut rng);
639        let log_base = keypair.public().as_element();
640        let admissible_values = [RistrettoPoint::identity(), Ristretto::generator()];
641
642        for _ in 0..20 {
643            let mut transcript = Transcript::new(b"test_ring_encryption");
644            RingProof::initialize_transcript(&mut transcript, keypair.public());
645
646            let mut ring_responses = vec![Scalar25519::default(); ring_count * 2];
647
648            let (ciphertexts, rings): (Vec<_>, Vec<_>) = ring_responses
649                .chunks_mut(2)
650                .enumerate()
651                .map(|(ring_index, ring_responses)| {
652                    let val: u32 = rng.random_range(0..=1);
653                    let element_val = Ristretto::mul_generator(&Scalar25519::from(val));
654                    let ext_ciphertext =
655                        ExtendedCiphertext::new(element_val, keypair.public(), &mut rng);
656                    let ciphertext = ext_ciphertext.inner;
657
658                    let signature_ring = Ring::new(
659                        ring_index,
660                        log_base,
661                        ext_ciphertext,
662                        &admissible_values,
663                        val as usize,
664                        &transcript,
665                        ring_responses,
666                        &mut rng,
667                    );
668
669                    (ciphertext, signature_ring)
670                })
671                .unzip();
672
673            let common_challenge = Ring::aggregate(rings, log_base, &mut transcript, &mut rng);
674
675            RingProof::new(common_challenge, ring_responses)
676                .verify(
677                    keypair.public(),
678                    iter::repeat_n(&admissible_values as &[_], ring_count),
679                    ciphertexts.into_iter(),
680                    &mut Transcript::new(b"test_ring_encryption"),
681                )
682                .unwrap();
683        }
684    }
685
686    #[test]
687    fn multiple_rings_with_base4_value_encoding_work() {
688        // We're testing ciphertexts of `u8` integers, hence 4 rings with 4 elements (=2 bits) each.
689        const RING_COUNT: u8 = 4;
690
691        // Admissible values are `[O, G, [2]G, [3]G]` for the first ring,
692        // `[O, [4]G, [8]G, [12]G]` for the second ring, etc.
693        let admissible_values: Vec<_> = (0..RING_COUNT)
694            .map(|ring_index| {
695                let power: u32 = 1 << (2 * u32::from(ring_index));
696                [
697                    RistrettoPoint::identity(),
698                    Ristretto::mul_generator(&Scalar25519::from(power)),
699                    Ristretto::mul_generator(&Scalar25519::from(power * 2)),
700                    Ristretto::mul_generator(&Scalar25519::from(power * 3)),
701                ]
702            })
703            .collect();
704
705        let mut rng = rand::rng();
706        let keypair = Keypair::generate(&mut rng);
707        let log_base = keypair.public().as_element();
708
709        for _ in 0..20 {
710            let overall_value: u8 = rng.random();
711            let mut transcript = Transcript::new(b"test_ring_encryption");
712            RingProof::initialize_transcript(&mut transcript, keypair.public());
713
714            let mut ring_responses = vec![Scalar25519::default(); RING_COUNT as usize * 4];
715
716            let (ciphertexts, rings): (Vec<_>, Vec<_>) = ring_responses
717                .chunks_mut(4)
718                .enumerate()
719                .map(|(ring_index, ring_responses)| {
720                    let mask = 3 << (2 * ring_index);
721                    let val = overall_value & mask;
722                    let val_index = (val >> (2 * ring_index)) as usize;
723                    assert!(val_index < 4);
724
725                    let element_val = Ristretto::mul_generator(&Scalar25519::from(val));
726                    let ext_ciphertext =
727                        ExtendedCiphertext::new(element_val, keypair.public(), &mut rng);
728                    let ciphertext = ext_ciphertext.inner;
729
730                    let signature_ring = Ring::new(
731                        ring_index,
732                        log_base,
733                        ext_ciphertext,
734                        &admissible_values[ring_index],
735                        val_index,
736                        &transcript,
737                        ring_responses,
738                        &mut rng,
739                    );
740
741                    (ciphertext, signature_ring)
742                })
743                .unzip();
744
745            let common_challenge = Ring::aggregate(rings, log_base, &mut transcript, &mut rng);
746            let admissible_values = admissible_values.iter().map(|values| values as &[_]);
747
748            RingProof::new(common_challenge, ring_responses)
749                .verify(
750                    keypair.public(),
751                    admissible_values,
752                    ciphertexts.into_iter(),
753                    &mut Transcript::new(b"test_ring_encryption"),
754                )
755                .unwrap();
756        }
757    }
758
759    #[test_casing(5, 3..=7)]
760    #[allow(clippy::needless_collect)]
761    // ^-- false positive; `builder` is captured by the iterator and moved by creating a `proof`
762    fn proof_builder_works(ring_count: usize) {
763        let mut rng = rand::rng();
764        let keypair = Keypair::generate(&mut rng);
765        let mut transcript = Transcript::new(b"test_ring_encryption");
766        let admissible_values = [RistrettoPoint::identity(), Ristretto::generator()];
767        let mut ring_responses = vec![Scalar25519::default(); ring_count * 2];
768
769        let mut builder = RingProofBuilder::new(
770            keypair.public(),
771            ring_count,
772            &mut ring_responses,
773            &mut transcript,
774            &mut rng,
775        );
776        let ciphertexts: Vec<_> = (0..ring_count)
777            .map(|i| builder.add_value(&admissible_values, i & 1).inner)
778            .collect();
779
780        RingProof::new(builder.build(), ring_responses)
781            .verify(
782                keypair.public(),
783                iter::repeat_n(&admissible_values as &[_], ring_count),
784                ciphertexts.into_iter(),
785                &mut Transcript::new(b"test_ring_encryption"),
786            )
787            .unwrap();
788    }
789}