jf_relation/gadgets/
emulated.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//! Emulate arithmetic operations on a large prime field.
8//! To emulate arithmetic operations on F_q when the native field is F_p where p
9//! < q, we represent the elements in F_q using CRT modulus [p, 2^T] where p *
10//! 2^T > q^2 + q. This constraint is required to emulate the F_q multiplication
11//! by checking a * b - k * q = c (mod 2^T * p) without any overflow. The second
12//! component, with modulus 2^T, will be divided into limbs each with B bits
13//! where 2^{2B} < p.
14
15use crate::{BoolVar, Circuit, CircuitError, PlonkCircuit, Variable};
16use ark_ff::PrimeField;
17use ark_std::{string::ToString, vec, vec::Vec, One, Zero};
18use core::marker::PhantomData;
19use itertools::izip;
20use num_bigint::BigUint;
21
22/// Parameters needed for emulating field operations over [`PrimeField`].
23pub trait EmulationConfig<F: PrimeField>: PrimeField {
24    /// Log2 of the other CRT modulus is 2^T.
25    const T: usize;
26    /// Bit length of each limbs.
27    const B: usize;
28    /// `B * NUM_LIMBS` should equals to `T`.
29    const NUM_LIMBS: usize;
30}
31
32/// A struct that can be serialized into `Vec` of field elements.
33pub trait SerializableEmulatedStruct<F: PrimeField> {
34    /// Serialize into a `Vec` of field elements.
35    fn serialize_to_native_elements(&self) -> Vec<F>;
36}
37
38fn biguint_to_limbs<F: PrimeField>(val: &BigUint, b: usize, num_limbs: usize) -> Vec<F> {
39    let mut result = vec![];
40    let b_pow = BigUint::one() << b;
41    let mut val = val.clone();
42
43    // Since q < 2^T, no need to perform mod 2^T
44    for _ in 0..num_limbs {
45        result.push(F::from(&val % &b_pow));
46        val /= &b_pow;
47    }
48    result
49}
50
51/// Convert an element in the emulated field to a list of native field elements.
52pub fn from_emulated_field<E, F>(val: E) -> Vec<F>
53where
54    E: EmulationConfig<F>,
55    F: PrimeField,
56{
57    biguint_to_limbs(&val.into(), E::B, E::NUM_LIMBS)
58}
59
60/// Inverse conversion of the [`from_emulated_field`]
61pub fn to_emulated_field<E, F>(vals: &[F]) -> Result<E, CircuitError>
62where
63    E: EmulationConfig<F>,
64    F: PrimeField,
65{
66    if vals.len() != E::NUM_LIMBS {
67        return Err(CircuitError::FieldAlgebraError(
68            "Malformed structure for emulated field element conversion.".to_string(),
69        ));
70    }
71    let b_pow = BigUint::one() << E::B;
72    Ok(E::from(
73        vals.iter().rfold(BigUint::zero(), |result, &val| {
74            result * &b_pow + <F as Into<BigUint>>::into(val)
75        }),
76    ))
77}
78
79/// The variable represents an element in the emulated field.
80#[derive(Debug, Clone)]
81pub struct EmulatedVariable<E: PrimeField>(pub(crate) Vec<Variable>, pub PhantomData<E>);
82
83impl<E: PrimeField> EmulatedVariable<E> {
84    /// Return the list of variables that simulate the field element
85    pub fn native_vars(&self) -> Vec<Variable> {
86        self.0.clone()
87    }
88}
89
90impl<F: PrimeField> PlonkCircuit<F> {
91    /// Return the witness point for the circuit
92    pub fn emulated_witness<E: EmulationConfig<F>>(
93        &self,
94        var: &EmulatedVariable<E>,
95    ) -> Result<E, CircuitError> {
96        let values = var
97            .0
98            .iter()
99            .map(|&v| self.witness(v))
100            .collect::<Result<Vec<_>, CircuitError>>()?;
101        to_emulated_field(&values)
102    }
103
104    /// Add an emulated variable
105    pub fn create_emulated_variable<E: EmulationConfig<F>>(
106        &mut self,
107        val: E,
108    ) -> Result<EmulatedVariable<E>, CircuitError> {
109        let var = self.create_emulated_variable_unchecked(val)?;
110        for &v in &var.0 {
111            self.enforce_in_range(v, E::B)?;
112        }
113        Ok(var)
114    }
115
116    /// Add an emulated variable without enforcing the validity check
117    fn create_emulated_variable_unchecked<E: EmulationConfig<F>>(
118        &mut self,
119        val: E,
120    ) -> Result<EmulatedVariable<E>, CircuitError> {
121        Ok(EmulatedVariable::<E>(
122            from_emulated_field(val)
123                .into_iter()
124                .map(|v| self.create_variable(v))
125                .collect::<Result<Vec<_>, CircuitError>>()?,
126            PhantomData,
127        ))
128    }
129    /// Add a constant emulated variable
130    pub fn create_constant_emulated_variable<E: EmulationConfig<F>>(
131        &mut self,
132        val: E,
133    ) -> Result<EmulatedVariable<E>, CircuitError> {
134        Ok(EmulatedVariable::<E>(
135            from_emulated_field(val)
136                .into_iter()
137                .map(|v| self.create_constant_variable(v))
138                .collect::<Result<Vec<_>, CircuitError>>()?,
139            PhantomData,
140        ))
141    }
142
143    /// Add a public emulated variable
144    pub fn create_public_emulated_variable<E: EmulationConfig<F>>(
145        &mut self,
146        val: E,
147    ) -> Result<EmulatedVariable<E>, CircuitError> {
148        Ok(EmulatedVariable::<E>(
149            from_emulated_field(val)
150                .into_iter()
151                .map(|v| self.create_public_variable(v))
152                .collect::<Result<Vec<_>, CircuitError>>()?,
153            PhantomData,
154        ))
155    }
156
157    /// Constrain that a*b=c in the emulated field.
158    /// Checking that a * b - k * E::MODULUS = c.
159    /// This function doesn't perform emulated variable validity check on the
160    /// input a, b and c. We assume that they are already performed elsewhere.
161    pub fn emulated_mul_gate<E: EmulationConfig<F>>(
162        &mut self,
163        a: &EmulatedVariable<E>,
164        b: &EmulatedVariable<E>,
165        c: &EmulatedVariable<E>,
166    ) -> Result<(), CircuitError> {
167        self.check_vars_bound(&a.0)?;
168        self.check_vars_bound(&b.0)?;
169        self.check_vars_bound(&c.0)?;
170
171        let val_a: BigUint = self.emulated_witness(a)?.into();
172        let val_b: BigUint = self.emulated_witness(b)?.into();
173        let val_k = E::from(&val_a * &val_b / E::MODULUS.into());
174        let k = self.create_emulated_variable(val_k)?;
175        let a_limbs = biguint_to_limbs::<F>(&val_a, E::B, E::NUM_LIMBS);
176        let b_limbs = biguint_to_limbs::<F>(&val_b, E::B, E::NUM_LIMBS);
177        let k_limbs = from_emulated_field(val_k);
178        let b_pow = F::from(2u32).pow([E::B as u64]);
179        let val_expected = E::from(val_a) * E::from(val_b);
180        let val_expected_limbs = from_emulated_field(val_expected);
181
182        let neg_modulus = biguint_to_limbs::<F>(
183            &(BigUint::from(2u32).pow(E::T as u32) - E::MODULUS.into()),
184            E::B,
185            E::NUM_LIMBS,
186        );
187
188        // enforcing a * b - k * E::MODULUS = c mod 2^t
189
190        // first compare the first limb
191        let mut val_carry_out =
192            (a_limbs[0] * b_limbs[0] + k_limbs[0] * neg_modulus[0] - val_expected_limbs[0]) / b_pow;
193        let mut carry_out = self.create_variable(val_carry_out)?;
194        // checking that the carry_out has at most [`E::B`] + 1 bits
195        self.enforce_in_range(carry_out, E::B + 1)?;
196        // enforcing that a0 * b0 - k0 * modulus[0] - carry_out * 2^E::B = c0
197        self.quad_poly_gate(
198            &[a.0[0], b.0[0], k.0[0], carry_out, c.0[0]],
199            &[F::zero(), F::zero(), neg_modulus[0], -b_pow],
200            &[F::one(), F::zero()],
201            F::one(),
202            F::zero(),
203        )?;
204
205        for i in 1..E::NUM_LIMBS {
206            // compare the i-th limb
207
208            // calculate the next carry out
209            let val_next_carry_out = ((0..=i)
210                .map(|j| k_limbs[j] * neg_modulus[i - j] + a_limbs[j] * b_limbs[i - j])
211                .sum::<F>()
212                + val_carry_out
213                - val_expected_limbs[i])
214                / b_pow;
215            let next_carry_out = self.create_variable(val_next_carry_out)?;
216
217            // range checking for this carry out.
218            // let a = 2^B - 1. The maximum possible value of `next_carry_out` is ((i + 1) *
219            // 2 * a^2 + a) / 2^B.
220            let num_vals = 2u64 * (i as u64) + 2;
221            let log_num_vals = (u64::BITS - num_vals.leading_zeros()) as usize;
222            self.enforce_in_range(next_carry_out, E::B + log_num_vals)?;
223
224            // k * E::MODULUS part, waiting for summation
225            let mut stack = (0..=i)
226                .map(|j| (k.0[j], neg_modulus[i - j]))
227                .collect::<Vec<_>>();
228            // carry out from last limb
229            stack.push((carry_out, F::one()));
230            stack.push((next_carry_out, -b_pow));
231
232            // part of the summation \sum_j a_i * b_{i-j}
233            for j in (0..i).step_by(2) {
234                let t = self.mul_add(
235                    &[a.0[j], b.0[i - j], a.0[j + 1], b.0[i - j - 1]],
236                    &[F::one(), F::one()],
237                )?;
238                stack.push((t, F::one()));
239            }
240
241            // last item of the summation \sum_j a_i * b_{i-j}
242            if i % 2 == 0 {
243                let t1 = stack.pop().unwrap();
244                let t2 = stack.pop().unwrap();
245                let t = self.gen_quad_poly(
246                    &[a.0[i], b.0[0], t1.0, t2.0],
247                    &[F::zero(), F::zero(), t1.1, t2.1],
248                    &[F::one(), F::zero()],
249                    F::zero(),
250                )?;
251                stack.push((t, F::one()));
252            }
253
254            // linear combination of all items in the stack
255            while stack.len() > 4 {
256                let t1 = stack.pop().unwrap();
257                let t2 = stack.pop().unwrap();
258                let t3 = stack.pop().unwrap();
259                let t4 = stack.pop().unwrap();
260                let t = self.lc(&[t1.0, t2.0, t3.0, t4.0], &[t1.1, t2.1, t3.1, t4.1])?;
261                stack.push((t, F::one()));
262            }
263            let t1 = stack.pop().unwrap_or((self.zero(), F::zero()));
264            let t2 = stack.pop().unwrap_or((self.zero(), F::zero()));
265            let t3 = stack.pop().unwrap_or((self.zero(), F::zero()));
266            let t4 = stack.pop().unwrap_or((self.zero(), F::zero()));
267
268            // checking that the summation equals to i-th limb of c
269            self.lc_gate(&[t1.0, t2.0, t3.0, t4.0, c.0[i]], &[t1.1, t2.1, t3.1, t4.1])?;
270
271            val_carry_out = val_next_carry_out;
272            carry_out = next_carry_out;
273        }
274
275        // enforcing a * b - k * E::MODULUS = c mod F::MODULUS
276        let a_mod = self.mod_to_native_field(a)?;
277        let b_mod = self.mod_to_native_field(b)?;
278        let k_mod = self.mod_to_native_field(&k)?;
279        let c_mod = self.mod_to_native_field(c)?;
280        let e_mod_f = F::from(E::MODULUS.into());
281        self.quad_poly_gate(
282            &[a_mod, b_mod, k_mod, self.zero(), c_mod],
283            &[F::zero(), F::zero(), -e_mod_f, F::zero()],
284            &[F::one(), F::zero()],
285            F::one(),
286            F::zero(),
287        )?;
288
289        Ok(())
290    }
291
292    /// Return an [`EmulatedVariable`] which equals to a*b.
293    pub fn emulated_mul<E: EmulationConfig<F>>(
294        &mut self,
295        a: &EmulatedVariable<E>,
296        b: &EmulatedVariable<E>,
297    ) -> Result<EmulatedVariable<E>, CircuitError> {
298        let c = self.emulated_witness(a)? * self.emulated_witness(b)?;
299        let c = self.create_emulated_variable(c)?;
300        self.emulated_mul_gate(a, b, &c)?;
301        Ok(c)
302    }
303
304    /// Constrain that a*b=c in the emulated field for a constant b.
305    /// This function doesn't perform emulated variable validity check on the
306    /// input a and c. We assume that they are already performed elsewhere.
307    pub fn emulated_mul_constant_gate<E: EmulationConfig<F>>(
308        &mut self,
309        a: &EmulatedVariable<E>,
310        b: E,
311        c: &EmulatedVariable<E>,
312    ) -> Result<(), CircuitError> {
313        self.check_vars_bound(&a.0)?;
314        self.check_vars_bound(&c.0)?;
315
316        let val_a: BigUint = self.emulated_witness(a)?.into();
317        let val_b: BigUint = b.into();
318        let val_k = E::from(&val_a * &val_b / E::MODULUS.into());
319        let k = self.create_emulated_variable(val_k)?;
320        let a_limbs = biguint_to_limbs::<F>(&val_a, E::B, E::NUM_LIMBS);
321        let b_limbs = biguint_to_limbs::<F>(&val_b, E::B, E::NUM_LIMBS);
322        let k_limbs = from_emulated_field(val_k);
323        let b_pow = F::from(2u32).pow([E::B as u64]);
324        let val_expected = E::from(val_a) * b;
325        let val_expected_limbs = from_emulated_field(val_expected);
326
327        let neg_modulus = biguint_to_limbs::<F>(
328            &(BigUint::from(2u32).pow(E::T as u32) - E::MODULUS.into()),
329            E::B,
330            E::NUM_LIMBS,
331        );
332
333        // enforcing a * b - k * E::MODULUS = c mod 2^t
334
335        // first compare the first limb
336        let mut val_carry_out =
337            (a_limbs[0] * b_limbs[0] + k_limbs[0] * neg_modulus[0] - val_expected_limbs[0]) / b_pow;
338        let mut carry_out = self.create_variable(val_carry_out)?;
339        // checking that the carry_out has at most [`E::B`] bits
340        self.enforce_in_range(carry_out, E::B + 1)?;
341        // enforcing that a0 * b0 - k0 * modulus[0] - carry_out * 2^E::B = c0
342        self.lc_gate(
343            &[a.0[0], k.0[0], carry_out, self.zero(), c.0[0]],
344            &[b_limbs[0], neg_modulus[0], -b_pow, F::zero()],
345        )?;
346
347        for i in 1..E::NUM_LIMBS {
348            // compare the i-th limb
349
350            // calculate the next carry out
351            let val_next_carry_out = ((0..=i)
352                .map(|j| k_limbs[j] * neg_modulus[i - j] + a_limbs[j] * b_limbs[i - j])
353                .sum::<F>()
354                + val_carry_out
355                - val_expected_limbs[i])
356                / b_pow;
357            let next_carry_out = self.create_variable(val_next_carry_out)?;
358
359            // range checking for this carry out.
360            let num_vals = 2u64 * (i as u64) + 2;
361            let log_num_vals = (u64::BITS - num_vals.leading_zeros()) as usize;
362            self.enforce_in_range(next_carry_out, E::B + log_num_vals)?;
363
364            // k * E::MODULUS part, waiting for summation
365            let mut stack = (0..=i)
366                .map(|j| (k.0[j], neg_modulus[i - j]))
367                .collect::<Vec<_>>();
368            // a * b part
369            (0..=i).for_each(|j| stack.push((a.0[j], b_limbs[i - j])));
370            // carry out from last limb
371            stack.push((carry_out, F::one()));
372            stack.push((next_carry_out, -b_pow));
373
374            // linear combination of all items in the stack
375            while stack.len() > 4 {
376                let t1 = stack.pop().unwrap();
377                let t2 = stack.pop().unwrap();
378                let t3 = stack.pop().unwrap();
379                let t4 = stack.pop().unwrap();
380                let t = self.lc(&[t1.0, t2.0, t3.0, t4.0], &[t1.1, t2.1, t3.1, t4.1])?;
381                stack.push((t, F::one()));
382            }
383            let t1 = stack.pop().unwrap_or((self.zero(), F::zero()));
384            let t2 = stack.pop().unwrap_or((self.zero(), F::zero()));
385            let t3 = stack.pop().unwrap_or((self.zero(), F::zero()));
386            let t4 = stack.pop().unwrap_or((self.zero(), F::zero()));
387
388            // checking that the summation equals to i-th limb of c
389            self.lc_gate(&[t1.0, t2.0, t3.0, t4.0, c.0[i]], &[t1.1, t2.1, t3.1, t4.1])?;
390
391            val_carry_out = val_next_carry_out;
392            carry_out = next_carry_out;
393        }
394
395        // enforcing a * b - k * E::MODULUS = c mod F::MODULUS
396        let a_mod = self.mod_to_native_field(a)?;
397        let b_mod = F::from(val_b);
398        let k_mod = self.mod_to_native_field(&k)?;
399        let c_mod = self.mod_to_native_field(c)?;
400        let e_mod_f = F::from(E::MODULUS.into());
401        self.lc_gate(
402            &[a_mod, k_mod, self.zero(), self.zero(), c_mod],
403            &[b_mod, -e_mod_f, F::zero(), F::zero()],
404        )?;
405
406        Ok(())
407    }
408
409    /// Return an [`EmulatedVariable`] which equals to a*b.
410    pub fn emulated_mul_constant<E: EmulationConfig<F>>(
411        &mut self,
412        a: &EmulatedVariable<E>,
413        b: E,
414    ) -> Result<EmulatedVariable<E>, CircuitError> {
415        let c = self.emulated_witness(a)? * b;
416        let c = self.create_emulated_variable(c)?;
417        self.emulated_mul_constant_gate(a, b, &c)?;
418        Ok(c)
419    }
420
421    /// Constrain that a+b=c in the emulated field.
422    /// Checking whether a + b = k * E::MODULUS + c
423    /// This function doesn't perform emulated variable validity check on the
424    /// input a, b and c. We assume that they are already performed elsewhere.
425    pub fn emulated_add_gate<E: EmulationConfig<F>>(
426        &mut self,
427        a: &EmulatedVariable<E>,
428        b: &EmulatedVariable<E>,
429        c: &EmulatedVariable<E>,
430    ) -> Result<(), CircuitError> {
431        self.check_vars_bound(&a.0)?;
432        self.check_vars_bound(&b.0)?;
433        self.check_vars_bound(&c.0)?;
434
435        let val_a: BigUint = self.emulated_witness(a)?.into();
436        let val_b: BigUint = self.emulated_witness(b)?.into();
437        let modulus: BigUint = E::MODULUS.into();
438        let b_pow = BigUint::from(2u32).pow(E::B as u32);
439        let add_no_mod = &val_a + &val_b;
440        let k = if add_no_mod >= modulus { 1u32 } else { 0u32 };
441        let var_k = self.create_boolean_variable(add_no_mod >= modulus)?.0;
442        let modulus_limbs = biguint_to_limbs::<F>(&modulus, E::B, E::NUM_LIMBS);
443
444        let add_no_mod_limbs = biguint_to_limbs::<F>(&add_no_mod, E::B, E::NUM_LIMBS)
445            .into_iter()
446            .map(|val| self.create_variable(val))
447            .collect::<Result<Vec<_>, CircuitError>>()?;
448
449        // Checking whether a + b = add_no_mod_limbs
450        let mut carry_out = self.zero();
451        for (a, b, c) in izip!(&a.0, &b.0, &add_no_mod_limbs) {
452            let next_carry_out =
453                F::from(<F as Into<BigUint>>::into(self.witness(*a)? + self.witness(*b)?) / &b_pow);
454            let next_carry_out = self.create_variable(next_carry_out)?;
455            self.enforce_bool(next_carry_out)?;
456
457            let wires = [*a, *b, carry_out, next_carry_out, *c];
458            let coeffs = [F::one(), F::one(), F::one(), -F::from(b_pow.clone())];
459            self.lc_gate(&wires, &coeffs)?;
460            carry_out = next_carry_out;
461
462            self.enforce_in_range(*c, E::B)?;
463        }
464
465        // Checking whether k * E::MODULUS + c = add_no_mod_limbs
466        carry_out = self.zero();
467        for (a, b, c) in izip!(modulus_limbs, &c.0, &add_no_mod_limbs) {
468            let next_carry_out =
469                F::from(<F as Into<BigUint>>::into(a * F::from(k) + self.witness(*b)?) / &b_pow);
470            let next_carry_out = self.create_variable(next_carry_out)?;
471            self.enforce_bool(next_carry_out)?;
472
473            let wires = [var_k, *b, carry_out, next_carry_out, *c];
474            let coeffs = [a, F::one(), F::one(), -F::from(b_pow.clone())];
475            self.lc_gate(&wires, &coeffs)?;
476            carry_out = next_carry_out;
477        }
478        Ok(())
479    }
480
481    /// Return an [`EmulatedVariable`] which equals to a+b.
482    pub fn emulated_add<E: EmulationConfig<F>>(
483        &mut self,
484        a: &EmulatedVariable<E>,
485        b: &EmulatedVariable<E>,
486    ) -> Result<EmulatedVariable<E>, CircuitError> {
487        let c = self.emulated_witness(a)? + self.emulated_witness(b)?;
488        let c = self.create_emulated_variable(c)?;
489        self.emulated_add_gate(a, b, &c)?;
490        Ok(c)
491    }
492
493    /// Return an [`EmulatedVariable`] which equals to a-b.
494    pub fn emulated_sub<E: EmulationConfig<F>>(
495        &mut self,
496        a: &EmulatedVariable<E>,
497        b: &EmulatedVariable<E>,
498    ) -> Result<EmulatedVariable<E>, CircuitError> {
499        let c = self.emulated_witness(a)? - self.emulated_witness(b)?;
500        let c = self.create_emulated_variable(c)?;
501        self.emulated_add_gate(&c, b, a)?;
502        Ok(c)
503    }
504
505    /// Constrain that a+b=c in the emulated field.
506    /// This function doesn't perform emulated variable validity check on the
507    /// input a and c. We assume that they are already performed elsewhere.
508    pub fn emulated_add_constant_gate<E: EmulationConfig<F>>(
509        &mut self,
510        a: &EmulatedVariable<E>,
511        b: E,
512        c: &EmulatedVariable<E>,
513    ) -> Result<(), CircuitError> {
514        self.check_vars_bound(&a.0)?;
515        self.check_vars_bound(&c.0)?;
516
517        let val_a: BigUint = self.emulated_witness(a)?.into();
518        let val_b: BigUint = b.into();
519        let q: BigUint = E::MODULUS.into();
520        let b_pow = BigUint::from(2u32).pow(E::B as u32);
521        let add_no_mod = &val_a + &val_b;
522        let k = if add_no_mod >= q { 1u32 } else { 0u32 };
523        let var_k = self.create_boolean_variable(add_no_mod >= q)?.0;
524        let q_limbs = biguint_to_limbs::<F>(&q, E::B, E::NUM_LIMBS);
525        let b_limbs = biguint_to_limbs::<F>(&val_b, E::B, E::NUM_LIMBS);
526
527        let add_no_mod_limbs = biguint_to_limbs::<F>(&add_no_mod, E::B, E::NUM_LIMBS)
528            .into_iter()
529            .map(|val| self.create_variable(val))
530            .collect::<Result<Vec<_>, CircuitError>>()?;
531
532        // Checking whether a + b = add_no_mod_limbs
533        let mut carry_out = self.zero();
534        for (a, b, c) in izip!(&a.0, b_limbs, &add_no_mod_limbs) {
535            let next_carry_out =
536                F::from(<F as Into<BigUint>>::into(self.witness(*a)? + b) / &b_pow);
537            let next_carry_out = self.create_variable(next_carry_out)?;
538            self.enforce_bool(next_carry_out)?;
539
540            let wires = [*a, self.one(), carry_out, next_carry_out, *c];
541            let coeffs = [F::one(), b, F::one(), -F::from(b_pow.clone())];
542            self.lc_gate(&wires, &coeffs)?;
543            carry_out = next_carry_out;
544
545            self.enforce_in_range(*c, E::B)?;
546        }
547
548        // Checking whether k * q + c = add_no_mod_limbs
549        carry_out = self.zero();
550        for (a, b, c) in izip!(q_limbs, &c.0, &add_no_mod_limbs) {
551            let next_carry_out =
552                F::from(<F as Into<BigUint>>::into(a * F::from(k) + self.witness(*b)?) / &b_pow);
553            let next_carry_out = self.create_variable(next_carry_out)?;
554            self.enforce_bool(next_carry_out)?;
555
556            let wires = [var_k, *b, carry_out, next_carry_out, *c];
557            let coeffs = [a, F::one(), F::one(), -F::from(b_pow.clone())];
558            self.lc_gate(&wires, &coeffs)?;
559            carry_out = next_carry_out;
560        }
561        Ok(())
562    }
563
564    /// Return an [`EmulatedVariable`] which equals to a + b where b is a
565    /// constant.
566    pub fn emulated_add_constant<E: EmulationConfig<F>>(
567        &mut self,
568        a: &EmulatedVariable<E>,
569        b: E,
570    ) -> Result<EmulatedVariable<E>, CircuitError> {
571        let c = self.emulated_witness(a)? + b;
572        let c = self.create_emulated_variable(c)?;
573        self.emulated_add_constant_gate(a, b, &c)?;
574        Ok(c)
575    }
576
577    /// Return an [`EmulatedVariable`] which equals to a - b where b is a
578    /// constant.
579    pub fn emulated_sub_constant<E: EmulationConfig<F>>(
580        &mut self,
581        a: &EmulatedVariable<E>,
582        b: E,
583    ) -> Result<EmulatedVariable<E>, CircuitError> {
584        let c = self.emulated_witness(a)? - b;
585        let c = self.create_emulated_variable(c)?;
586        self.emulated_add_constant_gate(&c, b, a)?;
587        Ok(c)
588    }
589    /// Obtain an emulated variable of the conditional selection from 2 emulated
590    /// variables. `b` is a boolean variable that indicates selection of P_b
591    /// from (P0, P1).
592    /// Return error if invalid input parameters are provided.
593    pub fn conditional_select_emulated<E: EmulationConfig<F>>(
594        &mut self,
595        b: BoolVar,
596        p0: &EmulatedVariable<E>,
597        p1: &EmulatedVariable<E>,
598    ) -> Result<EmulatedVariable<E>, CircuitError> {
599        self.check_var_bound(b.into())?;
600        self.check_vars_bound(&p0.0[..])?;
601        self.check_vars_bound(&p1.0[..])?;
602
603        let mut vals = vec![];
604        for (&x_0, &x_1) in p0.0.iter().zip(p1.0.iter()) {
605            let selected = self.conditional_select(b, x_0, x_1)?;
606            vals.push(selected);
607        }
608
609        Ok(EmulatedVariable::<E>(vals, PhantomData::<E>))
610    }
611
612    /// Constrain two emulated variables to be the same.
613    /// Return error if the input variables are invalid.
614    pub fn enforce_emulated_var_equal<E: EmulationConfig<F>>(
615        &mut self,
616        a: &EmulatedVariable<E>,
617        b: &EmulatedVariable<E>,
618    ) -> Result<(), CircuitError> {
619        self.check_vars_bound(&a.0[..])?;
620        self.check_vars_bound(&b.0[..])?;
621        for (&a, &b) in a.0.iter().zip(b.0.iter()) {
622            self.enforce_equal(a, b)?;
623        }
624        Ok(())
625    }
626
627    /// Obtain a bool variable representing whether two input emulated variables
628    /// are equal. Return error if variables are invalid.
629    pub fn is_emulated_var_equal<E: EmulationConfig<F>>(
630        &mut self,
631        a: &EmulatedVariable<E>,
632        b: &EmulatedVariable<E>,
633    ) -> Result<BoolVar, CircuitError> {
634        self.check_vars_bound(&a.0[..])?;
635        self.check_vars_bound(&b.0[..])?;
636        let c =
637            a.0.iter()
638                .zip(b.0.iter())
639                .map(|(&a, &b)| self.is_equal(a, b))
640                .collect::<Result<Vec<_>, _>>()?;
641        self.logic_and_all(&c)
642    }
643
644    /// Obtain a bool variable representing whether the input emulated variable
645    /// is zero. Return error if variables are invalid.
646    pub fn is_emulated_var_zero<E: EmulationConfig<F>>(
647        &mut self,
648        a: &EmulatedVariable<E>,
649    ) -> Result<BoolVar, CircuitError> {
650        self.check_vars_bound(&a.0[..])?;
651        let c =
652            a.0.iter()
653                .map(|&a| self.is_zero(a))
654                .collect::<Result<Vec<_>, _>>()?;
655        self.logic_and_all(&c)
656    }
657
658    /// Given an emulated field element `a`, return `a mod F::MODULUS` in the
659    /// native field.
660    fn mod_to_native_field<E: EmulationConfig<F>>(
661        &mut self,
662        a: &EmulatedVariable<E>,
663    ) -> Result<Variable, CircuitError> {
664        let b_pow = F::from(2u32).pow([E::B as u64]);
665        let double_b_pow = b_pow * b_pow;
666        let triple_b_pow = double_b_pow * b_pow;
667        let zero = self.zero();
668        let a0 = a.0.first().unwrap_or(&zero);
669        let a1 = a.0.get(1).unwrap_or(&zero);
670        let a2 = a.0.get(2).unwrap_or(&zero);
671        let a3 = a.0.get(3).unwrap_or(&zero);
672
673        let mut result = self.lc(
674            &[*a0, *a1, *a2, *a3],
675            &[F::one(), b_pow, double_b_pow, triple_b_pow],
676        )?;
677
678        if E::NUM_LIMBS > 4 {
679            let mut cur_pow = triple_b_pow * b_pow;
680            for i in (4..E::NUM_LIMBS).step_by(3) {
681                let a0 = a.0.get(i).unwrap_or(&zero);
682                let a1 = a.0.get(i + 1).unwrap_or(&zero);
683                let a2 = a.0.get(i + 2).unwrap_or(&zero);
684                result = self.lc(
685                    &[result, *a0, *a1, *a2],
686                    &[F::one(), cur_pow, cur_pow * b_pow, cur_pow * double_b_pow],
687                )?;
688                cur_pow *= triple_b_pow;
689            }
690        }
691        Ok(result)
692    }
693}
694
695impl EmulationConfig<ark_bn254::Fr> for ark_bls12_377::Fq {
696    const T: usize = 500;
697    const B: usize = 100;
698    const NUM_LIMBS: usize = 5;
699}
700
701impl EmulationConfig<ark_bn254::Fr> for ark_bn254::Fq {
702    const T: usize = 300;
703    const B: usize = 100;
704    const NUM_LIMBS: usize = 3;
705}
706
707#[cfg(test)]
708mod tests {
709    use super::EmulationConfig;
710    use crate::{gadgets::from_emulated_field, Circuit, PlonkCircuit};
711    use ark_bls12_377::Fq as Fq377;
712    use ark_bn254::{Fq as Fq254, Fr as Fr254};
713    use ark_ff::{MontFp, PrimeField};
714
715    #[test]
716    fn test_basics() {
717        test_basics_helper::<Fq377, Fr254>();
718        test_basics_helper::<Fq254, Fr254>();
719    }
720
721    fn test_basics_helper<E, F>()
722    where
723        E: EmulationConfig<F>,
724        F: PrimeField,
725    {
726        let mut circuit = PlonkCircuit::<F>::new_turbo_plonk();
727        let var_x = circuit.create_emulated_variable(E::one()).unwrap();
728        let overflow = E::from(F::MODULUS.into() * 2u64 + 1u64);
729        let var_y = circuit.create_emulated_variable(overflow).unwrap();
730        assert_eq!(circuit.emulated_witness(&var_x).unwrap(), E::one());
731        assert_eq!(circuit.emulated_witness(&var_y).unwrap(), overflow);
732        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
733    }
734
735    #[test]
736    fn test_emulated_add() {
737        test_emulated_add_helper::<Fq377, Fr254>();
738        test_emulated_add_helper::<Fq254, Fr254>();
739    }
740
741    fn test_emulated_add_helper<E, F>()
742    where
743        E: EmulationConfig<F>,
744        F: PrimeField,
745    {
746        let mut circuit = PlonkCircuit::<F>::new_turbo_plonk();
747        let var_x = circuit.create_public_emulated_variable(E::one()).unwrap();
748        let overflow = E::from(E::MODULUS.into() - 1u64);
749        let var_y = circuit.create_emulated_variable(overflow).unwrap();
750        let var_z = circuit.emulated_add(&var_x, &var_y).unwrap();
751        assert_eq!(circuit.emulated_witness(&var_x).unwrap(), E::one());
752        assert_eq!(circuit.emulated_witness(&var_y).unwrap(), overflow);
753        assert_eq!(circuit.emulated_witness(&var_z).unwrap(), E::zero());
754
755        let var_z = circuit.emulated_add_constant(&var_z, overflow).unwrap();
756        assert_eq!(circuit.emulated_witness(&var_z).unwrap(), overflow);
757
758        let x = from_emulated_field(E::one());
759        assert!(circuit.check_circuit_satisfiability(&x).is_ok());
760
761        let var_z = circuit.create_emulated_variable(E::one()).unwrap();
762        circuit.emulated_add_gate(&var_x, &var_y, &var_z).unwrap();
763        assert!(circuit.check_circuit_satisfiability(&x).is_err());
764    }
765
766    #[test]
767    fn test_emulated_mul() {
768        test_emulated_mul_helper::<Fq377, Fr254>();
769        test_emulated_mul_helper::<Fq254, Fr254>();
770
771        // test for issue (https://github.com/EspressoSystems/jellyfish/issues/306)
772        let x : Fq377= MontFp!("218393408942992446968589193493746660101651787560689350338764189588519393175121782177906966561079408675464506489966");
773        let y : Fq377 = MontFp!("122268283598675559488486339158635529096981886914877139579534153582033676785385790730042363341236035746924960903179");
774
775        let mut circuit = PlonkCircuit::<Fr254>::new_turbo_plonk();
776        let var_x = circuit.create_emulated_variable(x).unwrap();
777        let _ = circuit.emulated_mul_constant(&var_x, y).unwrap();
778        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
779    }
780
781    fn test_emulated_mul_helper<E, F>()
782    where
783        E: EmulationConfig<F>,
784        F: PrimeField,
785    {
786        let mut circuit = PlonkCircuit::<F>::new_ultra_plonk(20);
787        let x = E::from(6732u64);
788        let y = E::from(E::MODULUS.into() - 12387u64);
789        let expected = x * y;
790        let var_x = circuit.create_public_emulated_variable(x).unwrap();
791        let var_y = circuit.create_emulated_variable(y).unwrap();
792        let var_z = circuit.emulated_mul(&var_x, &var_y).unwrap();
793        assert_eq!(circuit.emulated_witness(&var_x).unwrap(), x);
794        assert_eq!(circuit.emulated_witness(&var_y).unwrap(), y);
795        assert_eq!(circuit.emulated_witness(&var_z).unwrap(), expected);
796        assert!(circuit
797            .check_circuit_satisfiability(&from_emulated_field(x))
798            .is_ok());
799
800        let var_y_z = circuit.emulated_mul(&var_y, &var_z).unwrap();
801        assert_eq!(circuit.emulated_witness(&var_y_z).unwrap(), expected * y);
802        assert!(circuit
803            .check_circuit_satisfiability(&from_emulated_field(x))
804            .is_ok());
805
806        let var_z = circuit.emulated_mul_constant(&var_z, expected).unwrap();
807        assert_eq!(
808            circuit.emulated_witness(&var_z).unwrap(),
809            expected * expected
810        );
811        assert!(circuit
812            .check_circuit_satisfiability(&from_emulated_field(x))
813            .is_ok());
814
815        let var_z = circuit.create_emulated_variable(E::one()).unwrap();
816        circuit.emulated_mul_gate(&var_x, &var_y, &var_z).unwrap();
817        assert!(circuit
818            .check_circuit_satisfiability(&from_emulated_field(x))
819            .is_err());
820    }
821
822    #[test]
823    fn test_select() {
824        test_select_helper::<Fq377, Fr254>();
825        test_select_helper::<Fq254, Fr254>();
826    }
827
828    fn test_select_helper<E, F>()
829    where
830        E: EmulationConfig<F>,
831        F: PrimeField,
832    {
833        let mut circuit = PlonkCircuit::<F>::new_turbo_plonk();
834        let var_x = circuit.create_emulated_variable(E::one()).unwrap();
835        let overflow = E::from(E::MODULUS.into() - 1u64);
836        let var_y = circuit.create_emulated_variable(overflow).unwrap();
837        let b = circuit.create_boolean_variable(true).unwrap();
838        let var_z = circuit
839            .conditional_select_emulated(b, &var_x, &var_y)
840            .unwrap();
841        assert_eq!(circuit.emulated_witness(&var_z).unwrap(), overflow);
842        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
843        *circuit.witness_mut(var_z.0[0]) = F::zero();
844        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
845    }
846
847    #[test]
848    fn test_enforce_equal() {
849        test_enforce_equal_helper::<Fq377, Fr254>();
850        test_enforce_equal_helper::<Fq254, Fr254>();
851    }
852
853    fn test_enforce_equal_helper<E, F>()
854    where
855        E: EmulationConfig<F>,
856        F: PrimeField,
857    {
858        let mut circuit = PlonkCircuit::<F>::new_turbo_plonk();
859        let var_x = circuit.create_emulated_variable(E::one()).unwrap();
860        let overflow = E::from(E::MODULUS.into() - 1u64);
861        let var_y = circuit.create_emulated_variable(overflow).unwrap();
862        let var_z = circuit.create_emulated_variable(overflow).unwrap();
863        circuit.enforce_emulated_var_equal(&var_y, &var_z).unwrap();
864        assert!(circuit.check_circuit_satisfiability(&[]).is_ok());
865        circuit.enforce_emulated_var_equal(&var_x, &var_y).unwrap();
866        assert!(circuit.check_circuit_satisfiability(&[]).is_err());
867    }
868}