poly_commitment/
commitment.rs

1//! This module implements Dlog-based polynomial commitment schema.
2//! The following functionality is implemented
3//!
4//! 1. Commit to polynomial with its max degree
5//! 2. Open polynomial commitment batch at the given evaluation point and
6//!    scaling factor scalar producing the batched opening proof
7//! 3. Verify batch of batched opening proofs
8
9use ark_ec::{
10    models::short_weierstrass::Affine as SWJAffine, short_weierstrass::SWCurveConfig, AffineRepr,
11    CurveGroup, VariableBaseMSM,
12};
13use ark_ff::{BigInteger, Field, One, PrimeField, Zero};
14use ark_poly::univariate::DensePolynomial;
15use ark_serialize::{CanonicalDeserialize, CanonicalSerialize};
16use groupmap::{BWParameters, GroupMap};
17use mina_poseidon::{sponge::ScalarChallenge, FqSponge};
18use o1_utils::{field_helpers::product, ExtendedDensePolynomial as _};
19use rayon::prelude::*;
20use serde::{de::Visitor, Deserialize, Serialize};
21use serde_with::{
22    de::DeserializeAsWrap, ser::SerializeAsWrap, serde_as, DeserializeAs, SerializeAs,
23};
24use std::{
25    iter::Iterator,
26    marker::PhantomData,
27    ops::{Add, AddAssign, Sub},
28};
29
30/// Represent a polynomial commitment when the type is instantiated with a
31/// curve.
32///
33/// The structure also handles chunking, i.e. when we aim to handle polynomials
34/// whose degree is higher than the SRS size. For this reason, we do use a
35/// vector for the field `chunks`.
36///
37/// Note that the parameter `C` is not constrained to be a curve, therefore in
38/// some places in the code, `C` can refer to a scalar field element. For
39/// instance, `PolyComm<G::ScalarField>` is used to represent the evaluation of
40/// the polynomial bound by a specific commitment, at a particular evaluation
41/// point.
42#[serde_as]
43#[derive(Clone, Debug, Serialize, Deserialize, PartialEq, Eq)]
44#[serde(bound = "C: CanonicalDeserialize + CanonicalSerialize")]
45pub struct PolyComm<C> {
46    #[serde_as(as = "Vec<o1_utils::serialization::SerdeAs>")]
47    pub chunks: Vec<C>,
48}
49
50impl<C> PolyComm<C>
51where
52    C: CommitmentCurve,
53{
54    /// Multiplies each commitment chunk of f with powers of zeta^n
55    pub fn chunk_commitment(&self, zeta_n: C::ScalarField) -> Self {
56        let mut res = C::Group::zero();
57        // use Horner's to compute chunk[0] + z^n chunk[1] + z^2n chunk[2] + ...
58        // as ( chunk[-1] * z^n + chunk[-2] ) * z^n + chunk[-3]
59        // (https://en.wikipedia.org/wiki/Horner%27s_method)
60        for chunk in self.chunks.iter().rev() {
61            res *= zeta_n;
62            res.add_assign(chunk);
63        }
64
65        PolyComm {
66            chunks: vec![res.into_affine()],
67        }
68    }
69}
70
71impl<F> PolyComm<F>
72where
73    F: Field,
74{
75    /// Multiplies each blinding chunk of f with powers of zeta^n
76    pub fn chunk_blinding(&self, zeta_n: F) -> F {
77        let mut res = F::zero();
78        // use Horner's to compute chunk[0] + z^n chunk[1] + z^2n chunk[2] + ...
79        // as ( chunk[-1] * z^n + chunk[-2] ) * z^n + chunk[-3]
80        // (https://en.wikipedia.org/wiki/Horner%27s_method)
81        for chunk in self.chunks.iter().rev() {
82            res *= zeta_n;
83            res += chunk
84        }
85        res
86    }
87}
88
89impl<'a, G> IntoIterator for &'a PolyComm<G> {
90    type Item = &'a G;
91    type IntoIter = std::slice::Iter<'a, G>;
92
93    fn into_iter(self) -> Self::IntoIter {
94        self.chunks.iter()
95    }
96}
97
98/// A commitment to a polynomial with some blinding factors.
99#[derive(Clone, Debug, Serialize, Deserialize)]
100pub struct BlindedCommitment<G>
101where
102    G: CommitmentCurve,
103{
104    pub commitment: PolyComm<G>,
105    pub blinders: PolyComm<G::ScalarField>,
106}
107
108impl<T> PolyComm<T> {
109    pub fn new(chunks: Vec<T>) -> Self {
110        Self { chunks }
111    }
112}
113
114impl<T, U> SerializeAs<PolyComm<T>> for PolyComm<U>
115where
116    U: SerializeAs<T>,
117{
118    fn serialize_as<S>(source: &PolyComm<T>, serializer: S) -> Result<S::Ok, S::Error>
119    where
120        S: serde::Serializer,
121    {
122        serializer.collect_seq(
123            source
124                .chunks
125                .iter()
126                .map(|e| SerializeAsWrap::<T, U>::new(e)),
127        )
128    }
129}
130
131impl<'de, T, U> DeserializeAs<'de, PolyComm<T>> for PolyComm<U>
132where
133    U: DeserializeAs<'de, T>,
134{
135    fn deserialize_as<D>(deserializer: D) -> Result<PolyComm<T>, D::Error>
136    where
137        D: serde::Deserializer<'de>,
138    {
139        struct SeqVisitor<T, U> {
140            marker: PhantomData<(T, U)>,
141        }
142
143        impl<'de, T, U> Visitor<'de> for SeqVisitor<T, U>
144        where
145            U: DeserializeAs<'de, T>,
146        {
147            type Value = PolyComm<T>;
148
149            fn expecting(&self, formatter: &mut core::fmt::Formatter<'_>) -> core::fmt::Result {
150                formatter.write_str("a sequence")
151            }
152
153            fn visit_seq<A>(self, mut seq: A) -> Result<Self::Value, A::Error>
154            where
155                A: serde::de::SeqAccess<'de>,
156            {
157                #[allow(clippy::redundant_closure_call)]
158                let mut chunks = vec![];
159
160                while let Some(value) = seq
161                    .next_element()?
162                    .map(|v: DeserializeAsWrap<T, U>| v.into_inner())
163                {
164                    chunks.push(value);
165                }
166
167                Ok(PolyComm::new(chunks))
168            }
169        }
170
171        let visitor = SeqVisitor::<T, U> {
172            marker: PhantomData,
173        };
174        deserializer.deserialize_seq(visitor)
175    }
176}
177
178impl<A: Copy + Clone + CanonicalDeserialize + CanonicalSerialize> PolyComm<A> {
179    pub fn map<B, F>(&self, mut f: F) -> PolyComm<B>
180    where
181        F: FnMut(A) -> B,
182        B: CanonicalDeserialize + CanonicalSerialize,
183    {
184        let chunks = self.chunks.iter().map(|x| f(*x)).collect();
185        PolyComm::new(chunks)
186    }
187
188    /// Returns the number of chunks.
189    pub fn len(&self) -> usize {
190        self.chunks.len()
191    }
192
193    /// Returns `true` if the commitment is empty.
194    pub fn is_empty(&self) -> bool {
195        self.chunks.is_empty()
196    }
197
198    // TODO: if all callers end up calling unwrap, just call this zip_eq and
199    // panic here (and document the panic)
200    pub fn zip<B: Copy + CanonicalDeserialize + CanonicalSerialize>(
201        &self,
202        other: &PolyComm<B>,
203    ) -> Option<PolyComm<(A, B)>> {
204        if self.chunks.len() != other.chunks.len() {
205            return None;
206        }
207        let chunks = self
208            .chunks
209            .iter()
210            .zip(other.chunks.iter())
211            .map(|(x, y)| (*x, *y))
212            .collect();
213        Some(PolyComm::new(chunks))
214    }
215
216    /// Return only the first chunk
217    /// Getting this single value is relatively common in the codebase, even
218    /// though we should not do this, and abstract the chunks in the structure.
219    pub fn get_first_chunk(&self) -> A {
220        self.chunks[0]
221    }
222}
223
224/// Inside the circuit, we have a specialized scalar multiplication which
225/// computes either
226///
227/// ```ignore
228/// |g: G, x: G::ScalarField| g.scale(x + 2^n)
229/// ```
230///
231/// if the scalar field of G is greater than the size of the base field
232/// and
233///
234/// ```ignore
235/// |g: G, x: G::ScalarField| g.scale(2*x + 2^n)
236/// ```
237///
238/// otherwise. So, if we want to actually scale by `x`, we need to apply the
239/// inverse function of `|x| x + 2^n` (or of `|x| 2*x + 2^n` in the other case),
240/// before supplying the scalar to our in-circuit scalar-multiplication
241/// function. This computes that inverse function.
242/// Namely,
243///
244/// ```ignore
245/// |x: G::ScalarField| x - 2^n
246/// ```
247///
248/// when the scalar field is larger than the base field and
249///
250/// ```ignore
251/// |x: G::ScalarField| (x - 2^n) / 2
252/// ```
253///
254/// in the other case.
255pub fn shift_scalar<G: AffineRepr>(x: G::ScalarField) -> G::ScalarField
256where
257    G::BaseField: PrimeField,
258{
259    let n1 = <G::ScalarField as PrimeField>::MODULUS;
260    let n2 = <G::ScalarField as PrimeField>::BigInt::from_bits_le(
261        &<G::BaseField as PrimeField>::MODULUS.to_bits_le()[..],
262    );
263    let two: G::ScalarField = (2u64).into();
264    let two_pow = two.pow([<G::ScalarField as PrimeField>::MODULUS_BIT_SIZE as u64]);
265    if n1 < n2 {
266        (x - (two_pow + G::ScalarField::one())) / two
267    } else {
268        x - two_pow
269    }
270}
271
272impl<'a, 'b, C: AffineRepr> Add<&'a PolyComm<C>> for &'b PolyComm<C> {
273    type Output = PolyComm<C>;
274
275    fn add(self, other: &'a PolyComm<C>) -> PolyComm<C> {
276        let mut chunks = vec![];
277        let n1 = self.chunks.len();
278        let n2 = other.chunks.len();
279        for i in 0..std::cmp::max(n1, n2) {
280            let pt = if i < n1 && i < n2 {
281                (self.chunks[i] + other.chunks[i]).into_affine()
282            } else if i < n1 {
283                self.chunks[i]
284            } else {
285                other.chunks[i]
286            };
287            chunks.push(pt);
288        }
289        PolyComm::new(chunks)
290    }
291}
292
293impl<'a, 'b, C: AffineRepr + Sub<Output = C::Group>> Sub<&'a PolyComm<C>> for &'b PolyComm<C> {
294    type Output = PolyComm<C>;
295
296    fn sub(self, other: &'a PolyComm<C>) -> PolyComm<C> {
297        let mut chunks = vec![];
298        let n1 = self.chunks.len();
299        let n2 = other.chunks.len();
300        for i in 0..std::cmp::max(n1, n2) {
301            let pt = if i < n1 && i < n2 {
302                (self.chunks[i] - other.chunks[i]).into_affine()
303            } else if i < n1 {
304                self.chunks[i]
305            } else {
306                other.chunks[i]
307            };
308            chunks.push(pt);
309        }
310        PolyComm::new(chunks)
311    }
312}
313
314impl<C: AffineRepr> PolyComm<C> {
315    pub fn scale(&self, c: C::ScalarField) -> PolyComm<C> {
316        PolyComm {
317            chunks: self.chunks.iter().map(|g| g.mul(c).into_affine()).collect(),
318        }
319    }
320
321    /// Performs a multi-scalar multiplication between scalars `elm` and
322    /// commitments `com`. If both are empty, returns a commitment of length 1
323    /// containing the point at infinity.
324    ///
325    /// ## Panics
326    ///
327    /// Panics if `com` and `elm` are not of the same size.
328    pub fn multi_scalar_mul(com: &[&PolyComm<C>], elm: &[C::ScalarField]) -> Self {
329        assert_eq!(com.len(), elm.len());
330
331        if com.is_empty() || elm.is_empty() {
332            return Self::new(vec![C::zero()]);
333        }
334
335        let all_scalars: Vec<_> = elm.iter().map(|s| s.into_bigint()).collect();
336
337        let elems_size = Iterator::max(com.iter().map(|c| c.chunks.len())).unwrap();
338
339        let chunks = (0..elems_size)
340            .map(|chunk| {
341                let (points, scalars): (Vec<_>, Vec<_>) = com
342                    .iter()
343                    .zip(&all_scalars)
344                    // get rid of scalars that don't have an associated chunk
345                    .filter_map(|(com, scalar)| com.chunks.get(chunk).map(|c| (c, scalar)))
346                    .unzip();
347
348                // Splitting into 2 chunks seems optimal; but in
349                // practice elems_size is almost always 1
350                //
351                // (see the comment to the `benchmark_msm_parallel_vesta` MSM benchmark)
352                let subchunk_size = std::cmp::max(points.len() / 2, 1);
353
354                points
355                    .into_par_iter()
356                    .chunks(subchunk_size)
357                    .zip(scalars.into_par_iter().chunks(subchunk_size))
358                    .map(|(psc, ssc)| C::Group::msm_bigint(&psc, &ssc).into_affine())
359                    .reduce(C::zero, |x, y| (x + y).into())
360            })
361            .collect();
362
363        Self::new(chunks)
364    }
365}
366
367/// Returns `(1 + chal[-1] x)(1 + chal[-2] x^2)(1 + chal[-3] x^4) ...`. It's
368/// "step 8: Define the univariate polynomial" of
369/// appendix A.2 of <https://eprint.iacr.org/2020/499>
370pub fn b_poly<F: Field>(chals: &[F], x: F) -> F {
371    let k = chals.len();
372
373    let mut pow_twos = vec![x];
374
375    for i in 1..k {
376        pow_twos.push(pow_twos[i - 1].square());
377    }
378
379    product((0..k).map(|i| (F::one() + (chals[i] * pow_twos[k - 1 - i]))))
380}
381
382pub fn b_poly_coefficients<F: Field>(chals: &[F]) -> Vec<F> {
383    let rounds = chals.len();
384    let s_length = 1 << rounds;
385    let mut s = vec![F::one(); s_length];
386    let mut k: usize = 0;
387    let mut pow: usize = 1;
388    for i in 1..s_length {
389        k += if i == pow { 1 } else { 0 };
390        pow <<= if i == pow { 1 } else { 0 };
391        s[i] = s[i - (pow >> 1)] * chals[rounds - 1 - (k - 1)];
392    }
393    s
394}
395
396pub fn squeeze_prechallenge<Fq: Field, G, Fr: Field, EFqSponge: FqSponge<Fq, G, Fr>>(
397    sponge: &mut EFqSponge,
398) -> ScalarChallenge<Fr> {
399    ScalarChallenge(sponge.challenge())
400}
401
402pub fn squeeze_challenge<Fq: Field, G, Fr: PrimeField, EFqSponge: FqSponge<Fq, G, Fr>>(
403    endo_r: &Fr,
404    sponge: &mut EFqSponge,
405) -> Fr {
406    squeeze_prechallenge(sponge).to_field(endo_r)
407}
408
409pub fn absorb_commitment<Fq: Field, G: Clone, Fr: PrimeField, EFqSponge: FqSponge<Fq, G, Fr>>(
410    sponge: &mut EFqSponge,
411    commitment: &PolyComm<G>,
412) {
413    sponge.absorb_g(&commitment.chunks);
414}
415
416/// A useful trait extending AffineRepr for commitments.
417/// Unfortunately, we can't specify that `AffineRepr<BaseField : PrimeField>`,
418/// so usage of this traits must manually bind `G::BaseField: PrimeField`.
419pub trait CommitmentCurve: AffineRepr + Sub<Output = Self::Group> {
420    type Params: SWCurveConfig;
421    type Map: GroupMap<Self::BaseField>;
422
423    fn to_coordinates(&self) -> Option<(Self::BaseField, Self::BaseField)>;
424    fn of_coordinates(x: Self::BaseField, y: Self::BaseField) -> Self;
425}
426
427/// A trait extending CommitmentCurve for endomorphisms.
428/// Unfortunately, we can't specify that `AffineRepr<BaseField : PrimeField>`,
429/// so usage of this traits must manually bind `G::BaseField: PrimeField`.
430pub trait EndoCurve: CommitmentCurve {
431    /// Combine where x1 = one
432    fn combine_one(g1: &[Self], g2: &[Self], x2: Self::ScalarField) -> Vec<Self> {
433        crate::combine::window_combine(g1, g2, Self::ScalarField::one(), x2)
434    }
435
436    /// Combine where x1 = one
437    fn combine_one_endo(
438        endo_r: Self::ScalarField,
439        _endo_q: Self::BaseField,
440        g1: &[Self],
441        g2: &[Self],
442        x2: ScalarChallenge<Self::ScalarField>,
443    ) -> Vec<Self> {
444        crate::combine::window_combine(g1, g2, Self::ScalarField::one(), x2.to_field(&endo_r))
445    }
446
447    fn combine(
448        g1: &[Self],
449        g2: &[Self],
450        x1: Self::ScalarField,
451        x2: Self::ScalarField,
452    ) -> Vec<Self> {
453        crate::combine::window_combine(g1, g2, x1, x2)
454    }
455}
456
457impl<P: SWCurveConfig + Clone> CommitmentCurve for SWJAffine<P> {
458    type Params = P;
459    type Map = BWParameters<P>;
460
461    fn to_coordinates(&self) -> Option<(Self::BaseField, Self::BaseField)> {
462        if self.infinity {
463            None
464        } else {
465            Some((self.x, self.y))
466        }
467    }
468
469    fn of_coordinates(x: P::BaseField, y: P::BaseField) -> SWJAffine<P> {
470        SWJAffine::<P>::new_unchecked(x, y)
471    }
472}
473
474impl<P: SWCurveConfig + Clone> EndoCurve for SWJAffine<P> {
475    fn combine_one(g1: &[Self], g2: &[Self], x2: Self::ScalarField) -> Vec<Self> {
476        crate::combine::affine_window_combine_one(g1, g2, x2)
477    }
478
479    fn combine_one_endo(
480        _endo_r: Self::ScalarField,
481        endo_q: Self::BaseField,
482        g1: &[Self],
483        g2: &[Self],
484        x2: ScalarChallenge<Self::ScalarField>,
485    ) -> Vec<Self> {
486        crate::combine::affine_window_combine_one_endo(endo_q, g1, g2, x2)
487    }
488
489    fn combine(
490        g1: &[Self],
491        g2: &[Self],
492        x1: Self::ScalarField,
493        x2: Self::ScalarField,
494    ) -> Vec<Self> {
495        crate::combine::affine_window_combine(g1, g2, x1, x2)
496    }
497}
498
499/// Computes the linearization of the evaluations of a (potentially
500/// split) polynomial.
501///
502/// Each polynomial in `polys` is represented by a matrix where the
503/// rows correspond to evaluated points, and the columns represent
504/// potential segments (if a polynomial was split in several parts).
505///
506/// Elements in `evaluation_points` are several discrete points on which
507/// we evaluate polynomials, e.g. `[zeta,zeta*w]`. See `PointEvaluations`.
508///
509/// Note that if one of the polynomial comes specified with a degree
510/// bound, the evaluation for the last segment is potentially shifted
511/// to meet the proof.
512///
513/// Returns
514/// ```text
515/// |polys| |segments[k]|
516///    Σ         Σ         polyscale^{k*n+i} (Σ polys[k][j][i] * evalscale^j)
517///  k = 1     i = 1                          j
518/// ```
519#[allow(clippy::type_complexity)]
520pub fn combined_inner_product<F: PrimeField>(
521    polyscale: &F,
522    evalscale: &F,
523    // TODO(mimoo): needs a type that can get you evaluations or segments
524    polys: &[Vec<Vec<F>>],
525) -> F {
526    // final combined evaluation result
527    let mut res = F::zero();
528    // polyscale^i
529    let mut polyscale_i = F::one();
530
531    for evals_tr in polys.iter().filter(|evals_tr| !evals_tr[0].is_empty()) {
532        // Transpose the evaluations.
533        // evals[i] = {evals_tr[j][i]}_j now corresponds to a column in
534        // evals_tr, representing a segment.
535        let evals: Vec<_> = (0..evals_tr[0].len())
536            .map(|i| evals_tr.iter().map(|v| v[i]).collect::<Vec<_>>())
537            .collect();
538
539        // Iterating over the polynomial segments.
540        // Each segment gets its own polyscale^i, each segment element j is
541        // multiplied by evalscale^j. Given that polyscale_i = polyscale^i0 at
542        // this point, after this loop we have:
543        //
544        //    res += Σ polyscale^{i0+i} ( Σ evals_tr[j][i] * evalscale^j )
545        //           i                    j
546        //
547        for eval in &evals {
548            // p_i(evalscale)
549            let term = DensePolynomial::<F>::eval_polynomial(eval, *evalscale);
550            res += &(polyscale_i * term);
551            polyscale_i *= polyscale;
552        }
553    }
554    res
555}
556
557/// Contains the evaluation of a polynomial commitment at a set of points.
558pub struct Evaluation<G>
559where
560    G: AffineRepr,
561{
562    /// The commitment of the polynomial being evaluated.
563    /// Note that PolyComm contains a vector of commitments, which handles the
564    /// case when chunking is used, i.e. when the polynomial degree is higher
565    /// than the SRS size.
566    pub commitment: PolyComm<G>,
567
568    /// Contains an evaluation table. For instance, for vanilla PlonK, it
569    /// would be a vector of (chunked) evaluations at ζ and ζω.
570    /// The outer vector would be the evaluations at the different points (e.g.
571    /// ζ and ζω for vanilla PlonK) and the inner vector would be the chunks of
572    /// the polynomial.
573    pub evaluations: Vec<Vec<G::ScalarField>>,
574}
575
576/// Contains the batch evaluation
577pub struct BatchEvaluationProof<'a, G, EFqSponge, OpeningProof>
578where
579    G: AffineRepr,
580    EFqSponge: FqSponge<G::BaseField, G, G::ScalarField>,
581{
582    /// Sponge used to coin and absorb values and simulate
583    /// non-interactivity using the Fiat-Shamir transformation.
584    pub sponge: EFqSponge,
585    /// A list of evaluations, each supposed to correspond to a different
586    /// polynomial.
587    pub evaluations: Vec<Evaluation<G>>,
588    /// The actual evaluation points. Each field `evaluations` of each structure
589    /// of `Evaluation` should have the same (outer) length.
590    pub evaluation_points: Vec<G::ScalarField>,
591    /// A challenge to combine polynomials. Powers of this point will be used,
592    /// hence the name.
593    pub polyscale: G::ScalarField,
594    /// A challenge to aggregate multiple evaluation points.
595    pub evalscale: G::ScalarField,
596    /// The opening proof.
597    pub opening: &'a OpeningProof,
598    pub combined_inner_product: G::ScalarField,
599}
600
601/// This function populates the parameters `scalars` and `points`.
602/// It iterates over the evaluations and adds each commitment to the
603/// vector `points`.
604/// The parameter `scalars` is populated with the values:
605/// `rand_base * polyscale^i` for each commitment.
606/// For instance, if we have 3 commitments, the `scalars` vector will
607/// contain the values
608/// ```text
609/// [rand_base, rand_base * polyscale, rand_base * polyscale^2]`
610/// ```
611/// and the vector `points` will contain the commitments.
612///
613/// Note that the function skips the commitments that are empty.
614///
615/// If more than one commitment is present in a single evaluation (i.e. if
616/// `elems` is larger than one), it means that probably chunking was used (i.e.
617/// it is a commitment to a polynomial larger than the SRS).
618pub fn combine_commitments<G: CommitmentCurve>(
619    evaluations: &[Evaluation<G>],
620    scalars: &mut Vec<G::ScalarField>,
621    points: &mut Vec<G>,
622    polyscale: G::ScalarField,
623    rand_base: G::ScalarField,
624) {
625    // will contain the power of polyscale
626    let mut polyscale_i = G::ScalarField::one();
627
628    for Evaluation { commitment, .. } in evaluations.iter().filter(|x| !x.commitment.is_empty()) {
629        // iterating over the polynomial segments
630        for comm_ch in &commitment.chunks {
631            scalars.push(rand_base * polyscale_i);
632            points.push(*comm_ch);
633
634            // compute next power of polyscale
635            polyscale_i *= polyscale;
636        }
637    }
638}
639
640#[cfg(feature = "ocaml_types")]
641pub mod caml {
642    // polynomial commitment
643    use super::PolyComm;
644    use ark_ec::AffineRepr;
645
646    #[derive(Clone, Debug, ocaml::IntoValue, ocaml::FromValue, ocaml_gen::Struct)]
647    pub struct CamlPolyComm<CamlG> {
648        pub unshifted: Vec<CamlG>,
649        pub shifted: Option<CamlG>,
650    }
651
652    // handy conversions
653
654    impl<G, CamlG> From<PolyComm<G>> for CamlPolyComm<CamlG>
655    where
656        G: AffineRepr,
657        CamlG: From<G>,
658    {
659        fn from(polycomm: PolyComm<G>) -> Self {
660            Self {
661                unshifted: polycomm.chunks.into_iter().map(CamlG::from).collect(),
662                shifted: None,
663            }
664        }
665    }
666
667    impl<'a, G, CamlG> From<&'a PolyComm<G>> for CamlPolyComm<CamlG>
668    where
669        G: AffineRepr,
670        CamlG: From<G> + From<&'a G>,
671    {
672        fn from(polycomm: &'a PolyComm<G>) -> Self {
673            Self {
674                unshifted: polycomm.chunks.iter().map(Into::<CamlG>::into).collect(),
675                shifted: None,
676            }
677        }
678    }
679
680    impl<G, CamlG> From<CamlPolyComm<CamlG>> for PolyComm<G>
681    where
682        G: AffineRepr + From<CamlG>,
683    {
684        fn from(camlpolycomm: CamlPolyComm<CamlG>) -> PolyComm<G> {
685            assert!(
686                camlpolycomm.shifted.is_none(),
687                "mina#14628: Shifted commitments are deprecated and must not be used"
688            );
689            PolyComm {
690                chunks: camlpolycomm
691                    .unshifted
692                    .into_iter()
693                    .map(Into::<G>::into)
694                    .collect(),
695            }
696        }
697    }
698
699    impl<'a, G, CamlG> From<&'a CamlPolyComm<CamlG>> for PolyComm<G>
700    where
701        G: AffineRepr + From<&'a CamlG> + From<CamlG>,
702    {
703        fn from(camlpolycomm: &'a CamlPolyComm<CamlG>) -> PolyComm<G> {
704            assert!(
705                camlpolycomm.shifted.is_none(),
706                "mina#14628: Shifted commitments are deprecated and must not be used"
707            );
708            PolyComm {
709                //FIXME something with as_ref()
710                chunks: camlpolycomm.unshifted.iter().map(Into::into).collect(),
711            }
712        }
713    }
714}