snarkvm_algorithms/polycommit/kzg10/
mod.rs

1// Copyright 2024 Aleo Network Foundation
2// This file is part of the snarkVM library.
3
4// Licensed under the Apache License, Version 2.0 (the "License");
5// you may not use this file except in compliance with the License.
6// You may obtain a copy of the License at:
7
8// http://www.apache.org/licenses/LICENSE-2.0
9
10// Unless required by applicable law or agreed to in writing, software
11// distributed under the License is distributed on an "AS IS" BASIS,
12// WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
13// See the License for the specific language governing permissions and
14// limitations under the License.
15
16//! Here we construct a polynomial commitment that enables users to commit to a
17//! single polynomial `p`, and then later provide an evaluation proof that
18//! convinces verifiers that a claimed value `v` is the true evaluation of `p`
19//! at a chosen point `x`. Our construction follows the template of the construction
20//! proposed by Kate, Zaverucha, and Goldberg ([KZG11](http://cacr.uwaterloo.ca/techreports/2010/cacr2010-10.pdf)).
21//! This construction achieves extractability in the algebraic group model (AGM).
22
23use crate::{
24    fft::{DensePolynomial, Polynomial},
25    msm::VariableBase,
26    polycommit::PCError,
27};
28use snarkvm_curves::traits::{AffineCurve, PairingCurve, PairingEngine, ProjectiveCurve};
29use snarkvm_fields::{One, PrimeField, Zero};
30use snarkvm_utilities::{BitIteratorBE, cfg_iter, cfg_iter_mut, rand::Uniform};
31
32use anyhow::{Result, anyhow, ensure};
33use core::{marker::PhantomData, ops::Mul};
34use itertools::Itertools;
35use rand_core::RngCore;
36
37#[cfg(not(feature = "serial"))]
38use rayon::prelude::*;
39
40mod data_structures;
41pub use data_structures::*;
42
43use super::sonic_pc::LabeledPolynomialWithBasis;
44
45#[derive(Debug, PartialEq, Eq)]
46pub enum KZGDegreeBounds {
47    All,
48    Varuna,
49    List(Vec<usize>),
50    None,
51}
52
53impl KZGDegreeBounds {
54    pub fn get_list<F: PrimeField>(&self, max_degree: usize) -> Vec<usize> {
55        match self {
56            KZGDegreeBounds::All => (0..max_degree).collect(),
57            KZGDegreeBounds::Varuna => {
58                // In Varuna, the degree bounds are all of the forms `domain_size - 2`.
59                // Consider that we are using radix-2 FFT,
60                // there are only a few possible domain sizes and therefore degree bounds.
61                //
62                // We do not consider mixed-radix FFT for simplicity, as the curves that we
63                // are using have very high two-arity.
64
65                let mut radix_2_possible_domain_sizes = vec![];
66
67                let mut cur = 2usize;
68                while cur - 2 <= max_degree {
69                    radix_2_possible_domain_sizes.push(cur - 2);
70                    cur *= 2;
71                }
72
73                radix_2_possible_domain_sizes
74            }
75            KZGDegreeBounds::List(v) => v.clone(),
76            KZGDegreeBounds::None => vec![],
77        }
78    }
79}
80
81/// `KZG10` is an implementation of the polynomial commitment scheme of
82/// [Kate, Zaverucha and Goldbgerg][kzg10]
83///
84/// [kzg10]: http://cacr.uwaterloo.ca/techreports/2010/cacr2010-10.pdf
85#[derive(Clone, Debug)]
86pub struct KZG10<E: PairingEngine>(PhantomData<E>);
87
88impl<E: PairingEngine> KZG10<E> {
89    /// Constructs public parameters when given as input the maximum degree `degree`
90    /// for the polynomial commitment scheme.
91    pub fn load_srs(max_degree: usize) -> Result<UniversalParams<E>, PCError> {
92        let params = UniversalParams::load()?;
93        params.download_powers_for(0..(max_degree + 1))?;
94        Ok(params)
95    }
96
97    /// Outputs a commitment to `polynomial`.
98    pub fn commit(
99        powers: &Powers<E>,
100        polynomial: &Polynomial<'_, E::Fr>,
101        hiding_bound: Option<usize>,
102        rng: Option<&mut dyn RngCore>,
103    ) -> Result<(KZGCommitment<E>, KZGRandomness<E>), PCError> {
104        Self::check_degree_is_too_large(polynomial.degree(), powers.size())?;
105
106        let commit_time = start_timer!(|| format!(
107            "Committing to polynomial of degree {} with hiding_bound: {:?}",
108            polynomial.degree(),
109            hiding_bound,
110        ));
111
112        let mut commitment = match polynomial {
113            Polynomial::Dense(polynomial) => {
114                let (num_leading_zeros, plain_coeffs) = skip_leading_zeros_and_convert_to_bigints(polynomial);
115
116                let bases = &powers.powers_of_beta_g[num_leading_zeros..(num_leading_zeros + plain_coeffs.len())];
117
118                let msm_time = start_timer!(|| "MSM to compute commitment to plaintext poly");
119                let commitment = VariableBase::msm(bases, &plain_coeffs);
120                end_timer!(msm_time);
121
122                commitment
123            }
124            Polynomial::Sparse(polynomial) => polynomial
125                .coeffs()
126                .map(|(i, coeff)| {
127                    powers.powers_of_beta_g[*i].mul_bits(BitIteratorBE::new_without_leading_zeros(coeff.to_bigint()))
128                })
129                .sum(),
130        };
131
132        let mut randomness = KZGRandomness::empty();
133        if let Some(hiding_degree) = hiding_bound {
134            let mut rng = rng.ok_or(PCError::MissingRng)?;
135            let sample_random_poly_time =
136                start_timer!(|| format!("Sampling a random polynomial of degree {hiding_degree}"));
137
138            randomness = KZGRandomness::rand(hiding_degree, false, &mut rng);
139            Self::check_hiding_bound(
140                randomness.blinding_polynomial.degree(),
141                powers.powers_of_beta_times_gamma_g.len(),
142            )?;
143            end_timer!(sample_random_poly_time);
144        }
145
146        let random_ints = convert_to_bigints(&randomness.blinding_polynomial.coeffs);
147        let msm_time = start_timer!(|| "MSM to compute commitment to random poly");
148        let random_commitment =
149            VariableBase::msm(&powers.powers_of_beta_times_gamma_g, random_ints.as_slice()).to_affine();
150        end_timer!(msm_time);
151
152        commitment.add_assign_mixed(&random_commitment);
153
154        end_timer!(commit_time);
155        Ok((KZGCommitment(commitment.into()), randomness))
156    }
157
158    /// Outputs a commitment to `polynomial`.
159    pub fn commit_lagrange(
160        lagrange_basis: &LagrangeBasis<E>,
161        evaluations: &[E::Fr],
162        hiding_bound: Option<usize>,
163        rng: Option<&mut dyn RngCore>,
164    ) -> Result<(KZGCommitment<E>, KZGRandomness<E>), PCError> {
165        Self::check_degree_is_too_large(evaluations.len() - 1, lagrange_basis.size())?;
166        assert_eq!(
167            evaluations.len().checked_next_power_of_two().ok_or(PCError::LagrangeBasisSizeIsTooLarge)?,
168            lagrange_basis.size()
169        );
170
171        let commit_time = start_timer!(|| format!(
172            "Committing to polynomial of degree {} with hiding_bound: {:?}",
173            evaluations.len() - 1,
174            hiding_bound,
175        ));
176
177        let evaluations = evaluations.iter().map(|e| e.to_bigint()).collect::<Vec<_>>();
178        let msm_time = start_timer!(|| "MSM to compute commitment to plaintext poly");
179        let mut commitment = VariableBase::msm(&lagrange_basis.lagrange_basis_at_beta_g, &evaluations);
180        end_timer!(msm_time);
181
182        let mut randomness = KZGRandomness::empty();
183        if let Some(hiding_degree) = hiding_bound {
184            let mut rng = rng.ok_or(PCError::MissingRng)?;
185            let sample_random_poly_time =
186                start_timer!(|| format!("Sampling a random polynomial of degree {hiding_degree}"));
187
188            randomness = KZGRandomness::rand(hiding_degree, false, &mut rng);
189            Self::check_hiding_bound(
190                randomness.blinding_polynomial.degree(),
191                lagrange_basis.powers_of_beta_times_gamma_g.len(),
192            )?;
193            end_timer!(sample_random_poly_time);
194        }
195
196        let random_ints = convert_to_bigints(&randomness.blinding_polynomial.coeffs);
197        let msm_time = start_timer!(|| "MSM to compute commitment to random poly");
198        let random_commitment =
199            VariableBase::msm(&lagrange_basis.powers_of_beta_times_gamma_g, random_ints.as_slice()).to_affine();
200        end_timer!(msm_time);
201
202        commitment.add_assign_mixed(&random_commitment);
203
204        end_timer!(commit_time);
205        Ok((KZGCommitment(commitment.into()), randomness))
206    }
207
208    /// Compute witness polynomial.
209    ///
210    /// The witness polynomial w(x) the quotient of the division (p(x) - p(z)) / (x - z)
211    /// Observe that this quotient does not change with z because
212    /// p(z) is the remainder term. We can therefore omit p(z) when computing the quotient.
213    pub fn compute_witness_polynomial(
214        polynomial: &DensePolynomial<E::Fr>,
215        point: E::Fr,
216        randomness: &KZGRandomness<E>,
217    ) -> Result<(DensePolynomial<E::Fr>, Option<DensePolynomial<E::Fr>>), PCError> {
218        let divisor = DensePolynomial::from_coefficients_vec(vec![-point, E::Fr::one()]);
219
220        let witness_time = start_timer!(|| "Computing witness polynomial");
221        let witness_polynomial = polynomial / &divisor;
222        end_timer!(witness_time);
223
224        let random_witness_polynomial = if randomness.is_hiding() {
225            let random_p = &randomness.blinding_polynomial;
226
227            let witness_time = start_timer!(|| "Computing random witness polynomial");
228            let random_witness_polynomial = random_p / &divisor;
229            end_timer!(witness_time);
230            Some(random_witness_polynomial)
231        } else {
232            None
233        };
234
235        Ok((witness_polynomial, random_witness_polynomial))
236    }
237
238    pub(crate) fn open_with_witness_polynomial(
239        powers: &Powers<E>,
240        point: E::Fr,
241        randomness: &KZGRandomness<E>,
242        witness_polynomial: &DensePolynomial<E::Fr>,
243        hiding_witness_polynomial: Option<&DensePolynomial<E::Fr>>,
244    ) -> Result<KZGProof<E>, PCError> {
245        Self::check_degree_is_too_large(witness_polynomial.degree(), powers.size())?;
246        let (num_leading_zeros, witness_coeffs) = skip_leading_zeros_and_convert_to_bigints(witness_polynomial);
247
248        let bases = &powers.powers_of_beta_g[num_leading_zeros..(num_leading_zeros + witness_coeffs.len())];
249
250        let witness_comm_time = start_timer!(|| "Computing commitment to witness polynomial");
251        let mut w = VariableBase::msm(bases, &witness_coeffs);
252        end_timer!(witness_comm_time);
253
254        let random_v = if let Some(hiding_witness_polynomial) = hiding_witness_polynomial {
255            let blinding_p = &randomness.blinding_polynomial;
256            let blinding_eval_time = start_timer!(|| "Evaluating random polynomial");
257            let blinding_evaluation = blinding_p.evaluate(point);
258            end_timer!(blinding_eval_time);
259
260            let random_witness_coeffs = convert_to_bigints(&hiding_witness_polynomial.coeffs);
261            let witness_comm_time = start_timer!(|| "Computing commitment to random witness polynomial");
262            w += &VariableBase::msm(&powers.powers_of_beta_times_gamma_g, &random_witness_coeffs);
263            end_timer!(witness_comm_time);
264            Some(blinding_evaluation)
265        } else {
266            None
267        };
268
269        Ok(KZGProof { w: w.to_affine(), random_v })
270    }
271
272    /// On input a polynomial `p` in Lagrange basis, and a point `point`,
273    /// outputs an evaluation proof for the same.
274    pub fn open_lagrange(
275        lagrange_basis: &LagrangeBasis<E>,
276        domain_elements: &[E::Fr],
277        evaluations: &[E::Fr],
278        point: E::Fr,
279        evaluation_at_point: E::Fr,
280    ) -> Result<KZGProof<E>> {
281        Self::check_degree_is_too_large(evaluations.len() - 1, lagrange_basis.size())?;
282        // Ensure that the point is not in the domain
283        if lagrange_basis.domain.evaluate_vanishing_polynomial(point).is_zero() {
284            Err(anyhow!("Point cannot be in the domain"))?;
285        }
286        if evaluations.len().checked_next_power_of_two().ok_or_else(|| anyhow!("Evaluations length is too large"))?
287            != lagrange_basis.size()
288        {
289            Err(anyhow!("`evaluations.len()` must equal `domain.size()`"))?;
290        }
291
292        let mut divisor_evals = cfg_iter!(domain_elements).map(|&e| e - point).collect::<Vec<_>>();
293        snarkvm_fields::batch_inversion(&mut divisor_evals);
294        ensure!(divisor_evals.len() == evaluations.len());
295        cfg_iter_mut!(divisor_evals).zip_eq(evaluations).for_each(|(divisor_eval, &eval)| {
296            *divisor_eval *= eval - evaluation_at_point;
297        });
298        let (witness_comm, _) = Self::commit_lagrange(lagrange_basis, &divisor_evals, None, None)?;
299
300        Ok(KZGProof { w: witness_comm.0, random_v: None })
301    }
302
303    /// On input a polynomial `p` and a point `point`, outputs a proof for the same.
304    pub fn open(
305        powers: &Powers<E>,
306        polynomial: &DensePolynomial<E::Fr>,
307        point: E::Fr,
308        rand: &KZGRandomness<E>,
309    ) -> Result<KZGProof<E>, PCError> {
310        Self::check_degree_is_too_large(polynomial.degree(), powers.size())?;
311        let open_time = start_timer!(|| format!("Opening polynomial of degree {}", polynomial.degree()));
312
313        let witness_time = start_timer!(|| "Computing witness polynomials");
314        let (witness_poly, hiding_witness_poly) = Self::compute_witness_polynomial(polynomial, point, rand)?;
315        end_timer!(witness_time);
316
317        let proof =
318            Self::open_with_witness_polynomial(powers, point, rand, &witness_poly, hiding_witness_poly.as_ref());
319
320        end_timer!(open_time);
321        proof
322    }
323
324    /// Verifies that `value` is the evaluation at `point` of the polynomial
325    /// committed inside `commitment`.
326    pub fn check(
327        vk: &VerifierKey<E>,
328        commitment: &KZGCommitment<E>,
329        point: E::Fr,
330        value: E::Fr,
331        proof: &KZGProof<E>,
332    ) -> Result<bool, PCError> {
333        let check_time = start_timer!(|| "Checking evaluation");
334        let mut inner = commitment.0.to_projective() - vk.g.to_projective().mul(value);
335        if let Some(random_v) = proof.random_v {
336            inner -= &vk.gamma_g.mul(random_v);
337        }
338        let lhs = E::pairing(inner, vk.h);
339
340        let inner = vk.beta_h.to_projective() - vk.h.mul(point);
341        let rhs = E::pairing(proof.w, inner);
342
343        end_timer!(check_time, || format!("Result: {}", lhs == rhs));
344        Ok(lhs == rhs)
345    }
346
347    /// Check that each `proof_i` in `proofs` is a valid proof of evaluation for
348    /// `commitment_i` at `point_i`.
349    pub fn batch_check<R: RngCore>(
350        vk: &VerifierKey<E>,
351        commitments: &[KZGCommitment<E>],
352        points: &[E::Fr],
353        values: &[E::Fr],
354        proofs: &[KZGProof<E>],
355        rng: &mut R,
356    ) -> Result<bool> {
357        let check_time = start_timer!(|| format!("Checking {} evaluation proofs", commitments.len()));
358        let g = vk.g.to_projective();
359        let gamma_g = vk.gamma_g.to_projective();
360
361        let mut total_c = <E::G1Projective>::zero();
362        let mut total_w = <E::G1Projective>::zero();
363
364        let combination_time = start_timer!(|| "Combining commitments and proofs");
365        let mut randomizer = E::Fr::one();
366        // Instead of multiplying g and gamma_g in each turn, we simply accumulate
367        // their coefficients and perform a final multiplication at the end.
368        let mut g_multiplier = E::Fr::zero();
369        let mut gamma_g_multiplier = E::Fr::zero();
370        ensure!(commitments.len() == points.len());
371        ensure!(commitments.len() == values.len());
372        ensure!(commitments.len() == proofs.len());
373        for (((c, z), v), proof) in commitments.iter().zip_eq(points).zip_eq(values).zip_eq(proofs) {
374            let w = proof.w;
375            let mut temp = w.mul(*z);
376            temp.add_assign_mixed(&c.0);
377            let c = temp;
378            g_multiplier += &(randomizer * v);
379            if let Some(random_v) = proof.random_v {
380                gamma_g_multiplier += &(randomizer * random_v);
381            }
382            total_c += &c.mul(randomizer);
383            total_w += &w.mul(randomizer);
384            // We don't need to sample randomizers from the full field,
385            // only from 128-bit strings.
386            randomizer = u128::rand(rng).into();
387        }
388        total_c -= &g.mul(g_multiplier);
389        total_c -= &gamma_g.mul(gamma_g_multiplier);
390        end_timer!(combination_time);
391
392        let to_affine_time = start_timer!(|| "Converting results to affine for pairing");
393        let affine_points = E::G1Projective::batch_normalization_into_affine(vec![-total_w, total_c]);
394        let (total_w, total_c) = (affine_points[0], affine_points[1]);
395        end_timer!(to_affine_time);
396
397        let pairing_time = start_timer!(|| "Performing product of pairings");
398        let result = E::product_of_pairings(
399            [(&total_w.prepare(), &vk.prepared_beta_h), (&total_c.prepare(), &vk.prepared_h)].iter().copied(),
400        )
401        .is_one();
402        end_timer!(pairing_time);
403        end_timer!(check_time, || format!("Result: {result}"));
404        Ok(result)
405    }
406
407    pub(crate) fn check_degree_is_too_large(degree: usize, num_powers: usize) -> Result<(), PCError> {
408        let num_coefficients = degree + 1;
409        if num_coefficients > num_powers {
410            Err(PCError::TooManyCoefficients { num_coefficients, num_powers })
411        } else {
412            Ok(())
413        }
414    }
415
416    pub(crate) fn check_hiding_bound(hiding_poly_degree: usize, num_powers: usize) -> Result<(), PCError> {
417        if hiding_poly_degree == 0 {
418            Err(PCError::HidingBoundIsZero)
419        } else if hiding_poly_degree >= num_powers {
420            // The above check uses `>=` because committing to a hiding poly with
421            // degree `hiding_poly_degree` requires `hiding_poly_degree + 1` powers.
422            Err(PCError::HidingBoundToolarge { hiding_poly_degree, num_powers })
423        } else {
424            Ok(())
425        }
426    }
427
428    pub(crate) fn check_degrees_and_bounds<'a>(
429        max_degree: usize,
430        enforced_degree_bounds: Option<&[usize]>,
431        p: impl Into<LabeledPolynomialWithBasis<'a, E::Fr>>,
432    ) -> Result<(), PCError> {
433        let p = p.into();
434        if let Some(bound) = p.degree_bound() {
435            let enforced_degree_bounds = enforced_degree_bounds.ok_or(PCError::UnsupportedDegreeBound(bound))?;
436
437            if enforced_degree_bounds.binary_search(&bound).is_err() {
438                Err(PCError::UnsupportedDegreeBound(bound))
439            } else if bound < p.degree() || bound > max_degree {
440                return Err(PCError::IncorrectDegreeBound {
441                    poly_degree: p.degree(),
442                    degree_bound: p.degree_bound().unwrap(),
443                    max_degree,
444                    label: p.label().to_string(),
445                });
446            } else {
447                Ok(())
448            }
449        } else {
450            Ok(())
451        }
452    }
453}
454
455fn skip_leading_zeros_and_convert_to_bigints<F: PrimeField>(p: &DensePolynomial<F>) -> (usize, Vec<F::BigInteger>) {
456    if p.coeffs.is_empty() {
457        (0, vec![])
458    } else {
459        let num_leading_zeros = p.coeffs.iter().take_while(|c| c.is_zero()).count();
460        let coeffs = if num_leading_zeros == p.coeffs.len() {
461            vec![]
462        } else {
463            convert_to_bigints(&p.coeffs[num_leading_zeros..])
464        };
465        (num_leading_zeros, coeffs)
466    }
467}
468
469fn convert_to_bigints<F: PrimeField>(p: &[F]) -> Vec<F::BigInteger> {
470    let to_bigint_time = start_timer!(|| "Converting polynomial coeffs to bigints");
471    let coeffs = cfg_iter!(p).map(|s| s.to_bigint()).collect::<Vec<_>>();
472    end_timer!(to_bigint_time);
473    coeffs
474}
475
476#[cfg(test)]
477mod tests {
478    #![allow(non_camel_case_types)]
479    #![allow(clippy::needless_borrow)]
480    use super::*;
481    use snarkvm_curves::bls12_377::{Bls12_377, Fr};
482    use snarkvm_utilities::{FromBytes, ToBytes, rand::TestRng};
483
484    use std::borrow::Cow;
485
486    type KZG_Bls12_377 = KZG10<Bls12_377>;
487
488    impl<E: PairingEngine> KZG10<E> {
489        /// Specializes the public parameters for a given maximum degree `d` for polynomials
490        /// `d` should be less that `pp.max_degree()`.
491        pub(crate) fn trim(
492            pp: &UniversalParams<E>,
493            mut supported_degree: usize,
494            hiding_bound: Option<usize>,
495        ) -> (Powers<E>, VerifierKey<E>) {
496            if supported_degree == 1 {
497                supported_degree += 1;
498            }
499            let powers_of_beta_g = pp.powers_of_beta_g(0, supported_degree + 1).unwrap().to_vec();
500
501            let powers_of_beta_times_gamma_g = if let Some(hiding_bound) = hiding_bound {
502                (0..=(hiding_bound + 1)).map(|i| pp.powers_of_beta_times_gamma_g()[&i]).collect()
503            } else {
504                vec![]
505            };
506
507            let powers = Powers {
508                powers_of_beta_g: Cow::Owned(powers_of_beta_g),
509                powers_of_beta_times_gamma_g: Cow::Owned(powers_of_beta_times_gamma_g),
510            };
511            let vk = VerifierKey {
512                g: pp.power_of_beta_g(0).unwrap(),
513                gamma_g: pp.powers_of_beta_times_gamma_g()[&0],
514                h: pp.h,
515                beta_h: pp.beta_h(),
516                prepared_h: pp.prepared_h.clone(),
517                prepared_beta_h: pp.prepared_beta_h.clone(),
518            };
519            (powers, vk)
520        }
521    }
522
523    #[test]
524    fn test_kzg10_universal_params_serialization() {
525        let degree = 4;
526        let pp = KZG_Bls12_377::load_srs(degree).unwrap();
527
528        let pp_bytes = pp.to_bytes_le().unwrap();
529        let pp_recovered: UniversalParams<Bls12_377> = FromBytes::read_le(&pp_bytes[..]).unwrap();
530        let pp_recovered_bytes = pp_recovered.to_bytes_le().unwrap();
531
532        assert_eq!(&pp_bytes, &pp_recovered_bytes);
533    }
534
535    fn end_to_end_test_template<E: PairingEngine>() -> Result<(), PCError> {
536        let rng = &mut TestRng::default();
537        for _ in 0..100 {
538            let mut degree = 0;
539            while degree <= 1 {
540                degree = usize::rand(rng) % 20;
541            }
542            let pp = KZG10::<E>::load_srs(degree)?;
543            let hiding_bound = Some(1);
544            let (ck, vk) = KZG10::trim(&pp, degree, hiding_bound);
545            let p = DensePolynomial::rand(degree, rng);
546            let (comm, rand) = KZG10::<E>::commit(&ck, &(&p).into(), hiding_bound, Some(rng))?;
547            let point = E::Fr::rand(rng);
548            let value = p.evaluate(point);
549            let proof = KZG10::<E>::open(&ck, &p, point, &rand)?;
550            assert!(
551                KZG10::<E>::check(&vk, &comm, point, value, &proof)?,
552                "proof was incorrect for max_degree = {}, polynomial_degree = {}, hiding_bound = {:?}",
553                degree,
554                p.degree(),
555                hiding_bound,
556            );
557        }
558        Ok(())
559    }
560
561    fn linear_polynomial_test_template<E: PairingEngine>() -> Result<(), PCError> {
562        let rng = &mut TestRng::default();
563        for _ in 0..100 {
564            let degree = 50;
565            let pp = KZG10::<E>::load_srs(degree)?;
566            let hiding_bound = Some(1);
567            let (ck, vk) = KZG10::trim(&pp, 2, hiding_bound);
568            let p = DensePolynomial::rand(1, rng);
569            let (comm, rand) = KZG10::<E>::commit(&ck, &(&p).into(), hiding_bound, Some(rng))?;
570            let point = E::Fr::rand(rng);
571            let value = p.evaluate(point);
572            let proof = KZG10::<E>::open(&ck, &p, point, &rand)?;
573            assert!(
574                KZG10::<E>::check(&vk, &comm, point, value, &proof)?,
575                "proof was incorrect for max_degree = {}, polynomial_degree = {}, hiding_bound = {:?}",
576                degree,
577                p.degree(),
578                hiding_bound,
579            );
580        }
581        Ok(())
582    }
583
584    fn batch_check_test_template<E: PairingEngine>() -> Result<(), PCError> {
585        let rng = &mut TestRng::default();
586        for _ in 0..10 {
587            let hiding_bound = Some(1);
588            let mut degree = 0;
589            while degree <= 1 {
590                degree = usize::rand(rng) % 20;
591            }
592            let pp = KZG10::<E>::load_srs(degree)?;
593            let (ck, vk) = KZG10::trim(&pp, degree, hiding_bound);
594
595            let mut comms = Vec::new();
596            let mut values = Vec::new();
597            let mut points = Vec::new();
598            let mut proofs = Vec::new();
599
600            for _ in 0..10 {
601                let p = DensePolynomial::rand(degree, rng);
602                let (comm, rand) = KZG10::<E>::commit(&ck, &(&p).into(), hiding_bound, Some(rng))?;
603                let point = E::Fr::rand(rng);
604                let value = p.evaluate(point);
605                let proof = KZG10::<E>::open(&ck, &p, point, &rand)?;
606
607                assert!(KZG10::<E>::check(&vk, &comm, point, value, &proof)?);
608                comms.push(comm);
609                values.push(value);
610                points.push(point);
611                proofs.push(proof);
612            }
613            assert!(KZG10::<E>::batch_check(&vk, &comms, &points, &values, &proofs, rng)?);
614        }
615        Ok(())
616    }
617
618    #[test]
619    fn test_end_to_end() {
620        end_to_end_test_template::<Bls12_377>().expect("test failed for bls12-377");
621    }
622
623    #[test]
624    fn test_linear_polynomial() {
625        linear_polynomial_test_template::<Bls12_377>().expect("test failed for bls12-377");
626    }
627
628    #[test]
629    fn test_batch_check() {
630        batch_check_test_template::<Bls12_377>().expect("test failed for bls12-377");
631    }
632
633    #[test]
634    fn test_degree_is_too_large() {
635        let rng = &mut TestRng::default();
636
637        let max_degree = 123;
638        let pp = KZG_Bls12_377::load_srs(max_degree).unwrap();
639        let (powers, _) = KZG_Bls12_377::trim(&pp, max_degree, None);
640
641        let p = DensePolynomial::<Fr>::rand(max_degree + 1, rng);
642        assert!(p.degree() > max_degree);
643        assert!(KZG_Bls12_377::check_degree_is_too_large(p.degree(), powers.size()).is_err());
644    }
645}