jf_relation/gadgets/ecc/
mod.rs

1// Copyright (c) 2022 Espresso Systems (espressosys.com)
2// This file is part of the Jellyfish library.
3
4// You should have received a copy of the MIT License
5// along with the Jellyfish library. If not, see <https://mit-license.org/>.
6
7//! Elliptic curve related gates and gadgets. Including both native and
8//! non-native fields.
9
10use super::{from_emulated_field, EmulationConfig, SerializableEmulatedStruct};
11use crate::{gates::*, BoolVar, Circuit, CircuitError, PlonkCircuit, Variable};
12use ark_ec::{
13    twisted_edwards::{Affine, Projective, TECurveConfig as Config},
14    AffineRepr, CurveConfig, CurveGroup, ScalarMul,
15};
16use ark_ff::PrimeField;
17use ark_std::{borrow::ToOwned, boxed::Box, string::ToString, vec, vec::Vec};
18use core::marker::PhantomData;
19
20mod conversion;
21pub mod emulated;
22mod glv;
23mod msm;
24pub use conversion::*;
25pub use msm::*;
26
27#[derive(Debug, Copy, Clone, Eq, PartialEq)]
28/// An elliptic curve point in twisted Edwards affine form (x, y).
29pub struct TEPoint<F: PrimeField>(F, F);
30
31impl<F: PrimeField> Default for TEPoint<F> {
32    fn default() -> Self {
33        Self(F::zero(), F::one())
34    }
35}
36
37impl<F, P> From<Affine<P>> for TEPoint<F>
38where
39    F: PrimeField,
40    P: Config<BaseField = F>,
41{
42    fn from(p: Affine<P>) -> Self {
43        if p.is_zero() {
44            // separately treat point of infinity since maliciously constructed Affine
45            // could be (0,0,true) which is still a valid infinity point but would result in
46            // `Point(0, 0)` which might lead to problems as seen in precedence:
47            // https://cryptosubtlety.medium.com/00-8d4adcf4d255
48            let inf = Affine::<P>::zero();
49            TEPoint(inf.x, inf.y)
50        } else {
51            TEPoint(p.x, p.y)
52        }
53    }
54}
55
56impl<F: PrimeField> TEPoint<F> {
57    /// Get the x coordinate of the point.
58    pub fn get_x(&self) -> F {
59        self.0
60    }
61
62    /// Get the y coordinate of the point.
63    pub fn get_y(&self) -> F {
64        self.1
65    }
66
67    /// The inverse point for the edward form.
68    pub fn inverse(&self) -> Self {
69        Self(-self.0, self.1)
70    }
71}
72
73impl<F, P> From<Projective<P>> for TEPoint<F>
74where
75    F: PrimeField,
76    P: Config<BaseField = F>,
77{
78    fn from(p: Projective<P>) -> Self {
79        let affine_repr = p.into_affine();
80        TEPoint(affine_repr.x, affine_repr.y)
81    }
82}
83
84impl<F, P> From<TEPoint<F>> for Affine<P>
85where
86    F: PrimeField,
87    P: Config<BaseField = F>,
88{
89    fn from(p: TEPoint<F>) -> Self {
90        Self::new(p.0, p.1)
91    }
92}
93
94impl<F, P> From<TEPoint<F>> for Projective<P>
95where
96    F: PrimeField,
97    P: Config<BaseField = F>,
98{
99    fn from(p: TEPoint<F>) -> Self {
100        let affine_point: Affine<P> = p.into();
101        affine_point.into_group()
102    }
103}
104
105impl<E, F> SerializableEmulatedStruct<F> for TEPoint<E>
106where
107    E: EmulationConfig<F>,
108    F: PrimeField,
109{
110    fn serialize_to_native_elements(&self) -> Vec<F> {
111        let mut result = from_emulated_field(self.0);
112        result.extend(from_emulated_field(self.1));
113        result
114    }
115}
116
117#[derive(Debug, Copy, Clone, Eq, PartialEq)]
118/// Represent variable of an EC point.
119pub struct PointVariable(Variable, Variable);
120
121impl PointVariable {
122    /// Get the variable representing the x coordinate of the point.
123    pub fn get_x(&self) -> Variable {
124        self.0
125    }
126
127    /// Get the variable representing the y coordinate of the point.
128    pub fn get_y(&self) -> Variable {
129        self.1
130    }
131}
132
133// ECC related gates
134impl<F: PrimeField> PlonkCircuit<F> {
135    /// Return the witness point for the circuit
136    pub fn point_witness(&self, point_var: &PointVariable) -> Result<TEPoint<F>, CircuitError> {
137        self.check_point_var_bound(point_var)?;
138        let x = self.witness(point_var.0)?;
139        let y = self.witness(point_var.1)?;
140        Ok(TEPoint(x, y))
141    }
142
143    /// Add a new EC point (as witness) to the circuit
144    pub fn create_point_variable(
145        &mut self,
146        point: TEPoint<F>,
147    ) -> Result<PointVariable, CircuitError> {
148        let x_var = self.create_variable(point.0)?;
149        let y_var = self.create_variable(point.1)?;
150        Ok(PointVariable(x_var, y_var))
151    }
152
153    /// Add a new EC point (as a constant) to the circuit
154    pub fn create_constant_point_variable(
155        &mut self,
156        point: TEPoint<F>,
157    ) -> Result<PointVariable, CircuitError> {
158        let x_var = self.create_constant_variable(point.0)?;
159        let y_var = self.create_constant_variable(point.1)?;
160        Ok(PointVariable(x_var, y_var))
161    }
162
163    /// Add a new EC point (as public input) to the circuit
164    pub fn create_public_point_variable(
165        &mut self,
166        point: TEPoint<F>,
167    ) -> Result<PointVariable, CircuitError> {
168        let x_var = self.create_public_variable(point.0)?;
169        let y_var = self.create_public_variable(point.1)?;
170        Ok(PointVariable(x_var, y_var))
171    }
172
173    /// Obtain a point variable of the conditional selection from 4 point
174    /// candidates, (b0, b1) are two boolean variables indicating the choice
175    /// P_b0+2b1 where P0 = (0, 1) the neutral point, P1, P2, P3 are input
176    /// parameters.
177    /// A bad PointVariable would be returned if (b0, b1) are not boolean
178    /// variables, that would ultimately failed to build a correct circuit.
179    fn quaternary_point_select<P: Config<BaseField = F>>(
180        &mut self,
181        b0: BoolVar,
182        b1: BoolVar,
183        point1: &TEPoint<F>,
184        point2: &TEPoint<F>,
185        point3: &TEPoint<F>,
186    ) -> Result<PointVariable, CircuitError> {
187        self.check_var_bound(b0.into())?;
188        self.check_var_bound(b1.into())?;
189
190        let selected_point = {
191            let selected = match (
192                self.witness(b0.into())? == F::one(),
193                self.witness(b1.into())? == F::one(),
194            ) {
195                (false, false) => TEPoint::from(Affine::<P>::zero()),
196                (true, false) => point1.to_owned(),
197                (false, true) => point2.to_owned(),
198                (true, true) => point3.to_owned(),
199            };
200            // create new point with the same (x, y) coordinates
201            self.create_point_variable(selected)?
202        };
203        let wire_vars_x = [b0.into(), b1.into(), 0, 0, selected_point.0];
204        self.insert_gate(
205            &wire_vars_x,
206            Box::new(QuaternaryPointSelectXGate {
207                x1: point1.0,
208                x2: point2.0,
209                x3: point3.0,
210            }),
211        )?;
212        let wire_vars_y = [b0.into(), b1.into(), 0, 0, selected_point.1];
213        self.insert_gate(
214            &wire_vars_y,
215            Box::new(QuaternaryPointSelectYGate {
216                y1: point1.1,
217                y2: point2.1,
218                y3: point3.1,
219            }),
220        )?;
221
222        Ok(selected_point)
223    }
224
225    /// Obtain a point variable of the conditional selection from 2 point
226    /// variables. `b` is a boolean variable that indicates selection of P_b
227    /// from (P0, P1).
228    /// Return error if invalid input parameters are provided.
229    pub fn binary_point_vars_select(
230        &mut self,
231        b: BoolVar,
232        point0: &PointVariable,
233        point1: &PointVariable,
234    ) -> Result<PointVariable, CircuitError> {
235        self.check_var_bound(b.into())?;
236        self.check_point_var_bound(point0)?;
237        self.check_point_var_bound(point1)?;
238
239        let selected_x = self.conditional_select(b, point0.0, point1.0)?;
240        let selected_y = self.conditional_select(b, point0.1, point1.1)?;
241        Ok(PointVariable(selected_x, selected_y))
242    }
243
244    /// Constrain two point variables to be the same.
245    /// Return error if the input point variables are invalid.
246    pub fn enforce_point_equal(
247        &mut self,
248        point0: &PointVariable,
249        point1: &PointVariable,
250    ) -> Result<(), CircuitError> {
251        self.check_point_var_bound(point0)?;
252        self.check_point_var_bound(point1)?;
253        self.enforce_equal(point0.0, point1.0)?;
254        self.enforce_equal(point0.1, point1.1)?;
255        Ok(())
256    }
257
258    /// Obtain a bool variable representing whether two point variables are
259    /// equal. Return error if point variables are invalid.
260    pub fn is_point_equal(
261        &mut self,
262        point0: &PointVariable,
263        point1: &PointVariable,
264    ) -> Result<BoolVar, CircuitError> {
265        self.check_point_var_bound(point0)?;
266        self.check_point_var_bound(point1)?;
267        let x_eq = self.is_equal(point0.0, point1.0)?;
268        let y_eq = self.is_equal(point0.1, point1.1)?;
269        self.logic_and(x_eq, y_eq)
270    }
271}
272
273impl<F: PrimeField> PlonkCircuit<F> {
274    /// Inverse a point variable
275    pub fn inverse_point(
276        &mut self,
277        point_var: &PointVariable,
278    ) -> Result<PointVariable, CircuitError> {
279        let x_neg = self.sub(self.zero(), point_var.0)?;
280        Ok(PointVariable(x_neg, point_var.1))
281    }
282
283    /// Return the point variable for the infinity point
284    /// in the TE form: (0, 1)
285    pub fn neutral_point_variable(&self) -> PointVariable {
286        PointVariable(self.zero(), self.one())
287    }
288
289    /// Constrain a point variable to be a neutral point (0, 1) if
290    /// `expected_neutral` == 1
291    /// Constrain a point variable to be NOT a neutral point if
292    /// `expected_neutral` == 0
293    /// Note that `expected_neutral` is already constrained as a bool variable
294    fn neutral_point_gate(
295        &mut self,
296        point_var: &PointVariable,
297        expected_neutral: BoolVar,
298    ) -> Result<(), CircuitError> {
299        self.check_point_var_bound(point_var)?;
300        self.check_var_bound(expected_neutral.into())?;
301
302        // constraint 1: b_x = is_equal(x, 0);
303        let b_x = self.is_equal(point_var.0, self.zero())?;
304        // constraint 2: b_y = is_equal(y, 1);
305        let b_y = self.is_equal(point_var.1, self.one())?;
306        // constraint 3: b = b_x * b_y;
307        self.mul_gate(b_x.into(), b_y.into(), expected_neutral.into())?;
308        Ok(())
309    }
310
311    /// Obtain a boolean variable indicating whether a point is the neutral
312    /// point (0, 1) Return variable with value 1 if it is, or 0 otherwise
313    /// Return error if input variables are invalid
314    pub fn is_neutral_point<P>(
315        &mut self,
316        point_var: &PointVariable,
317    ) -> Result<BoolVar, CircuitError>
318    where
319        P: Config<BaseField = F>,
320    {
321        self.check_point_var_bound(point_var)?;
322
323        let b = {
324            if self.point_witness(point_var)? == TEPoint::from(Affine::<P>::zero()) {
325                self.create_boolean_variable(true)?
326            } else {
327                self.create_boolean_variable(false)?
328            }
329        };
330
331        self.neutral_point_gate(point_var, b)?;
332        Ok(b)
333    }
334    /// Constrain a point to be on certain curve, namely its coordinates satisfy
335    /// the curve equation, which is curve-dependent. Currently we only support
336    /// checks of a `Affine::<P>` over a base field which is the bls12-381
337    /// scalar field
338    ///
339    /// Returns error if input variables are invalid
340    pub fn enforce_on_curve<P: Config<BaseField = F>>(
341        &mut self,
342        point_var: &PointVariable,
343    ) -> Result<(), CircuitError> {
344        self.check_point_var_bound(point_var)?;
345
346        let (x, y) = (point_var.0, point_var.1);
347        let wire_vars = [x, x, y, y, 1];
348        self.insert_gate(
349            &wire_vars,
350            Box::new(EdwardsCurveEquationGate::<P> {
351                _phantom: PhantomData,
352            }),
353        )?;
354        Ok(())
355    }
356
357    /// Constrain variable `point_c` to be the point addition of `point_a` and
358    /// `point_b` over an elliptic curve.
359    /// Currently only supports Affine::<P> addition.
360    ///
361    /// Returns error if the input variables are invalid.
362    fn ecc_add_gate<P: Config<BaseField = F>>(
363        &mut self,
364        point_a: &PointVariable,
365        point_b: &PointVariable,
366        point_c: &PointVariable,
367    ) -> Result<(), CircuitError> {
368        self.check_point_var_bound(point_a)?;
369        self.check_point_var_bound(point_b)?;
370        self.check_point_var_bound(point_c)?;
371
372        let (x_1, y_1) = (point_a.0, point_a.1);
373        let (x_2, y_2) = (point_b.0, point_b.1);
374        let (x_3, y_3) = (point_c.0, point_c.1);
375
376        let x_coordinate_wire_vars = [x_1, y_2, x_2, y_1, x_3];
377        self.insert_gate(
378            &x_coordinate_wire_vars,
379            Box::new(CurvePointXAdditionGate::<P> {
380                _phantom: PhantomData,
381            }),
382        )?;
383        let y_coordinate_wire_vars = [x_1, x_2, y_1, y_2, y_3];
384        self.insert_gate(
385            &y_coordinate_wire_vars,
386            Box::new(CurvePointYAdditionGate::<P> {
387                _phantom: PhantomData,
388            }),
389        )?;
390        Ok(())
391    }
392
393    /// Obtain a variable to the point addition result of `point_a` + `point_b`
394    /// where "+" is the group operation over an elliptic curve.
395    /// Currently only supports `Affine::<P>` addition.
396    ///
397    /// Returns error if inputs are invalid
398    pub fn ecc_add<P: Config<BaseField = F>>(
399        &mut self,
400        point_a: &PointVariable,
401        point_b: &PointVariable,
402    ) -> Result<PointVariable, CircuitError> {
403        let (x_1, y_1) = (self.witness(point_a.0)?, self.witness(point_a.1)?);
404        let (x_2, y_2) = (self.witness(point_b.0)?, self.witness(point_b.1)?);
405        let eq_gate = EdwardsCurveEquationGate::<P> {
406            _phantom: PhantomData,
407        };
408        let d: F = eq_gate.q_ecc();
409
410        let z = d * x_1 * y_1 * x_2 * y_2; // temporary intermediate value
411        let x_3 = (x_1 * y_2 + x_2 * y_1) / (F::one() + z);
412        let y_3 = (-P::COEFF_A * x_1 * x_2 + y_2 * y_1) / (F::one() - z);
413        let point_c = self.create_point_variable(TEPoint(x_3, y_3))?;
414        self.ecc_add_gate::<P>(point_a, point_b, &point_c)?;
415
416        Ok(point_c)
417    }
418
419    /// Obtain the fixed-based scalar multiplication result of `scalar` * `Base`
420    /// Currently only supports `Affine::<P>` scalar multiplication.
421    pub fn fixed_base_scalar_mul<P: Config<BaseField = F>>(
422        &mut self,
423        scalar: Variable,
424        base: &Affine<P>,
425    ) -> Result<PointVariable, CircuitError> {
426        self.check_var_bound(scalar)?;
427
428        let mut num_bits = <P as CurveConfig>::ScalarField::MODULUS_BIT_SIZE as usize;
429        // `num_bits` needs to be an even number
430        num_bits += num_bits & 1;
431        let scalar_bits_le = self.unpack(scalar, num_bits)?;
432        let fixed_bases = compute_base_points(&base.into_group(), num_bits / 2)?;
433        let mut accum = self.neutral_point_variable();
434        for i in 0..num_bits / 2 {
435            let b0 = scalar_bits_le.get(2 * i).ok_or_else(|| {
436                CircuitError::InternalError(
437                    "scalar binary representation has the wrong length".to_string(),
438                )
439            })?;
440            let b1 = scalar_bits_le.get(2 * i + 1).ok_or_else(|| {
441                CircuitError::InternalError(
442                    "scalar binary representation has the wrong length".to_string(),
443                )
444            })?;
445            let p1 = fixed_bases[0].get(i).ok_or_else(|| {
446                CircuitError::InternalError("fixed_bases_1 has the wrong length".to_string())
447            })?;
448            let p2 = fixed_bases[1].get(i).ok_or_else(|| {
449                CircuitError::InternalError("fixed_bases_2 has the wrong length".to_string())
450            })?;
451            let p3 = fixed_bases[2].get(i).ok_or_else(|| {
452                CircuitError::InternalError("fixed_bases_3 has the wrong length".to_string())
453            })?;
454            let selected = self.quaternary_point_select::<P>(
455                *b0,
456                *b1,
457                &TEPoint::from(*p1),
458                &TEPoint::from(*p2),
459                &TEPoint::from(*p3),
460            )?;
461            accum = self.ecc_add::<P>(&accum, &selected)?;
462        }
463        Ok(accum)
464    }
465
466    /// Obtain a variable of the result of a variable base scalar
467    /// multiplication. both `scalar` and `base` are variables.
468    /// Currently only supports `Affine::<P>`.
469    /// If the parameter is bandersnatch, we will use GLV multiplication.
470    pub fn variable_base_scalar_mul<P: Config<BaseField = F>>(
471        &mut self,
472        scalar: Variable,
473        base: &PointVariable,
474    ) -> Result<PointVariable, CircuitError> {
475        self.check_var_bound(scalar)?;
476        self.check_point_var_bound(base)?;
477
478        if self.support_lookup()
479            && P::ScalarField::MODULUS_BIT_SIZE == 253
480            && P::BaseField::MODULUS_BIT_SIZE == 255
481        {
482            // bandersnatch glv multiplication
483            // FIXME: we do not have an easier flag to tell if a parameter
484            // is bandersnatch or not, yet.
485            self.glv_mul::<P>(scalar, base)
486        } else {
487            // non-bandersantch multiplication
488            msm::MultiScalarMultiplicationCircuit::<F, P>::msm(self, &[*base], &[scalar])
489        }
490    }
491
492    /// Obtain a variable of the result of a variable base scalar
493    /// multiplication. Both `scalar_bits_le` and `base` are variables,
494    /// where `scalar_bits_le` is the little-endian form of the scalar.
495    /// Currently only supports `Affine::<P>`.
496    pub fn variable_base_binary_scalar_mul<P: Config<BaseField = F>>(
497        &mut self,
498        scalar_bits_le: &[BoolVar],
499        base: &PointVariable,
500    ) -> Result<PointVariable, CircuitError> {
501        for &bit in scalar_bits_le {
502            self.check_var_bound(bit.into())?;
503        }
504        self.check_point_var_bound(base)?;
505
506        let neutral_point_var = self.neutral_point_variable();
507        let mut accum = neutral_point_var;
508        for i in (0..scalar_bits_le.len()).rev() {
509            let z = self.binary_point_vars_select(scalar_bits_le[i], &neutral_point_var, base)?;
510            accum = self.ecc_add::<P>(&accum, &accum)?;
511            accum = self.ecc_add::<P>(&accum, &z)?;
512        }
513        Ok(accum)
514    }
515}
516
517// private helper functions
518impl<F: PrimeField> PlonkCircuit<F> {
519    fn check_point_var_bound(&self, point_var: &PointVariable) -> Result<(), CircuitError> {
520        self.check_var_bound(point_var.0)?;
521        self.check_var_bound(point_var.1)?;
522        Ok(())
523    }
524}
525
526// Given a base point [G] and a scalar s of length 2*n, denote as s[G] the
527// scalar multiplication.
528// The function computes:
529// {4^i * [G]}_{i=0..n-1}, {2 * 4^i * [G]}_{i=0..n-1}, and {3 * 4^i *
530// [G]}_{i=0..n-1}
531// TODO (tessico): this used to operate on Affine points, but now it takes in
532// Projective points. There are some known issues with outputting projective
533// points, we should make sure that the usage here is safe.
534fn compute_base_points<E: ScalarMul>(base: &E, len: usize) -> Result<[Vec<E>; 3], CircuitError> {
535    if len == 0 {
536        return Err(CircuitError::InternalError(
537            "compute base points length input parameter must be positive".to_string(),
538        ));
539    }
540    fn next_base<E: ScalarMul>(bases: &[E]) -> Result<E, CircuitError> {
541        let last = *bases.last().ok_or_else(|| {
542            CircuitError::InternalError(
543                "Initialize the fixed base vector before calling this function".to_string(),
544            )
545        })?;
546        Ok(last.double().double())
547    }
548    fn fill_bases<E: ScalarMul>(bases: &mut Vec<E>, len: usize) -> Result<(), CircuitError> {
549        for _ in 1..len {
550            bases.push(next_base(bases)?);
551        }
552        Ok(())
553    }
554
555    let mut b = *base;
556    // base1 = (B, 4*B, ..., 4^(l-1)*B)
557    let mut bases1 = vec![b];
558    b = b.double();
559    // base2 = (2*B, 2*4*B, ..., 2*4^(l-1)*B)
560    let mut bases2 = vec![b];
561    b += base;
562    // base3 = (3*B, 3*4*B, ..., 3*4^(l-1)*B)
563    let mut bases3 = vec![b];
564
565    #[cfg(feature = "parallel")]
566    {
567        rayon::join(
568            || {
569                rayon::join(
570                    || fill_bases(&mut bases1, len).ok(),
571                    || fill_bases(&mut bases2, len).ok(),
572                )
573            },
574            || fill_bases(&mut bases3, len).ok(),
575        );
576    }
577
578    #[cfg(not(feature = "parallel"))]
579    {
580        fill_bases(&mut bases1, len).ok();
581        fill_bases(&mut bases2, len).ok();
582        fill_bases(&mut bases3, len).ok();
583    }
584
585    // converting Affine -> Points here.
586    // Cannot do it earlier: in `fill_bases` we need to do `double`
587    // todo(ZZ): consider removing `Point<T>` completely and directly use
588    // `Affine<P>` let bases1 =
589    // bases1.iter().map(|e|Point::<F>::from(*e)).collect(); let bases2 =
590    // bases2.iter().map(|e|Point::<F>::from(*e)).collect(); let bases3 =
591    // bases3.iter().map(|e|Point::<F>::from(*e)).collect();
592
593    Ok([bases1, bases2, bases3])
594}
595
596#[cfg(test)]
597mod test {
598    use super::*;
599    use crate::gadgets::test_utils::test_variable_independence_for_circuit;
600    use ark_bls12_377::{g1::Config as Param761, Fq as Fq377};
601    use ark_ec::{twisted_edwards::TECurveConfig as Config, Group};
602    use ark_ed_on_bls12_377::{EdwardsConfig as Param377, Fq as FqEd377};
603    use ark_ed_on_bls12_381::{EdwardsConfig as Param381, Fq as FqEd381};
604    use ark_ed_on_bls12_381_bandersnatch::{EdwardsConfig as Param381b, Fq as FqEd381b};
605    use ark_ed_on_bn254::{EdwardsConfig as Param254, Fq as FqEd354};
606    use ark_ff::{One, UniformRand, Zero};
607    use ark_std::str::FromStr;
608    use jf_utils::fr_to_fq;
609
610    #[test]
611    fn test_is_neutral() -> Result<(), CircuitError> {
612        test_is_neutral_helper::<FqEd354, Param254>()?;
613        test_is_neutral_helper::<FqEd377, Param377>()?;
614        test_is_neutral_helper::<FqEd381, Param381>()?;
615        test_is_neutral_helper::<FqEd381b, Param381b>()?;
616        test_is_neutral_helper::<Fq377, Param761>()
617    }
618
619    fn test_is_neutral_helper<F, P>() -> Result<(), CircuitError>
620    where
621        F: PrimeField,
622        P: Config<BaseField = F>,
623    {
624        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
625        let p1 = circuit.create_point_variable(TEPoint(F::zero(), F::one()))?;
626        let p2 = circuit.create_point_variable(TEPoint(F::from(2353u32), F::one()))?;
627        let p1_check = circuit.is_neutral_point::<P>(&p1)?;
628        let p2_check = circuit.is_neutral_point::<P>(&p2)?;
629
630        assert_eq!(circuit.witness(p1_check.into())?, F::one());
631        assert_eq!(circuit.witness(p2_check.into())?, F::zero());
632        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
633        *circuit.witness_mut(p1.0) = F::one();
634        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
635        // Check variable out of bound error.
636        assert!(circuit
637            .is_neutral_point::<P>(&PointVariable(circuit.num_vars(), circuit.num_vars() - 1))
638            .is_err());
639
640        let circuit_1 = build_is_neutral_circuit::<F, P>(TEPoint(F::zero(), F::one()))?;
641        let circuit_2 = build_is_neutral_circuit::<F, P>(TEPoint(F::one(), F::zero()))?;
642        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
643
644        Ok(())
645    }
646
647    fn build_is_neutral_circuit<F, P>(point: TEPoint<F>) -> Result<PlonkCircuit<F>, CircuitError>
648    where
649        F: PrimeField,
650        P: Config<BaseField = F>,
651    {
652        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
653        let p = circuit.create_point_variable(point)?;
654        circuit.is_neutral_point::<P>(&p)?;
655        circuit.finalize_for_arithmetization()?;
656        Ok(circuit)
657    }
658
659    macro_rules! test_enforce_on_curve {
660        ($fq:tt, $param:tt, $pt:tt) => {
661            let mut circuit: PlonkCircuit<$fq> = PlonkCircuit::new_turbo_plonk();
662            let p1 = circuit.create_point_variable(TEPoint($fq::zero(), $fq::one()))?;
663            circuit.enforce_on_curve::<$param>(&p1)?;
664            let p2 = circuit.create_point_variable($pt)?;
665            circuit.enforce_on_curve::<$param>(&p2)?;
666            assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
667
668            let p3 = circuit.create_point_variable(TEPoint($fq::one(), $fq::one()))?;
669            circuit.enforce_on_curve::<$param>(&p3)?;
670            assert!(circuit.check_circuit_satisfiability(&[]).is_err());
671            // Check variable out of bound error.
672            assert!(circuit
673                .enforce_on_curve::<$param>(&PointVariable(
674                    circuit.num_vars(),
675                    circuit.num_vars() - 1
676                ))
677                .is_err());
678
679            let circuit_1 =
680                build_enforce_on_curve_circuit::<_, $param>(TEPoint($fq::zero(), $fq::one()))?;
681            let circuit_2 = build_enforce_on_curve_circuit::<_, $param>(TEPoint(
682                $fq::from(5u32),
683                $fq::from(89u32),
684            ))?;
685            test_variable_independence_for_circuit(circuit_1, circuit_2)?;
686        };
687    }
688
689    #[test]
690    fn test_enforce_on_curve() -> Result<(), CircuitError> {
691        // generator for ed_on_bn254 curve
692        let ed_on_254_gen = TEPoint(
693            FqEd354::from_str(
694                "19698561148652590122159747500897617769866003486955115824547446575314762165298",
695            )
696            .unwrap(),
697            FqEd354::from_str(
698                "19298250018296453272277890825869354524455968081175474282777126169995084727839",
699            )
700            .unwrap(),
701        );
702        // generator for ed_on_bls377 curve
703        let ed_on_377_gen = TEPoint(
704            FqEd377::from_str(
705                "4497879464030519973909970603271755437257548612157028181994697785683032656389",
706            )
707            .unwrap(),
708            FqEd377::from_str(
709                "4357141146396347889246900916607623952598927460421559113092863576544024487809",
710            )
711            .unwrap(),
712        );
713        // generator for ed_on_bls381 curve
714        let ed_on_381_gen = TEPoint(
715            FqEd381::from_str(
716                "8076246640662884909881801758704306714034609987455869804520522091855516602923",
717            )
718            .unwrap(),
719            FqEd381::from_str(
720                "13262374693698910701929044844600465831413122818447359594527400194675274060458",
721            )
722            .unwrap(),
723        );
724        // generator for ed_on_bls381_bandersnatch curve
725        let ed_on_381b_gen = TEPoint(
726            FqEd381::from_str(
727                "18886178867200960497001835917649091219057080094937609519140440539760939937304",
728            )
729            .unwrap(),
730            FqEd381::from_str(
731                "19188667384257783945677642223292697773471335439753913231509108946878080696678",
732            )
733            .unwrap(),
734        );
735        // generator for bls377 G1 curve
736        let bls377_gen = TEPoint(
737            Fq377::from_str(
738                "71222569531709137229370268896323705690285216175189308202338047559628438110820800641278662592954630774340654489393",
739            )
740            .unwrap(),
741            Fq377::from_str(
742                "6177051365529633638563236407038680211609544222665285371549726196884440490905471891908272386851767077598415378235",
743            )
744            .unwrap(),
745        );
746
747        test_enforce_on_curve!(FqEd354, Param254, ed_on_254_gen);
748        test_enforce_on_curve!(FqEd377, Param377, ed_on_377_gen);
749        test_enforce_on_curve!(FqEd381, Param381, ed_on_381_gen);
750        test_enforce_on_curve!(FqEd381b, Param381b, ed_on_381b_gen);
751        test_enforce_on_curve!(Fq377, Param761, bls377_gen);
752        Ok(())
753    }
754
755    fn build_enforce_on_curve_circuit<F, P>(
756        point: TEPoint<F>,
757    ) -> Result<PlonkCircuit<F>, CircuitError>
758    where
759        F: PrimeField,
760        P: Config<BaseField = F>,
761    {
762        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
763        let p = circuit.create_point_variable(point)?;
764        circuit.enforce_on_curve::<P>(&p)?;
765        circuit.finalize_for_arithmetization()?;
766        Ok(circuit)
767    }
768
769    #[test]
770    fn test_curve_point_addition() -> Result<(), CircuitError> {
771        test_curve_point_addition_helper::<FqEd354, Param254>()?;
772        test_curve_point_addition_helper::<FqEd377, Param377>()?;
773        test_curve_point_addition_helper::<FqEd381, Param381>()?;
774        test_curve_point_addition_helper::<FqEd381b, Param381b>()?;
775        test_curve_point_addition_helper::<Fq377, Param761>()
776    }
777
778    fn test_curve_point_addition_helper<F, P>() -> Result<(), CircuitError>
779    where
780        F: PrimeField,
781        P: Config<BaseField = F>,
782    {
783        let mut rng = jf_utils::test_rng();
784        let p1 = Affine::<P>::rand(&mut rng);
785        let p2 = Affine::<P>::rand(&mut rng);
786        let p3 = (p1 + p2).into_affine();
787
788        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
789        let p1_var = circuit.create_point_variable(TEPoint::from(p1))?;
790        let p2_var = circuit.create_point_variable(TEPoint::from(p2))?;
791        let p3_var = circuit.ecc_add::<P>(&p1_var, &p2_var)?;
792
793        assert_eq!(circuit.witness(p3_var.0)?, p3.x);
794        assert_eq!(circuit.witness(p3_var.1)?, p3.y);
795        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
796        *circuit.witness_mut(p3_var.0) = F::zero();
797        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
798        // Check variable out of bound error.
799        assert!(circuit
800            .ecc_add::<P>(&PointVariable(0, 0), &PointVariable(1, circuit.num_vars()))
801            .is_err());
802
803        let p1 = Affine::<P>::rand(&mut rng);
804        let p2 = Affine::<P>::rand(&mut rng);
805        let p3 = Affine::<P>::rand(&mut rng);
806        let p4 = Affine::<P>::rand(&mut rng);
807        let circuit_1 =
808            build_curve_point_addition_circuit::<F, P>(TEPoint::from(p1), TEPoint::from(p2))?;
809        let circuit_2 =
810            build_curve_point_addition_circuit::<F, P>(TEPoint::from(p3), TEPoint::from(p4))?;
811        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
812
813        Ok(())
814    }
815
816    fn build_curve_point_addition_circuit<F, P>(
817        p1: TEPoint<F>,
818        p2: TEPoint<F>,
819    ) -> Result<PlonkCircuit<F>, CircuitError>
820    where
821        F: PrimeField,
822        P: Config<BaseField = F>,
823    {
824        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
825        let p1_var = circuit.create_point_variable(p1)?;
826        let p2_var = circuit.create_point_variable(p2)?;
827        circuit.ecc_add::<P>(&p1_var, &p2_var)?;
828        circuit.finalize_for_arithmetization()?;
829        Ok(circuit)
830    }
831
832    #[test]
833    fn test_quaternary_point_select() -> Result<(), CircuitError> {
834        test_quaternary_point_select_helper::<FqEd354, Param254>()?;
835        test_quaternary_point_select_helper::<FqEd377, Param377>()?;
836        test_quaternary_point_select_helper::<FqEd381, Param381>()?;
837        test_quaternary_point_select_helper::<FqEd381b, Param381b>()?;
838        test_quaternary_point_select_helper::<Fq377, Param761>()
839    }
840
841    fn test_quaternary_point_select_helper<F, P>() -> Result<(), CircuitError>
842    where
843        F: PrimeField,
844        P: Config<BaseField = F>,
845    {
846        let mut rng = jf_utils::test_rng();
847        let p1 = Affine::<P>::rand(&mut rng);
848        let p2 = Affine::<P>::rand(&mut rng);
849        let p3 = Affine::<P>::rand(&mut rng);
850
851        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
852        let false_var = circuit.false_var();
853        let true_var = circuit.true_var();
854
855        let select_p0 = circuit.quaternary_point_select::<P>(
856            false_var,
857            false_var,
858            &TEPoint::from(p1),
859            &TEPoint::from(p2),
860            &TEPoint::from(p3),
861        )?;
862        assert_eq!(
863            TEPoint(F::zero(), F::one()),
864            TEPoint(circuit.witness(select_p0.0)?, circuit.witness(select_p0.1)?)
865        );
866        let select_p1 = circuit.quaternary_point_select::<P>(
867            true_var,
868            false_var,
869            &TEPoint::from(p1),
870            &TEPoint::from(p2),
871            &TEPoint::from(p3),
872        )?;
873        assert_eq!(TEPoint::from(p1), circuit.point_witness(&select_p1)?);
874        let select_p2 = circuit.quaternary_point_select::<P>(
875            false_var,
876            true_var,
877            &TEPoint::from(p1),
878            &TEPoint::from(p2),
879            &TEPoint::from(p3),
880        )?;
881        assert_eq!(TEPoint::from(p2), circuit.point_witness(&select_p2)?);
882        let select_p3 = circuit.quaternary_point_select::<P>(
883            true_var,
884            true_var,
885            &TEPoint::from(p1),
886            &TEPoint::from(p2),
887            &TEPoint::from(p3),
888        )?;
889        assert_eq!(TEPoint::from(p3), circuit.point_witness(&select_p3)?);
890
891        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
892
893        *circuit.witness_mut(select_p3.0) = p2.x;
894        *circuit.witness_mut(select_p3.1) = p2.y;
895        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
896
897        let circuit_1 = build_quaternary_select_gate::<F, P>(false, false)?;
898        let circuit_2 = build_quaternary_select_gate::<F, P>(true, true)?;
899        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
900        Ok(())
901    }
902
903    fn build_quaternary_select_gate<F, P>(
904        b0: bool,
905        b1: bool,
906    ) -> Result<PlonkCircuit<F>, CircuitError>
907    where
908        F: PrimeField,
909        P: Config<BaseField = F>,
910    {
911        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
912        let b0_var = circuit.create_boolean_variable(b0)?;
913        let b1_var = circuit.create_boolean_variable(b1)?;
914
915        let mut rng = jf_utils::test_rng();
916        let p1 = Affine::<P>::rand(&mut rng);
917        let p2 = Affine::<P>::rand(&mut rng);
918        let p3 = Affine::<P>::rand(&mut rng);
919        circuit.quaternary_point_select::<P>(
920            b0_var,
921            b1_var,
922            &TEPoint::from(p1),
923            &TEPoint::from(p2),
924            &TEPoint::from(p3),
925        )?;
926        circuit.finalize_for_arithmetization()?;
927        Ok(circuit)
928    }
929
930    #[test]
931    fn test_point_equal_gate() -> Result<(), CircuitError> {
932        test_point_equal_gate_helper::<FqEd354, Param254>()?;
933        test_point_equal_gate_helper::<FqEd377, Param377>()?;
934        test_point_equal_gate_helper::<FqEd381, Param381>()?;
935        test_point_equal_gate_helper::<FqEd381b, Param381b>()?;
936        test_point_equal_gate_helper::<Fq377, Param761>()
937    }
938
939    fn test_point_equal_gate_helper<F, P>() -> Result<(), CircuitError>
940    where
941        F: PrimeField,
942        P: Config<BaseField = F>,
943    {
944        let mut rng = jf_utils::test_rng();
945        let p = Affine::<P>::rand(&mut rng);
946
947        let mut circuit = PlonkCircuit::<F>::new_turbo_plonk();
948        let p1_var = circuit.create_point_variable(TEPoint::from(p))?;
949        let p2_var = circuit.create_point_variable(TEPoint::from(p))?;
950        circuit.enforce_point_equal(&p1_var, &p2_var)?;
951
952        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
953        *circuit.witness_mut(p2_var.0) = F::zero();
954        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
955        // Check variable out of bound error.
956        assert!(circuit
957            .enforce_point_equal(&PointVariable(0, 0), &PointVariable(1, circuit.num_vars()))
958            .is_err());
959
960        let new_p = Affine::<P>::rand(&mut rng);
961        let circuit_1 = build_point_equal_circuit(TEPoint::from(p), TEPoint::from(p))?;
962        let circuit_2 = build_point_equal_circuit(TEPoint::from(new_p), TEPoint::from(new_p))?;
963        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
964
965        Ok(())
966    }
967
968    fn build_point_equal_circuit<F: PrimeField>(
969        p1: TEPoint<F>,
970        p2: TEPoint<F>,
971    ) -> Result<PlonkCircuit<F>, CircuitError> {
972        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
973        let p1_var = circuit.create_point_variable(p1)?;
974        let p2_var = circuit.create_point_variable(p2)?;
975        circuit.enforce_point_equal(&p1_var, &p2_var)?;
976        circuit.finalize_for_arithmetization()?;
977        Ok(circuit)
978    }
979
980    #[test]
981    fn test_is_equal_point() -> Result<(), CircuitError> {
982        test_is_equal_point_helper::<FqEd354, Param254>()?;
983        test_is_equal_point_helper::<FqEd377, Param377>()?;
984        test_is_equal_point_helper::<FqEd381, Param381>()?;
985        test_is_equal_point_helper::<FqEd381b, Param381b>()?;
986        test_is_equal_point_helper::<Fq377, Param761>()
987    }
988
989    fn test_is_equal_point_helper<F, P>() -> Result<(), CircuitError>
990    where
991        F: PrimeField,
992        P: Config<BaseField = F>,
993    {
994        let mut rng = jf_utils::test_rng();
995        let p1 = Affine::<P>::rand(&mut rng);
996        let p2 = p1;
997        let p3 = Affine::<P>::rand(&mut rng);
998
999        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1000        let p1_var = circuit.create_point_variable(TEPoint::from(p1))?;
1001        let p2_var = circuit.create_point_variable(TEPoint::from(p2))?;
1002        let p3_var = circuit.create_point_variable(TEPoint::from(p3))?;
1003        let p1_p2_eq = circuit.is_point_equal(&p1_var, &p2_var)?;
1004        let p1_p3_eq = circuit.is_point_equal(&p1_var, &p3_var)?;
1005
1006        assert_eq!(circuit.witness(p1_p2_eq.into())?, F::one());
1007        assert_eq!(circuit.witness(p1_p3_eq.into())?, F::zero());
1008        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
1009        *circuit.witness_mut(p2_var.0) = F::zero();
1010        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
1011        // Check variable out of bound error.
1012        assert!(circuit
1013            .is_point_equal(&PointVariable(0, 0), &PointVariable(1, circuit.num_vars()))
1014            .is_err());
1015
1016        let circuit_1 =
1017            build_is_equal_point_circuit(TEPoint::from(p1), TEPoint::from(p2), TEPoint::from(p3))?;
1018        let circuit_2 =
1019            build_is_equal_point_circuit(TEPoint::from(p3), TEPoint::from(p3), TEPoint::from(p1))?;
1020        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
1021
1022        Ok(())
1023    }
1024
1025    fn build_is_equal_point_circuit<F: PrimeField>(
1026        p1: TEPoint<F>,
1027        p2: TEPoint<F>,
1028        p3: TEPoint<F>,
1029    ) -> Result<PlonkCircuit<F>, CircuitError> {
1030        let mut circuit = PlonkCircuit::new_turbo_plonk();
1031        let p1_var = circuit.create_point_variable(p1)?;
1032        let p2_var = circuit.create_point_variable(p2)?;
1033        let p3_var = circuit.create_point_variable(p3)?;
1034        circuit.is_point_equal(&p1_var, &p2_var)?;
1035        circuit.is_point_equal(&p1_var, &p3_var)?;
1036        circuit.finalize_for_arithmetization()?;
1037        Ok(circuit)
1038    }
1039
1040    #[test]
1041    fn test_compute_fixed_bases() -> Result<(), CircuitError> {
1042        test_compute_fixed_bases_helper::<FqEd354, Param254>()?;
1043        test_compute_fixed_bases_helper::<FqEd377, Param377>()?;
1044        test_compute_fixed_bases_helper::<FqEd381, Param381>()?;
1045        test_compute_fixed_bases_helper::<FqEd381b, Param381b>()?;
1046        test_compute_fixed_bases_helper::<Fq377, Param761>()
1047    }
1048
1049    fn test_compute_fixed_bases_helper<F, P>() -> Result<(), CircuitError>
1050    where
1051        F: PrimeField,
1052        P: Config<BaseField = F>,
1053    {
1054        fn check_base_list<F, P>(bases: &[Projective<P>])
1055        where
1056            F: PrimeField,
1057            P: Config<BaseField = F>,
1058        {
1059            bases
1060                .windows(2)
1061                .for_each(|neighbors| assert!(neighbors[1] == neighbors[0].double().double()));
1062        }
1063
1064        let mut rng = jf_utils::test_rng();
1065
1066        let base = Affine::<P>::rand(&mut rng);
1067        let base2 = base.into_group().double();
1068        let base3 = base + base2;
1069
1070        assert_eq!(
1071            compute_base_points(&base.into_group(), 1)?,
1072            [vec![base.into_group()], vec![base2], vec![base3]]
1073        );
1074        let size = 10;
1075        let result = compute_base_points(&base.into_group(), size)?;
1076        let bases1 = &result[0];
1077        assert_eq!(bases1.len(), size);
1078        let bases2 = &result[1];
1079        assert_eq!(bases2.len(), size);
1080        let bases3 = &result[2];
1081        assert_eq!(bases3.len(), size);
1082        check_base_list(bases1);
1083        check_base_list(bases2);
1084        check_base_list(bases3);
1085        bases1
1086            .iter()
1087            .zip(bases2.iter())
1088            .zip(bases3.iter())
1089            .for_each(|((&b1, &b2), &b3)| {
1090                assert!(b2 == b1.double());
1091                assert!(b3 == b1 + b2);
1092            });
1093
1094        Ok(())
1095    }
1096
1097    #[test]
1098    fn test_fixed_based_scalar_mul() -> Result<(), CircuitError> {
1099        test_fixed_based_scalar_mul_helper::<FqEd354, Param254>()?;
1100        test_fixed_based_scalar_mul_helper::<FqEd377, Param377>()?;
1101        test_fixed_based_scalar_mul_helper::<FqEd381, Param381>()?;
1102        test_fixed_based_scalar_mul_helper::<FqEd381b, Param381b>()?;
1103        test_fixed_based_scalar_mul_helper::<Fq377, Param761>()
1104    }
1105
1106    fn test_fixed_based_scalar_mul_helper<F, P>() -> Result<(), CircuitError>
1107    where
1108        F: PrimeField,
1109        P: Config<BaseField = F>,
1110    {
1111        let mut rng = jf_utils::test_rng();
1112        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1113
1114        for _ in 0..6 {
1115            let mut base = Affine::<P>::rand(&mut rng);
1116            let s = P::ScalarField::rand(&mut rng);
1117            let scalar = circuit.create_variable(fr_to_fq::<F, P>(&s))?;
1118            let result = circuit.fixed_base_scalar_mul(scalar, &base)?;
1119            base = (base * s).into();
1120            assert_eq!(TEPoint::from(base), circuit.point_witness(&result)?);
1121        }
1122        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
1123
1124        // wrong witness should fail
1125        *circuit.witness_mut(2) = F::rand(&mut rng);
1126        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
1127        // Check variable out of bound error.
1128        assert!(circuit
1129            .fixed_base_scalar_mul(circuit.num_vars(), &Affine::<P>::rand(&mut rng))
1130            .is_err());
1131
1132        let circuit_1 = build_fixed_based_scalar_mul_circuit::<F, P>(F::from(87u32))?;
1133        let circuit_2 = build_fixed_based_scalar_mul_circuit::<F, P>(F::from(2u32))?;
1134        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
1135        Ok(())
1136    }
1137
1138    fn build_fixed_based_scalar_mul_circuit<F, P>(
1139        scalar: F,
1140    ) -> Result<PlonkCircuit<F>, CircuitError>
1141    where
1142        F: PrimeField,
1143        P: Config<BaseField = F>,
1144    {
1145        let mut rng = jf_utils::test_rng();
1146        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1147        let base = Affine::<P>::rand(&mut rng);
1148        let scalar_var = circuit.create_variable(scalar)?;
1149        circuit.fixed_base_scalar_mul(scalar_var, &base)?;
1150        circuit.finalize_for_arithmetization()?;
1151        Ok(circuit)
1152    }
1153
1154    #[test]
1155    fn test_binary_point_vars_select() -> Result<(), CircuitError> {
1156        test_binary_point_vars_select_helper::<FqEd354, Param254>()?;
1157        test_binary_point_vars_select_helper::<FqEd377, Param377>()?;
1158        test_binary_point_vars_select_helper::<FqEd381, Param381>()?;
1159        test_binary_point_vars_select_helper::<FqEd381b, Param381b>()?;
1160        test_binary_point_vars_select_helper::<Fq377, Param761>()
1161    }
1162    fn test_binary_point_vars_select_helper<F, P>() -> Result<(), CircuitError>
1163    where
1164        F: PrimeField,
1165        P: Config<BaseField = F>,
1166    {
1167        let mut rng = jf_utils::test_rng();
1168        let p0 = Affine::<P>::rand(&mut rng);
1169        let p1 = Affine::<P>::rand(&mut rng);
1170        let p2 = Affine::<P>::rand(&mut rng);
1171
1172        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1173        let p0_var = circuit.create_point_variable(TEPoint::from(p0))?;
1174        let p1_var = circuit.create_point_variable(TEPoint::from(p1))?;
1175        let true_var = circuit.true_var();
1176        let false_var = circuit.false_var();
1177
1178        let select_p0 = circuit.binary_point_vars_select(false_var, &p0_var, &p1_var)?;
1179        assert_eq!(circuit.point_witness(&select_p0)?, TEPoint::from(p0));
1180        let select_p1 = circuit.binary_point_vars_select(true_var, &p0_var, &p1_var)?;
1181        assert_eq!(circuit.point_witness(&select_p1)?, TEPoint::from(p1));
1182        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
1183
1184        // wrong witness should fail
1185        *circuit.witness_mut(p1_var.0) = F::rand(&mut rng);
1186        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
1187        assert!(circuit
1188            .binary_point_vars_select(
1189                false_var,
1190                &p0_var,
1191                &PointVariable(p1_var.0, circuit.num_vars()),
1192            )
1193            .is_err());
1194
1195        let circuit_1 = build_binary_point_vars_select_circuit::<F>(
1196            true,
1197            TEPoint::from(p0),
1198            TEPoint::from(p1),
1199        )?;
1200        let circuit_2 = build_binary_point_vars_select_circuit::<F>(
1201            false,
1202            TEPoint::from(p1),
1203            TEPoint::from(p2),
1204        )?;
1205        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
1206
1207        Ok(())
1208    }
1209
1210    fn build_binary_point_vars_select_circuit<F>(
1211        b: bool,
1212        p0: TEPoint<F>,
1213        p1: TEPoint<F>,
1214    ) -> Result<PlonkCircuit<F>, CircuitError>
1215    where
1216        F: PrimeField,
1217    {
1218        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1219        let b_var = circuit.create_boolean_variable(b)?;
1220        let p0_var = circuit.create_point_variable(p0)?;
1221        let p1_var = circuit.create_point_variable(p1)?;
1222        circuit.binary_point_vars_select(b_var, &p0_var, &p1_var)?;
1223        circuit.finalize_for_arithmetization()?;
1224        Ok(circuit)
1225    }
1226
1227    #[test]
1228    fn test_variable_base_scalar_mul() -> Result<(), CircuitError> {
1229        test_variable_base_scalar_mul_helper::<FqEd354, Param254>()?;
1230        test_variable_base_scalar_mul_helper::<FqEd377, Param377>()?;
1231        test_variable_base_scalar_mul_helper::<FqEd381, Param381>()?;
1232        test_variable_base_scalar_mul_helper::<FqEd381b, Param381b>()?;
1233        test_variable_base_scalar_mul_helper::<Fq377, Param761>()
1234    }
1235    fn test_variable_base_scalar_mul_helper<F, P>() -> Result<(), CircuitError>
1236    where
1237        F: PrimeField,
1238        P: Config<BaseField = F>,
1239    {
1240        let mut rng = jf_utils::test_rng();
1241        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1242
1243        for _ in 0..6 {
1244            let mut base = Affine::<P>::rand(&mut rng);
1245            let s = P::ScalarField::rand(&mut rng);
1246            let s_var = circuit.create_variable(fr_to_fq::<F, P>(&s))?;
1247            let base_var = circuit.create_point_variable(TEPoint::from(base))?;
1248            base = (base * s).into();
1249            let result = circuit.variable_base_scalar_mul::<P>(s_var, &base_var)?;
1250            assert_eq!(TEPoint::from(base), circuit.point_witness(&result)?);
1251        }
1252        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
1253
1254        let base = Affine::<P>::rand(&mut rng);
1255        let s = P::ScalarField::rand(&mut rng);
1256        let s_var = circuit.create_variable(fr_to_fq::<F, P>(&s))?;
1257        let base_var = circuit.create_point_variable(TEPoint::from(base))?;
1258        // wrong witness should fail
1259        *circuit.witness_mut(2) = F::rand(&mut rng);
1260        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
1261        // Check variable out of bound error.
1262        assert!(circuit
1263            .variable_base_scalar_mul::<P>(circuit.num_vars(), &base_var)
1264            .is_err());
1265        assert!(circuit
1266            .variable_base_scalar_mul::<P>(
1267                s_var,
1268                &PointVariable(circuit.num_vars(), circuit.num_vars())
1269            )
1270            .is_err());
1271
1272        let circuit_1 =
1273            build_variable_base_scalar_mul_circuit::<F, P>(F::zero(), TEPoint::from(base))?;
1274        let circuit_2 = build_variable_base_scalar_mul_circuit::<F, P>(
1275            F::from(314u32),
1276            TEPoint::from(Affine::<P>::rand(&mut rng)),
1277        )?;
1278        test_variable_independence_for_circuit(circuit_1, circuit_2)?;
1279
1280        Ok(())
1281    }
1282
1283    fn build_variable_base_scalar_mul_circuit<F, P>(
1284        scalar: F,
1285        base: TEPoint<F>,
1286    ) -> Result<PlonkCircuit<F>, CircuitError>
1287    where
1288        F: PrimeField,
1289        P: Config<BaseField = F>,
1290    {
1291        let mut circuit: PlonkCircuit<F> = PlonkCircuit::new_turbo_plonk();
1292        let scalar_var = circuit.create_variable(scalar)?;
1293        let base_var = circuit.create_point_variable(base)?;
1294        circuit.variable_base_scalar_mul::<P>(scalar_var, &base_var)?;
1295        circuit.finalize_for_arithmetization()?;
1296        Ok(circuit)
1297    }
1298}