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