elastic_elgamal/proofs/
range.rs

1//! Range proofs for ElGamal ciphertexts.
2
3use core::{convert::TryFrom, fmt};
4
5use elliptic_curve::{
6    rand_core::{CryptoRng, RngCore},
7    zeroize::Zeroizing,
8};
9use merlin::Transcript;
10#[cfg(feature = "serde")]
11use serde::{Deserialize, Serialize};
12use subtle::{ConditionallySelectable, ConstantTimeGreater};
13
14use crate::{
15    Ciphertext, PublicKey, VerificationError,
16    alloc::{HashMap, ToString, Vec, vec},
17    encryption::{CiphertextWithValue, ExtendedCiphertext},
18    group::Group,
19    proofs::{RingProof, RingProofBuilder, TranscriptForGroup},
20};
21
22#[derive(Debug, Clone, Copy, PartialEq)]
23struct RingSpec {
24    size: u64,
25    step: u64,
26}
27
28/// Decomposition of an integer range `0..n` into one or more sub-ranges. Decomposing the range
29/// allows constructing [`RangeProof`]s with size / computational complexity `O(log n)`.
30///
31/// # Construction
32///
33/// To build efficient `RangeProof`s, we need to be able to decompose any value `x` in `0..n`
34/// into several components, with each of them being in a smaller predefined range; once we
35/// have such a decomposition, we can build a [`RingProof`] around it.
36/// To build a decomposition, we use the following generic construction:
37///
38/// ```text
39/// 0..n = 0..t_0 + k_0 * (0..t_1 + k_1 * (0..t_2 + …)),
40/// ```
41///
42/// where `t_i` and `k_i` are integers greater than 1. If `x` is a value in `0..n`,
43/// it is decomposed as
44///
45/// ```text
46/// x = x_0 + k_0 * x_1 + k_0 * k_1 * x_2 + …; x_i in 0..t_i.
47/// ```
48///
49/// For a decomposition to be valid (i.e., to represent any value in `0..n` and no other values),
50/// the following statements are sufficient:
51///
52/// - `t_i >= k_i` (no gaps in values)
53/// - `n = t_0 + k_0 * (t_1 - 1 + k_1 * …)` (exact upper bound).
54///
55/// The size of a `RingProof` is the sum of upper range bounds `t_i` (= number of responses) + 1
56/// (the common challenge). Additionally, we need a ciphertext per each sub-range `0..t_i`
57/// (i.e., for each ring in `RingProof`). In practice, proof size is logarithmic:
58///
59/// | Upper bound `n`| Optimal decomposition | Proof size |
60/// |---------------:|-----------------------|-----------:|
61/// | 5              | `0..5`                | 6 scalars  |
62/// | 10             | `0..5 * 2 + 0..2`     | 8 scalars, 2 elements |
63/// | 20             | `0..5 * 4 + 0..4`     | 10 scalars, 2 elements |
64/// | 50             | `(0..5 * 5 + 0..5) * 2 + 0..2` | 13 scalars, 4 elements |
65/// | 64             | `(0..4 * 4 + 0..4) * 4 + 0..4` | 13 scalars, 4 elements |
66/// | 100            | `(0..5 * 5 + 0..5) * 4 + 0..4` | 15 scalars, 4 elements |
67/// | 256            | `((0..4 * 4 + 0..4) * 4 + 0..4) * 4 + 0..4` | 17 scalars, 6 elements |
68/// | 1000           | `((0..8 * 5 + 0..5) * 5 + 0..5) * 5 + 0..5` | 24 scalars, 6 elements |
69///
70/// (We do not count one of sub-range ciphertexts since it can be restored from the other
71/// sub-range ciphertexts and the original ciphertext of the value.)
72///
73/// ## Notes
74///
75/// - Decomposition of some values may be non-unique, but this is fine for our purposes.
76/// - Encoding of a value in a certain base is a partial case, with all `t_i` and `k_i` equal
77///   to the base. It only works for `n` being a power of the base.
78/// - Other types of decompositions may perform better, but this one has a couple
79///   of nice properties. It works for all `n`s, and the optimal decomposition can be found
80///   recursively.
81/// - If we know how to create / verify range proofs for `0..N`, proofs for all ranges `0..n`,
82///   `n < N` can be constructed as a combination of 2 proofs: a proof that encrypted value `x`
83///   is in `0..N` and that `n - 1 - x` is in `0..N`. (The latter is proved for a ciphertext
84///   obtained by the matching linear transform of the original ciphertext of `x`.)
85///   This does not help us if proofs for `0..N` are constructed using [`RingProof`]s,
86///   but allows estimating for which `n` a [Bulletproofs]-like construction would become
87///   more efficient despite using 2 proofs. If we take `N = 2^(2^P)`
88///   and the "vanilla" Bulletproof length `2 * P + 9`, this threshold is around `n = 2000`.
89///
90/// [Bulletproofs]: https://crypto.stanford.edu/bulletproofs/
91///
92/// # Examples
93///
94/// Finding out the optimal decomposition for a certain range:
95///
96/// ```
97/// # use elastic_elgamal::RangeDecomposition;
98/// let range = RangeDecomposition::optimal(42);
99/// assert_eq!(range.to_string(), "6 * 0..7 + 0..6");
100/// assert_eq!(range.proof_size(), 16); // 14 scalars, 2 elements
101///
102/// let range = RangeDecomposition::optimal(100);
103/// assert_eq!(range.to_string(), "20 * 0..5 + 4 * 0..5 + 0..4");
104/// assert_eq!(range.proof_size(), 19); // 15 scalars, 4 elements
105/// ```
106///
107/// See [`RangeProof`] docs for an end-to-end example of usage.
108#[derive(Debug, Clone, PartialEq)]
109pub struct RangeDecomposition {
110    rings: Vec<RingSpec>,
111}
112
113impl fmt::Display for RangeDecomposition {
114    fn fmt(&self, formatter: &mut fmt::Formatter<'_>) -> fmt::Result {
115        for (i, ring_spec) in self.rings.iter().enumerate() {
116            if ring_spec.step > 1 {
117                write!(formatter, "{} * ", ring_spec.step)?;
118            }
119            write!(formatter, "0..{}", ring_spec.size)?;
120
121            if i + 1 < self.rings.len() {
122                formatter.write_str(" + ")?;
123            }
124        }
125        Ok(())
126    }
127}
128
129/// `RangeDecomposition` together with optimized parameters.
130#[derive(Debug, Clone)]
131struct OptimalDecomposition {
132    decomposition: RangeDecomposition,
133    optimal_len: u64,
134}
135
136#[allow(
137    clippy::cast_possible_truncation,
138    clippy::cast_precision_loss,
139    clippy::cast_sign_loss
140)]
141impl RangeDecomposition {
142    /// Finds an optimal decomposition of the range with the given `upper_bound` in terms
143    /// of space of the range proof.
144    ///
145    /// Empirically, this method has sublinear complexity, but may work slowly for large values
146    /// of `upper_bound` (say, larger than 1 billion).
147    ///
148    /// # Panics
149    ///
150    /// Panics if `upper_bound` is less than 2.
151    pub fn optimal(upper_bound: u64) -> Self {
152        assert!(upper_bound >= 2, "`upper_bound` must be greater than 1");
153
154        let mut optimal_values = HashMap::new();
155        Self::optimize(upper_bound, &mut optimal_values).decomposition
156    }
157
158    fn just(capacity: u64) -> Self {
159        let spec = RingSpec {
160            size: capacity,
161            step: 1,
162        };
163        Self { rings: vec![spec] }
164    }
165
166    fn combine_mul(self, new_ring_size: u64, multiplier: u64) -> Self {
167        let mut combined_rings = self.rings;
168        for spec in &mut combined_rings {
169            spec.step *= multiplier;
170        }
171        combined_rings.push(RingSpec {
172            size: new_ring_size,
173            step: 1,
174        });
175
176        Self {
177            rings: combined_rings,
178        }
179    }
180
181    /// Returns the exclusive upper bound of the range presentable by this decomposition.
182    pub fn upper_bound(&self) -> u64 {
183        self.rings
184            .iter()
185            .map(|spec| (spec.size - 1) * spec.step)
186            .sum::<u64>()
187            + 1
188    }
189
190    /// Returns the total number of items in all rings.
191    fn rings_size(&self) -> u64 {
192        self.rings.iter().map(|spec| spec.size).sum::<u64>()
193    }
194
195    /// Returns the size of [`RangeProof`]s using this decomposition, measured as a total number
196    /// of scalars and group elements in the proof. Computational complexity of creating and
197    /// verifying proofs is also linear w.r.t. this number.
198    pub fn proof_size(&self) -> u64 {
199        self.rings_size() + 2 * self.rings.len() as u64 - 1
200    }
201
202    fn decompose(&self, value_indexes: &mut Vec<usize>, mut secret_value: u64) {
203        for ring_spec in &self.rings {
204            let mut value_index = secret_value / ring_spec.step;
205            let ring_max_value = ring_spec.size - 1;
206            let overflow = value_index.ct_gt(&ring_max_value);
207            value_index.conditional_assign(&ring_max_value, overflow);
208            value_indexes.push(value_index as usize);
209            secret_value -= value_index * ring_spec.step;
210        }
211
212        debug_assert_eq!(secret_value, 0, "unused secret value for {self:?}");
213    }
214
215    /// We decompose our range `0..n` as `0..t + k * 0..T`, where `t >= 2`, `T >= 2`,
216    /// `k >= 2`. For all values in the range to be presentable, we need `t >= k` (otherwise,
217    /// there will be gaps) and
218    ///
219    /// ```text
220    /// n - 1 = t - 1 + k * (T - 1) <=> n = t + k * (T - 1)
221    /// ```
222    ///
223    /// (to accurately represent the upper bound). For valid decompositions, we apply the
224    /// same decomposition recursively to `0..T`. If `P(n)` is the optimal proof length for
225    /// range `0..n`, we thus obtain
226    ///
227    /// ```text
228    /// P(n) = min_(t, k) { t + 2 + P((n - t) / k + 1) }.
229    /// ```
230    ///
231    /// Here, `t` is the number of commitments (= number of scalars for ring `0..t`), plus
232    /// 2 group elements in a partial ElGamal ciphertext corresponding to the ring.
233    ///
234    /// We additionally trim the solution space using a lower-bound estimate
235    ///
236    /// ```text
237    /// P(n) >= 3 * log2(n),
238    /// ```
239    ///
240    /// which can be proven recursively.
241    fn optimize(
242        upper_bound: u64,
243        optimal_values: &mut HashMap<u64, OptimalDecomposition>,
244    ) -> OptimalDecomposition {
245        if let Some(opt) = optimal_values.get(&upper_bound) {
246            return opt.clone();
247        }
248
249        let mut opt = OptimalDecomposition {
250            optimal_len: upper_bound + 2,
251            decomposition: RangeDecomposition::just(upper_bound),
252        };
253
254        for first_ring_size in 2_u64.. {
255            if first_ring_size + 2 > opt.optimal_len {
256                // Any further estimate will be worse than the current optimum.
257                break;
258            }
259
260            let remaining_capacity = upper_bound - first_ring_size;
261            for multiplier in 2_u64..=first_ring_size {
262                if remaining_capacity % multiplier != 0 {
263                    continue;
264                }
265                let inner_upper_bound = remaining_capacity / multiplier + 1;
266                if inner_upper_bound < 2 {
267                    // Since `inner_upper_bound` decreases w.r.t. `multiplier`, we can
268                    // break here.
269                    break;
270                }
271
272                let best_estimate =
273                    first_ring_size + 2 + Self::lower_len_estimate(inner_upper_bound);
274                if best_estimate > opt.optimal_len {
275                    continue;
276                }
277
278                let inner_opt = Self::optimize(inner_upper_bound, optimal_values);
279                let candidate_len = first_ring_size + 2 + inner_opt.optimal_len;
280                let candidate_rings = 1 + inner_opt.decomposition.rings.len();
281
282                if candidate_len < opt.optimal_len
283                    || (candidate_len == opt.optimal_len
284                        && candidate_rings < opt.decomposition.rings.len())
285                {
286                    opt.optimal_len = candidate_len;
287                    opt.decomposition = inner_opt
288                        .decomposition
289                        .combine_mul(first_ring_size, multiplier);
290                }
291            }
292        }
293
294        debug_assert!(
295            opt.optimal_len >= Self::lower_len_estimate(upper_bound),
296            "Lower len estimate {est} is invalid for {bound}: {opt:?}",
297            est = Self::lower_len_estimate(upper_bound),
298            bound = upper_bound,
299            opt = opt
300        );
301        optimal_values.insert(upper_bound, opt.clone());
302        opt
303    }
304
305    #[cfg(feature = "std")]
306    fn lower_len_estimate(upper_bound: u64) -> u64 {
307        ((upper_bound as f64).log2() * 3.0).ceil() as u64
308    }
309
310    #[cfg(not(feature = "std"))]
311    fn lower_len_estimate(upper_bound: u64) -> u64 {
312        Self::int_lower_len_estimate(upper_bound)
313    }
314
315    // We may not have floating-point arithmetics on no-std targets; thus, we use
316    // a less precise estimate.
317    #[cfg(any(test, not(feature = "std")))]
318    #[inline]
319    fn int_lower_len_estimate(upper_bound: u64) -> u64 {
320        let log2_upper_bound = if upper_bound == 0 {
321            0
322        } else {
323            63 - u64::from(upper_bound.leading_zeros()) // rounded down
324        };
325        log2_upper_bound * 3
326    }
327}
328
329/// [`RangeDecomposition`] together with values precached for creating and/or verifying
330/// [`RangeProof`]s in a certain [`Group`].
331#[derive(Debug, Clone)]
332pub struct PreparedRange<G: Group> {
333    inner: RangeDecomposition,
334    admissible_values: Vec<Vec<G::Element>>,
335}
336
337impl<G: Group> From<RangeDecomposition> for PreparedRange<G> {
338    fn from(decomposition: RangeDecomposition) -> Self {
339        Self::new(decomposition)
340    }
341}
342
343impl<G: Group> PreparedRange<G> {
344    fn new(inner: RangeDecomposition) -> Self {
345        let admissible_values = Vec::with_capacity(inner.rings.len());
346        let admissible_values = inner.rings.iter().fold(admissible_values, |mut acc, spec| {
347            let ring_values: Vec<_> = (0..spec.size)
348                .map(|i| G::vartime_mul_generator(&(i * spec.step).into()))
349                .collect();
350            acc.push(ring_values);
351            acc
352        });
353
354        Self {
355            inner,
356            admissible_values,
357        }
358    }
359
360    /// Returns a reference to the contained decomposition.
361    pub fn decomposition(&self) -> &RangeDecomposition {
362        &self.inner
363    }
364
365    /// Decomposes the provided `secret_value` into value indexes in constituent rings.
366    fn decompose(&self, secret_value: u64) -> Zeroizing<Vec<usize>> {
367        assert!(
368            secret_value < self.inner.upper_bound(),
369            "Secret value must be in range 0..{}",
370            self.inner.upper_bound()
371        );
372        // We immediately allocate the necessary capacity for `decomposition`.
373        let mut decomposition = Zeroizing::new(Vec::with_capacity(self.admissible_values.len()));
374        self.inner.decompose(&mut decomposition, secret_value);
375        decomposition
376    }
377}
378
379/// Zero-knowledge proof that an ElGamal ciphertext encrypts a value into a certain range `0..n`.
380///
381/// # Construction
382///
383/// To make the proof more compact – `O(log n)` in terms of size and proving / verification
384/// complexity – we use the same trick as for [Pedersen commitments] (used, e.g., for confidential
385/// transaction amounts in [Elements]):
386///
387/// 1. Represent the encrypted value `x` as `x = x_0 + k_0 * x_1 + k_0 * k_1 * x_2 + …`,
388///    where `0 <= x_i < t_i` is the decomposition of `x` as per the [`RangeDecomposition`],
389///    `0..t_0 + k_0 * (0..t_1 + …)`.
390///    As an example, if `n` is a power of 2, one can choose a decomposition as
391///    the base-2 presentation of `x`, i.e., `t_i = k_i = 2` for all `i`.
392///    For brevity, denote a multiplier of `x_i` in `x` decomposition as `K_i`,
393///    `K_i = k_0 * … * k_{i-1}`; `K_0 = 1` by extension.
394/// 2. Split the ciphertext: `E = E_0 + E_1 + …`, where `E_i` encrypts `K_i * x_i`.
395/// 3. Produce a [`RingProof`] that for all `i` the encrypted scalar for `E_i`
396///    is among 0, `K_i`, …, `K_i * (t_i - 1)`. The range proof consists of all `E_i` ciphertexts
397///    and this `RingProof`.
398///
399/// As with range proofs for Pedersen commitments, this construction is not optimal
400/// in terms of space or proving / verification complexity for large ranges;
401/// it is linear w.r.t. the bit length of the range.
402/// (Constructions like [Bulletproofs] are *logarithmic* w.r.t. the bit length.)
403/// Still, it can be useful for small ranges.
404///
405/// [Pedersen commitments]: https://en.wikipedia.org/wiki/Commitment_scheme
406/// [Elements]: https://elementsproject.org/features/confidential-transactions/investigation
407/// [Bulletproofs]: https://crypto.stanford.edu/bulletproofs/
408///
409/// # Examples
410///
411/// ```
412/// # use elastic_elgamal::{
413/// #     group::Ristretto, DiscreteLogTable, Keypair, RangeDecomposition, RangeProof, Ciphertext,
414/// # };
415/// # use merlin::Transcript;
416/// # fn main() -> Result<(), Box<dyn std::error::Error>> {
417/// // Generate the ciphertext receiver.
418/// let mut rng = rand::rng();
419/// let receiver = Keypair::<Ristretto>::generate(&mut rng);
420/// // Find the optimal range decomposition for our range
421/// // and specialize it for the Ristretto group.
422/// let range = RangeDecomposition::optimal(100).into();
423///
424/// let (ciphertext, proof) = RangeProof::new(
425///     receiver.public(),
426///     &range,
427///     55,
428///     &mut Transcript::new(b"test_proof"),
429///     &mut rng,
430/// );
431/// let ciphertext = Ciphertext::from(ciphertext);
432///
433/// // Check that the ciphertext is valid
434/// let lookup = DiscreteLogTable::new(0..100);
435/// assert_eq!(receiver.secret().decrypt(ciphertext, &lookup), Some(55));
436/// // ...and that the proof verifies.
437/// proof.verify(
438///     receiver.public(),
439///     &range,
440///     ciphertext,
441///     &mut Transcript::new(b"test_proof"),
442/// )?;
443/// # Ok(())
444/// # }
445/// ```
446#[derive(Debug, Clone)]
447#[cfg_attr(feature = "serde", derive(Serialize, Deserialize))]
448#[cfg_attr(feature = "serde", serde(bound = ""))]
449pub struct RangeProof<G: Group> {
450    partial_ciphertexts: Vec<Ciphertext<G>>,
451    #[cfg_attr(feature = "serde", serde(flatten))]
452    inner: RingProof<G>,
453}
454
455impl<G: Group> RangeProof<G> {
456    /// Encrypts `value` for `receiver` and creates a zero-knowledge proof that the encrypted value
457    /// is in `range`.
458    ///
459    /// This is a lower-level operation; see [`PublicKey::encrypt_range()`] for a higher-level
460    /// alternative.
461    ///
462    /// # Panics
463    ///
464    /// Panics if `value` is outside the range specified by `range`.
465    pub fn new<R: RngCore + CryptoRng>(
466        receiver: &PublicKey<G>,
467        range: &PreparedRange<G>,
468        value: u64,
469        transcript: &mut Transcript,
470        rng: &mut R,
471    ) -> (CiphertextWithValue<G, u64>, Self) {
472        let ciphertext = CiphertextWithValue::new(value, receiver, rng);
473        let proof = Self::from_ciphertext(receiver, range, &ciphertext, transcript, rng);
474        (ciphertext, proof)
475    }
476
477    /// Creates a proof that a value in `ciphertext` is in the `range`.
478    ///
479    /// The caller is responsible for providing a `ciphertext` encrypted for the `receiver`;
480    /// if the ciphertext is encrypted for another public key, the resulting proof will not verify.
481    ///
482    /// # Panics
483    ///
484    /// Panics if `value` is outside the range specified by `range`.
485    pub fn from_ciphertext<R: RngCore + CryptoRng>(
486        receiver: &PublicKey<G>,
487        range: &PreparedRange<G>,
488        ciphertext: &CiphertextWithValue<G, u64>,
489        transcript: &mut Transcript,
490        rng: &mut R,
491    ) -> Self {
492        let value_indexes = range.decompose(*ciphertext.value());
493        debug_assert_eq!(value_indexes.len(), range.admissible_values.len());
494        transcript.start_proof(b"encryption_range_proof");
495        transcript.append_message(b"range", range.inner.to_string().as_bytes());
496
497        let ring_responses_size = usize::try_from(range.inner.rings_size())
498            .expect("Integer overflow when allocating ring responses");
499        let mut ring_responses = vec![G::Scalar::default(); ring_responses_size];
500
501        let mut proof_builder = RingProofBuilder::new(
502            receiver,
503            range.admissible_values.len(),
504            &mut ring_responses,
505            transcript,
506            rng,
507        );
508
509        let mut cumulative_ciphertext = ExtendedCiphertext::zero();
510        let mut it = value_indexes.iter().zip(&range.admissible_values);
511
512        let partial_ciphertexts = it
513            .by_ref()
514            .take(value_indexes.len() - 1)
515            .map(|(value_index, admissible_values)| {
516                let ciphertext = proof_builder.add_value(admissible_values, *value_index);
517                let inner = ciphertext.inner;
518                cumulative_ciphertext += ciphertext;
519                inner
520            })
521            .collect();
522
523        let last_partial_ciphertext =
524            ciphertext.extended_ciphertext().clone() - cumulative_ciphertext;
525        let (&value_index, admissible_values) = it.next().unwrap();
526        // ^ `unwrap()` is safe by construction
527        proof_builder.add_precomputed_value(
528            last_partial_ciphertext,
529            admissible_values,
530            value_index,
531        );
532
533        Self {
534            partial_ciphertexts,
535            inner: RingProof::new(proof_builder.build(), ring_responses),
536        }
537    }
538
539    /// Verifies this proof against `ciphertext` for `receiver` and the specified `range`.
540    ///
541    /// This is a lower-level operation; see [`PublicKey::verify_range()`] for a higher-level
542    /// alternative.
543    ///
544    /// For a proof to verify, all parameters must be identical to ones provided when creating
545    /// the proof. In particular, `range` must have the same decomposition.
546    ///
547    /// # Errors
548    ///
549    /// Returns an error if this proof does not verify.
550    pub fn verify(
551        &self,
552        receiver: &PublicKey<G>,
553        range: &PreparedRange<G>,
554        ciphertext: Ciphertext<G>,
555        transcript: &mut Transcript,
556    ) -> Result<(), VerificationError> {
557        // Check decomposition / proof consistency.
558        VerificationError::check_lengths(
559            "admissible values",
560            self.partial_ciphertexts.len() + 1,
561            range.admissible_values.len(),
562        )?;
563
564        transcript.start_proof(b"encryption_range_proof");
565        transcript.append_message(b"range", range.inner.to_string().as_bytes());
566
567        let ciphertext_sum = self
568            .partial_ciphertexts
569            .iter()
570            .fold(Ciphertext::zero(), |acc, ciphertext| acc + *ciphertext);
571        let ciphertexts = self
572            .partial_ciphertexts
573            .iter()
574            .copied()
575            .chain(Some(ciphertext - ciphertext_sum));
576
577        let admissible_values = range.admissible_values.iter().map(Vec::as_slice);
578        self.inner
579            .verify(receiver, admissible_values, ciphertexts, transcript)
580    }
581}
582
583#[cfg(test)]
584mod tests {
585    use rand::Rng;
586    use test_casing::test_casing;
587
588    use super::*;
589    use crate::{
590        Keypair,
591        group::{ElementOps, Ristretto},
592    };
593
594    #[test]
595    fn optimal_value_small() {
596        let value = RangeDecomposition::optimal(5);
597        assert_eq!(value.rings.as_ref(), [RingSpec { size: 5, step: 1 }]);
598
599        let value = RangeDecomposition::optimal(16);
600        assert_eq!(
601            value.rings.as_ref(),
602            [RingSpec { size: 4, step: 4 }, RingSpec { size: 4, step: 1 }]
603        );
604
605        let value = RangeDecomposition::optimal(60);
606        assert_eq!(
607            value.rings.as_ref(),
608            [
609                RingSpec { size: 5, step: 12 },
610                RingSpec { size: 4, step: 3 },
611                RingSpec { size: 3, step: 1 },
612            ]
613        );
614
615        let value = RangeDecomposition::optimal(1_000);
616        assert_eq!(
617            value.to_string(),
618            "125 * 0..8 + 25 * 0..5 + 5 * 0..5 + 0..5"
619        );
620    }
621
622    #[test]
623    fn optimal_values_with_additives() {
624        let value = RangeDecomposition::optimal(17);
625        assert_eq!(
626            value.rings.as_ref(),
627            [RingSpec { size: 4, step: 4 }, RingSpec { size: 5, step: 1 }]
628        );
629
630        let value = RangeDecomposition::optimal(101);
631        assert_eq!(
632            value.rings.as_ref(),
633            [
634                RingSpec { size: 5, step: 20 },
635                RingSpec { size: 5, step: 4 },
636                RingSpec { size: 5, step: 1 }
637            ]
638        );
639    }
640
641    #[test]
642    fn large_optimal_values() {
643        let value = RangeDecomposition::optimal(12_345);
644        assert_eq!(
645            value.to_string(),
646            "2880 * 0..4 + 720 * 0..5 + 90 * 0..9 + 15 * 0..7 + 3 * 0..5 + 0..3"
647        );
648        assert_eq!(value.upper_bound(), 12_345);
649
650        let value = RangeDecomposition::optimal(777_777);
651        assert_eq!(
652            value.to_string(),
653            "125440 * 0..6 + 25088 * 0..6 + 3136 * 0..8 + 784 * 0..4 + 196 * 0..4 + \
654             49 * 0..5 + 7 * 0..7 + 0..7"
655        );
656        assert_eq!(value.upper_bound(), 777_777);
657
658        let value = RangeDecomposition::optimal(12_345_678);
659        assert_eq!(
660            value.to_string(),
661            "3072000 * 0..4 + 768000 * 0..4 + 192000 * 0..4 + 48000 * 0..5 + 9600 * 0..6 + \
662             1200 * 0..8 + 300 * 0..4 + 75 * 0..5 + 15 * 0..5 + 3 * 0..6 + 0..3"
663        );
664        assert_eq!(value.upper_bound(), 12_345_678);
665    }
666
667    #[test_casing(4, [1_000, 9_999, 12_345, 54_321])]
668    fn decomposing_for_larger_range(upper_bound: u64) {
669        let decomposition = RangeDecomposition::optimal(upper_bound);
670        let mut rng = rand::rng();
671
672        let values = (0..1_000)
673            .map(|_| rng.random_range(0..upper_bound))
674            .chain(0..5)
675            .chain((upper_bound - 5)..upper_bound);
676
677        for secret_value in values {
678            let mut value_indexes = vec![];
679            decomposition.decompose(&mut value_indexes, secret_value);
680
681            let restored = value_indexes
682                .iter()
683                .zip(&decomposition.rings)
684                .fold(0, |acc, (&idx, spec)| acc + idx as u64 * spec.step);
685            assert_eq!(
686                restored, secret_value,
687                "Cannot restore secret value {secret_value}; decomposed as {value_indexes:?}"
688            );
689        }
690    }
691
692    #[test]
693    fn decomposing_for_small_range() {
694        let decomposition = RangeDecomposition::optimal(17);
695        assert_eq!(decomposition.to_string(), "4 * 0..4 + 0..5");
696        let mut value_indexes = vec![];
697        decomposition.decompose(&mut value_indexes, 16);
698        assert_eq!(value_indexes, [3, 4]);
699        // 3 * 4 + 4 = 16
700    }
701
702    #[test]
703    fn decomposing_for_range() {
704        let decomposition = RangeDecomposition::optimal(1_000);
705        let mut value_indexes = vec![];
706        decomposition.decompose(&mut value_indexes, 567);
707        assert_eq!(value_indexes, [4, 2, 3, 2]);
708        // 2 + 3 * 5 + 2 * 25 + 4 * 125 = 567
709    }
710
711    #[test_casing(4, [12, 15, 20, 50])]
712    fn range_proof_basics(upper_bound: u64) {
713        let decomposition = RangeDecomposition::optimal(upper_bound).into();
714
715        let mut rng = rand::rng();
716        let receiver = Keypair::<Ristretto>::generate(&mut rng);
717        let (ciphertext, proof) = RangeProof::new(
718            receiver.public(),
719            &decomposition,
720            10,
721            &mut Transcript::new(b"test"),
722            &mut rng,
723        );
724        let ciphertext = ciphertext.into();
725
726        proof
727            .verify(
728                receiver.public(),
729                &decomposition,
730                ciphertext,
731                &mut Transcript::new(b"test"),
732            )
733            .unwrap();
734
735        // Should not verify with another transcript context
736        assert!(
737            proof
738                .verify(
739                    receiver.public(),
740                    &decomposition,
741                    ciphertext,
742                    &mut Transcript::new(b"other"),
743                )
744                .is_err()
745        );
746
747        // ...or with another receiver
748        let other_receiver = Keypair::<Ristretto>::generate(&mut rng);
749        assert!(
750            proof
751                .verify(
752                    other_receiver.public(),
753                    &decomposition,
754                    ciphertext,
755                    &mut Transcript::new(b"test"),
756                )
757                .is_err()
758        );
759
760        // ...or with another ciphertext
761        let other_ciphertext = receiver.public().encrypt(10_u64, &mut rng);
762        assert!(
763            proof
764                .verify(
765                    receiver.public(),
766                    &decomposition,
767                    other_ciphertext,
768                    &mut Transcript::new(b"test"),
769                )
770                .is_err()
771        );
772
773        let mut mangled_ciphertext = ciphertext;
774        mangled_ciphertext.blinded_element += Ristretto::generator();
775        assert!(
776            proof
777                .verify(
778                    receiver.public(),
779                    &decomposition,
780                    mangled_ciphertext,
781                    &mut Transcript::new(b"test"),
782                )
783                .is_err()
784        );
785
786        // ...or with another decomposition
787        let other_decomposition = RangeDecomposition::just(15).into();
788        assert!(
789            proof
790                .verify(
791                    receiver.public(),
792                    &other_decomposition,
793                    ciphertext,
794                    &mut Transcript::new(b"test"),
795                )
796                .is_err()
797        );
798    }
799
800    #[test]
801    #[cfg(feature = "std")]
802    fn int_lower_len_estimate_is_always_not_more_than_exact() {
803        let samples = (0..1_000).chain((1..1_000).map(|i| i * 1_000));
804        for sample in samples {
805            let floating_point_estimate = RangeDecomposition::lower_len_estimate(sample);
806            let int_estimate = RangeDecomposition::int_lower_len_estimate(sample);
807            assert!(
808                floating_point_estimate >= int_estimate,
809                "Unexpected estimates for {sample}: floating-point = {floating_point_estimate}, \
810                 int = {int_estimate}"
811            );
812        }
813    }
814}