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