halo2_proofs/poly/
domain.rs

1//! Contains utilities for performing polynomial arithmetic over an evaluation
2//! domain that is of a suitable size for the application.
3
4use crate::{
5    arithmetic::{best_fft, parallelize, FieldExt, Group},
6    plonk::Assigned,
7};
8
9use super::{Coeff, ExtendedLagrangeCoeff, LagrangeCoeff, Polynomial, Rotation};
10
11use group::ff::{BatchInvert, Field, PrimeField};
12
13use std::marker::PhantomData;
14
15/// This structure contains precomputed constants and other details needed for
16/// performing operations on an evaluation domain of size $2^k$ and an extended
17/// domain of size $2^{k} * j$ with $j \neq 0$.
18#[derive(Clone, Debug)]
19pub struct EvaluationDomain<G: Group> {
20    n: u64,
21    k: u32,
22    extended_k: u32,
23    omega: G::Scalar,
24    omega_inv: G::Scalar,
25    extended_omega: G::Scalar,
26    extended_omega_inv: G::Scalar,
27    g_coset: G::Scalar,
28    g_coset_inv: G::Scalar,
29    quotient_poly_degree: u64,
30    ifft_divisor: G::Scalar,
31    extended_ifft_divisor: G::Scalar,
32    t_evaluations: Vec<G::Scalar>,
33    barycentric_weight: G::Scalar,
34}
35
36impl<G: Group> EvaluationDomain<G> {
37    /// This constructs a new evaluation domain object based on the provided
38    /// values $j, k$.
39    pub fn new(j: u32, k: u32) -> Self {
40        // quotient_poly_degree * params.n - 1 is the degree of the quotient polynomial
41        let quotient_poly_degree = (j - 1) as u64;
42
43        // n = 2^k
44        let n = 1u64 << k;
45
46        // We need to work within an extended domain, not params.k but params.k + i
47        // for some integer i such that 2^(params.k + i) is sufficiently large to
48        // describe the quotient polynomial.
49        let mut extended_k = k;
50        while (1 << extended_k) < (n * quotient_poly_degree) {
51            extended_k += 1;
52        }
53
54        let mut extended_omega = G::Scalar::root_of_unity();
55
56        // Get extended_omega, the 2^{extended_k}'th root of unity
57        // The loop computes extended_omega = omega^{2 ^ (S - extended_k)}
58        // Notice that extended_omega ^ {2 ^ extended_k} = omega ^ {2^S} = 1.
59        for _ in extended_k..G::Scalar::S {
60            extended_omega = extended_omega.square();
61        }
62        let extended_omega = extended_omega;
63        let mut extended_omega_inv = extended_omega; // Inversion computed later
64
65        // Get omega, the 2^{k}'th root of unity (i.e. n'th root of unity)
66        // The loop computes omega = extended_omega ^ {2 ^ (extended_k - k)}
67        //           = (omega^{2 ^ (S - extended_k)})  ^ {2 ^ (extended_k - k)}
68        //           = omega ^ {2 ^ (S - k)}.
69        // Notice that omega ^ {2^k} = omega ^ {2^S} = 1.
70        let mut omega = extended_omega;
71        for _ in k..extended_k {
72            omega = omega.square();
73        }
74        let omega = omega;
75        let mut omega_inv = omega; // Inversion computed later
76
77        // We use zeta here because we know it generates a coset, and it's available
78        // already.
79        // The coset evaluation domain is:
80        // zeta {1, extended_omega, extended_omega^2, ..., extended_omega^{(2^extended_k) - 1}}
81        let g_coset = G::Scalar::ZETA;
82        let g_coset_inv = g_coset.square();
83
84        let mut t_evaluations = Vec::with_capacity(1 << (extended_k - k));
85        {
86            // Compute the evaluations of t(X) = X^n - 1 in the coset evaluation domain.
87            // We don't have to compute all of them, because it will repeat.
88            let orig = G::Scalar::ZETA.pow_vartime(&[n as u64, 0, 0, 0]);
89            let step = extended_omega.pow_vartime(&[n as u64, 0, 0, 0]);
90            let mut cur = orig;
91            loop {
92                t_evaluations.push(cur);
93                cur *= &step;
94                if cur == orig {
95                    break;
96                }
97            }
98            assert_eq!(t_evaluations.len(), 1 << (extended_k - k));
99
100            // Subtract 1 from each to give us t_evaluations[i] = t(zeta * extended_omega^i)
101            for coeff in &mut t_evaluations {
102                *coeff -= &G::Scalar::one();
103            }
104
105            // Invert, because we're dividing by this polynomial.
106            // We invert in a batch, below.
107        }
108
109        let mut ifft_divisor = G::Scalar::from(1 << k); // Inversion computed later
110        let mut extended_ifft_divisor = G::Scalar::from(1 << extended_k); // Inversion computed later
111
112        // The barycentric weight of 1 over the evaluation domain
113        // 1 / \prod_{i != 0} (1 - omega^i)
114        let mut barycentric_weight = G::Scalar::from(n); // Inversion computed later
115
116        // Compute batch inversion
117        t_evaluations
118            .iter_mut()
119            .chain(Some(&mut ifft_divisor))
120            .chain(Some(&mut extended_ifft_divisor))
121            .chain(Some(&mut barycentric_weight))
122            .chain(Some(&mut extended_omega_inv))
123            .chain(Some(&mut omega_inv))
124            .batch_invert();
125
126        EvaluationDomain {
127            n,
128            k,
129            extended_k,
130            omega,
131            omega_inv,
132            extended_omega,
133            extended_omega_inv,
134            g_coset,
135            g_coset_inv,
136            quotient_poly_degree,
137            ifft_divisor,
138            extended_ifft_divisor,
139            t_evaluations,
140            barycentric_weight,
141        }
142    }
143
144    /// Obtains a polynomial in Lagrange form when given a vector of Lagrange
145    /// coefficients of size `n`; panics if the provided vector is the wrong
146    /// length.
147    pub fn lagrange_from_vec(&self, values: Vec<G>) -> Polynomial<G, LagrangeCoeff> {
148        assert_eq!(values.len(), self.n as usize);
149
150        Polynomial {
151            values,
152            _marker: PhantomData,
153        }
154    }
155
156    /// Obtains a polynomial in coefficient form when given a vector of
157    /// coefficients of size `n`; panics if the provided vector is the wrong
158    /// length.
159    pub fn coeff_from_vec(&self, values: Vec<G>) -> Polynomial<G, Coeff> {
160        assert_eq!(values.len(), self.n as usize);
161
162        Polynomial {
163            values,
164            _marker: PhantomData,
165        }
166    }
167
168    /// Returns an empty (zero) polynomial in the coefficient basis
169    pub fn empty_coeff(&self) -> Polynomial<G, Coeff> {
170        Polynomial {
171            values: vec![G::group_zero(); self.n as usize],
172            _marker: PhantomData,
173        }
174    }
175
176    /// Returns an empty (zero) polynomial in the Lagrange coefficient basis
177    pub fn empty_lagrange(&self) -> Polynomial<G, LagrangeCoeff> {
178        Polynomial {
179            values: vec![G::group_zero(); self.n as usize],
180            _marker: PhantomData,
181        }
182    }
183
184    /// Returns an empty (zero) polynomial in the Lagrange coefficient basis, with
185    /// deferred inversions.
186    pub(crate) fn empty_lagrange_assigned(&self) -> Polynomial<Assigned<G>, LagrangeCoeff>
187    where
188        G: Field,
189    {
190        Polynomial {
191            values: vec![G::group_zero().into(); self.n as usize],
192            _marker: PhantomData,
193        }
194    }
195
196    /// Returns a constant polynomial in the Lagrange coefficient basis
197    pub fn constant_lagrange(&self, scalar: G) -> Polynomial<G, LagrangeCoeff> {
198        Polynomial {
199            values: vec![scalar; self.n as usize],
200            _marker: PhantomData,
201        }
202    }
203
204    /// Returns an empty (zero) polynomial in the extended Lagrange coefficient
205    /// basis
206    pub fn empty_extended(&self) -> Polynomial<G, ExtendedLagrangeCoeff> {
207        Polynomial {
208            values: vec![G::group_zero(); self.extended_len()],
209            _marker: PhantomData,
210        }
211    }
212
213    /// Returns a constant polynomial in the extended Lagrange coefficient
214    /// basis
215    pub fn constant_extended(&self, scalar: G) -> Polynomial<G, ExtendedLagrangeCoeff> {
216        Polynomial {
217            values: vec![scalar; self.extended_len()],
218            _marker: PhantomData,
219        }
220    }
221
222    /// This takes us from an n-length vector into the coefficient form.
223    ///
224    /// This function will panic if the provided vector is not the correct
225    /// length.
226    pub fn lagrange_to_coeff(&self, mut a: Polynomial<G, LagrangeCoeff>) -> Polynomial<G, Coeff> {
227        assert_eq!(a.values.len(), 1 << self.k);
228
229        // Perform inverse FFT to obtain the polynomial in coefficient form
230        Self::ifft(&mut a.values, self.omega_inv, self.k, self.ifft_divisor);
231
232        Polynomial {
233            values: a.values,
234            _marker: PhantomData,
235        }
236    }
237
238    /// This takes us from an n-length coefficient vector into a coset of the extended
239    /// evaluation domain, rotating by `rotation` if desired.
240    pub fn coeff_to_extended(
241        &self,
242        mut a: Polynomial<G, Coeff>,
243    ) -> Polynomial<G, ExtendedLagrangeCoeff> {
244        assert_eq!(a.values.len(), 1 << self.k);
245
246        self.distribute_powers_zeta(&mut a.values, true);
247        a.values.resize(self.extended_len(), G::group_zero());
248        best_fft(&mut a.values, self.extended_omega, self.extended_k);
249
250        Polynomial {
251            values: a.values,
252            _marker: PhantomData,
253        }
254    }
255
256    /// Rotate the extended domain polynomial over the original domain.
257    pub fn rotate_extended(
258        &self,
259        poly: &Polynomial<G, ExtendedLagrangeCoeff>,
260        rotation: Rotation,
261    ) -> Polynomial<G, ExtendedLagrangeCoeff> {
262        let new_rotation = ((1 << (self.extended_k - self.k)) * rotation.0.abs()) as usize;
263
264        let mut poly = poly.clone();
265
266        if rotation.0 >= 0 {
267            poly.values.rotate_left(new_rotation);
268        } else {
269            poly.values.rotate_right(new_rotation);
270        }
271
272        poly
273    }
274
275    /// This takes us from the extended evaluation domain and gets us the
276    /// quotient polynomial coefficients.
277    ///
278    /// This function will panic if the provided vector is not the correct
279    /// length.
280    // TODO/FIXME: caller should be responsible for truncating
281    pub fn extended_to_coeff(&self, mut a: Polynomial<G, ExtendedLagrangeCoeff>) -> Vec<G> {
282        assert_eq!(a.values.len(), self.extended_len());
283
284        // Inverse FFT
285        Self::ifft(
286            &mut a.values,
287            self.extended_omega_inv,
288            self.extended_k,
289            self.extended_ifft_divisor,
290        );
291
292        // Distribute powers to move from coset; opposite from the
293        // transformation we performed earlier.
294        self.distribute_powers_zeta(&mut a.values, false);
295
296        // Truncate it to match the size of the quotient polynomial; the
297        // evaluation domain might be slightly larger than necessary because
298        // it always lies on a power-of-two boundary.
299        a.values
300            .truncate((&self.n * self.quotient_poly_degree) as usize);
301
302        a.values
303    }
304
305    /// This divides the polynomial (in the extended domain) by the vanishing
306    /// polynomial of the $2^k$ size domain.
307    pub fn divide_by_vanishing_poly(
308        &self,
309        mut a: Polynomial<G, ExtendedLagrangeCoeff>,
310    ) -> Polynomial<G, ExtendedLagrangeCoeff> {
311        assert_eq!(a.values.len(), self.extended_len());
312
313        // Divide to obtain the quotient polynomial in the coset evaluation
314        // domain.
315        parallelize(&mut a.values, |h, mut index| {
316            for h in h {
317                h.group_scale(&self.t_evaluations[index % self.t_evaluations.len()]);
318                index += 1;
319            }
320        });
321
322        Polynomial {
323            values: a.values,
324            _marker: PhantomData,
325        }
326    }
327
328    /// Given a slice of group elements `[a_0, a_1, a_2, ...]`, this returns
329    /// `[a_0, [zeta]a_1, [zeta^2]a_2, a_3, [zeta]a_4, [zeta^2]a_5, a_6, ...]`,
330    /// where zeta is a cube root of unity in the multiplicative subgroup with
331    /// order (p - 1), i.e. zeta^3 = 1.
332    ///
333    /// `into_coset` should be set to `true` when moving into the coset,
334    /// and `false` when moving out. This toggles the choice of `zeta`.
335    fn distribute_powers_zeta(&self, a: &mut [G], into_coset: bool) {
336        let coset_powers = if into_coset {
337            [self.g_coset, self.g_coset_inv]
338        } else {
339            [self.g_coset_inv, self.g_coset]
340        };
341        parallelize(a, |a, mut index| {
342            for a in a {
343                // Distribute powers to move into/from coset
344                let i = index % (coset_powers.len() + 1);
345                if i != 0 {
346                    a.group_scale(&coset_powers[i - 1]);
347                }
348                index += 1;
349            }
350        });
351    }
352
353    fn ifft(a: &mut [G], omega_inv: G::Scalar, log_n: u32, divisor: G::Scalar) {
354        best_fft(a, omega_inv, log_n);
355        parallelize(a, |a, _| {
356            for a in a {
357                // Finish iFFT
358                a.group_scale(&divisor);
359            }
360        });
361    }
362
363    /// Get the size of the extended domain
364    pub fn extended_len(&self) -> usize {
365        1 << self.extended_k
366    }
367
368    /// Get $\omega$, the generator of the $2^k$ order multiplicative subgroup.
369    pub fn get_omega(&self) -> G::Scalar {
370        self.omega
371    }
372
373    /// Get $\omega^{-1}$, the inverse of the generator of the $2^k$ order
374    /// multiplicative subgroup.
375    pub fn get_omega_inv(&self) -> G::Scalar {
376        self.omega_inv
377    }
378
379    /// Get the generator of the extended domain's multiplicative subgroup.
380    pub fn get_extended_omega(&self) -> G::Scalar {
381        self.extended_omega
382    }
383
384    /// Multiplies a value by some power of $\omega$, essentially rotating over
385    /// the domain.
386    pub fn rotate_omega(&self, value: G::Scalar, rotation: Rotation) -> G::Scalar {
387        let mut point = value;
388        if rotation.0 >= 0 {
389            point *= &self.get_omega().pow_vartime(&[rotation.0 as u64]);
390        } else {
391            point *= &self
392                .get_omega_inv()
393                .pow_vartime(&[(rotation.0 as i64).abs() as u64]);
394        }
395        point
396    }
397
398    /// Computes evaluations (at the point `x`, where `xn = x^n`) of Lagrange
399    /// basis polynomials `l_i(X)` defined such that `l_i(omega^i) = 1` and
400    /// `l_i(omega^j) = 0` for all `j != i` at each provided rotation `i`.
401    ///
402    /// # Implementation
403    ///
404    /// The polynomial
405    ///     $$\prod_{j=0,j \neq i}^{n - 1} (X - \omega^j)$$
406    /// has a root at all points in the domain except $\omega^i$, where it evaluates to
407    ///     $$\prod_{j=0,j \neq i}^{n - 1} (\omega^i - \omega^j)$$
408    /// and so we divide that polynomial by this value to obtain $l_i(X)$. Since
409    ///     $$\prod_{j=0,j \neq i}^{n - 1} (X - \omega^j)
410    ///       = \frac{X^n - 1}{X - \omega^i}$$
411    /// then $l_i(x)$ for some $x$ is evaluated as
412    ///     $$\left(\frac{x^n - 1}{x - \omega^i}\right)
413    ///       \cdot \left(\frac{1}{\prod_{j=0,j \neq i}^{n - 1} (\omega^i - \omega^j)}\right).$$
414    /// We refer to
415    ///     $$1 \over \prod_{j=0,j \neq i}^{n - 1} (\omega^i - \omega^j)$$
416    /// as the barycentric weight of $\omega^i$.
417    ///
418    /// We know that for $i = 0$
419    ///     $$\frac{1}{\prod_{j=0,j \neq i}^{n - 1} (\omega^i - \omega^j)} = \frac{1}{n}.$$
420    ///
421    /// If we multiply $(1 / n)$ by $\omega^i$ then we obtain
422    ///     $$\frac{1}{\prod_{j=0,j \neq 0}^{n - 1} (\omega^i - \omega^j)}
423    ///       = \frac{1}{\prod_{j=0,j \neq i}^{n - 1} (\omega^i - \omega^j)}$$
424    /// which is the barycentric weight of $\omega^i$.
425    pub fn l_i_range<I: IntoIterator<Item = i32> + Clone>(
426        &self,
427        x: G::Scalar,
428        xn: G::Scalar,
429        rotations: I,
430    ) -> Vec<G::Scalar> {
431        let mut results;
432        {
433            let rotations = rotations.clone().into_iter();
434            results = Vec::with_capacity(rotations.size_hint().1.unwrap_or(0));
435            for rotation in rotations {
436                let rotation = Rotation(rotation);
437                let result = x - self.rotate_omega(G::Scalar::one(), rotation);
438                results.push(result);
439            }
440            results.iter_mut().batch_invert();
441        }
442
443        let common = (xn - G::Scalar::one()) * self.barycentric_weight;
444        for (rotation, result) in rotations.into_iter().zip(results.iter_mut()) {
445            let rotation = Rotation(rotation);
446            *result = self.rotate_omega(*result * common, rotation);
447        }
448
449        results
450    }
451
452    /// Gets the quotient polynomial's degree (as a multiple of n)
453    pub fn get_quotient_poly_degree(&self) -> usize {
454        self.quotient_poly_degree as usize
455    }
456
457    /// Obtain a pinned version of this evaluation domain; a structure with the
458    /// minimal parameters needed to determine the rest of the evaluation
459    /// domain.
460    pub fn pinned(&self) -> PinnedEvaluationDomain<'_, G> {
461        PinnedEvaluationDomain {
462            k: &self.k,
463            extended_k: &self.extended_k,
464            omega: &self.omega,
465        }
466    }
467}
468
469/// Represents the minimal parameters that determine an `EvaluationDomain`.
470#[allow(dead_code)]
471#[derive(Debug)]
472pub struct PinnedEvaluationDomain<'a, G: Group> {
473    k: &'a u32,
474    extended_k: &'a u32,
475    omega: &'a G::Scalar,
476}
477
478#[test]
479fn test_rotate() {
480    use rand_core::OsRng;
481
482    use crate::arithmetic::eval_polynomial;
483    use crate::pasta::pallas::Scalar;
484
485    let domain = EvaluationDomain::<Scalar>::new(1, 3);
486    let rng = OsRng;
487
488    let mut poly = domain.empty_lagrange();
489    assert_eq!(poly.len(), 8);
490    for value in poly.iter_mut() {
491        *value = Scalar::random(rng);
492    }
493
494    let poly_rotated_cur = poly.rotate(Rotation::cur());
495    let poly_rotated_next = poly.rotate(Rotation::next());
496    let poly_rotated_prev = poly.rotate(Rotation::prev());
497
498    let poly = domain.lagrange_to_coeff(poly);
499    let poly_rotated_cur = domain.lagrange_to_coeff(poly_rotated_cur);
500    let poly_rotated_next = domain.lagrange_to_coeff(poly_rotated_next);
501    let poly_rotated_prev = domain.lagrange_to_coeff(poly_rotated_prev);
502
503    let x = Scalar::random(rng);
504
505    assert_eq!(
506        eval_polynomial(&poly[..], x),
507        eval_polynomial(&poly_rotated_cur[..], x)
508    );
509    assert_eq!(
510        eval_polynomial(&poly[..], x * domain.omega),
511        eval_polynomial(&poly_rotated_next[..], x)
512    );
513    assert_eq!(
514        eval_polynomial(&poly[..], x * domain.omega_inv),
515        eval_polynomial(&poly_rotated_prev[..], x)
516    );
517}
518
519#[test]
520fn test_l_i() {
521    use rand_core::OsRng;
522
523    use crate::arithmetic::{eval_polynomial, lagrange_interpolate};
524    use crate::pasta::pallas::Scalar;
525    let domain = EvaluationDomain::<Scalar>::new(1, 3);
526
527    let mut l = vec![];
528    let mut points = vec![];
529    for i in 0..8 {
530        points.push(domain.omega.pow(&[i, 0, 0, 0]));
531    }
532    for i in 0..8 {
533        let mut l_i = vec![Scalar::zero(); 8];
534        l_i[i] = Scalar::one();
535        let l_i = lagrange_interpolate(&points[..], &l_i[..]);
536        l.push(l_i);
537    }
538
539    let x = Scalar::random(OsRng);
540    let xn = x.pow(&[8, 0, 0, 0]);
541
542    let evaluations = domain.l_i_range(x, xn, -7..=7);
543    for i in 0..8 {
544        assert_eq!(eval_polynomial(&l[i][..], x), evaluations[7 + i]);
545        assert_eq!(eval_polynomial(&l[(8 - i) % 8][..], x), evaluations[7 - i]);
546    }
547}