sapling_crypto_ce/circuit/sapling/
mod.rs

1// use bellman::pairing::{
2//
3// };
4
5use bellman::pairing::ff::{
6    PrimeField,
7    PrimeFieldRepr,
8    Field,
9};
10
11use bellman::{
12    SynthesisError,
13    ConstraintSystem,
14    Circuit
15};
16
17use jubjub::{
18    JubjubEngine,
19    FixedGenerators
20};
21
22use constants;
23
24use primitives::{
25    ValueCommitment,
26    ProofGenerationKey,
27    PaymentAddress
28};
29
30use super::Assignment;
31use super::boolean;
32use super::ecc;
33use super::pedersen_hash;
34use super::blake2s;
35use super::num;
36use super::multipack;
37
38/// This is an instance of the `Spend` circuit.
39pub struct Spend<'a, E: JubjubEngine> {
40    pub params: &'a E::Params,
41
42    /// Pedersen commitment to the value being spent
43    pub value_commitment: Option<ValueCommitment<E>>,
44
45    /// Key required to construct proofs for spending notes
46    /// for a particular spending key
47    pub proof_generation_key: Option<ProofGenerationKey<E>>,
48
49    /// The payment address associated with the note
50    pub payment_address: Option<PaymentAddress<E>>,
51
52    /// The randomness of the note commitment
53    pub commitment_randomness: Option<E::Fs>,
54
55    /// Re-randomization of the public key
56    pub ar: Option<E::Fs>,
57
58    /// The authentication path of the commitment in the tree
59    pub auth_path: Vec<Option<(E::Fr, bool)>>,
60
61    /// The anchor; the root of the tree. If the note being
62    /// spent is zero-value, this can be anything.
63    pub anchor: Option<E::Fr>
64}
65
66/// This is an output circuit instance.
67pub struct Output<'a, E: JubjubEngine> {
68    pub params: &'a E::Params,
69
70    /// Pedersen commitment to the value being spent
71    pub value_commitment: Option<ValueCommitment<E>>,
72
73    /// The payment address of the recipient
74    pub payment_address: Option<PaymentAddress<E>>,
75
76    /// The randomness used to hide the note commitment data
77    pub commitment_randomness: Option<E::Fs>,
78
79    /// The ephemeral secret key for DH with recipient
80    pub esk: Option<E::Fs>
81}
82
83/// Exposes a Pedersen commitment to the value as an
84/// input to the circuit
85fn expose_value_commitment<E, CS>(
86    mut cs: CS,
87    value_commitment: Option<ValueCommitment<E>>,
88    params: &E::Params
89) -> Result<Vec<boolean::Boolean>, SynthesisError>
90    where E: JubjubEngine,
91          CS: ConstraintSystem<E>
92{
93    // Booleanize the value into little-endian bit order
94    let value_bits = boolean::u64_into_boolean_vec_le(
95        cs.namespace(|| "value"),
96        value_commitment.as_ref().map(|c| c.value)
97    )?;
98
99    // Compute the note value in the exponent
100    let value = ecc::fixed_base_multiplication(
101        cs.namespace(|| "compute the value in the exponent"),
102        FixedGenerators::ValueCommitmentValue,
103        &value_bits,
104        params
105    )?;
106
107    // Booleanize the randomness. This does not ensure
108    // the bit representation is "in the field" because
109    // it doesn't matter for security.
110    let rcv = boolean::field_into_boolean_vec_le(
111        cs.namespace(|| "rcv"),
112        value_commitment.as_ref().map(|c| c.randomness)
113    )?;
114
115    // Compute the randomness in the exponent
116    let rcv = ecc::fixed_base_multiplication(
117        cs.namespace(|| "computation of rcv"),
118        FixedGenerators::ValueCommitmentRandomness,
119        &rcv,
120        params
121    )?;
122
123    // Compute the Pedersen commitment to the value
124    let cv = value.add(
125        cs.namespace(|| "computation of cv"),
126        &rcv,
127        params
128    )?;
129
130    // Expose the commitment as an input to the circuit
131    cv.inputize(cs.namespace(|| "commitment point"))?;
132
133    Ok(value_bits)
134}
135
136impl<'a, E: JubjubEngine> Circuit<E> for Spend<'a, E> {
137    fn synthesize<CS: ConstraintSystem<E>>(self, cs: &mut CS) -> Result<(), SynthesisError>
138    {
139        // Prover witnesses ak (ensures that it's on the curve)
140        let ak = ecc::EdwardsPoint::witness(
141            cs.namespace(|| "ak"),
142            self.proof_generation_key.as_ref().map(|k| k.ak.clone()),
143            self.params
144        )?;
145
146        // There are no sensible attacks on small order points
147        // of ak (that we're aware of!) but it's a cheap check,
148        // so we do it.
149        ak.assert_not_small_order(
150            cs.namespace(|| "ak not small order"),
151            self.params
152        )?;
153
154        // Rerandomize ak and expose it as an input to the circuit
155        {
156            let ar = boolean::field_into_boolean_vec_le(
157                cs.namespace(|| "ar"),
158                self.ar
159            )?;
160
161            // Compute the randomness in the exponent
162            let ar = ecc::fixed_base_multiplication(
163                cs.namespace(|| "computation of randomization for the signing key"),
164                FixedGenerators::SpendingKeyGenerator,
165                &ar,
166                self.params
167            )?;
168
169            let rk = ak.add(
170                cs.namespace(|| "computation of rk"),
171                &ar,
172                self.params
173            )?;
174
175            rk.inputize(cs.namespace(|| "rk"))?;
176        }
177
178        // Compute nk = [nsk] ProofGenerationKey
179        let nk;
180        {
181            // Witness nsk as bits
182            let nsk = boolean::field_into_boolean_vec_le(
183                cs.namespace(|| "nsk"),
184                self.proof_generation_key.as_ref().map(|k| k.nsk.clone())
185            )?;
186
187            // NB: We don't ensure that the bit representation of nsk
188            // is "in the field" (Fs) because it's not used except to
189            // demonstrate the prover knows it. If they know a
190            // congruency then that's equivalent.
191
192            // Compute nk = [nsk] ProvingPublicKey
193            nk = ecc::fixed_base_multiplication(
194                cs.namespace(|| "computation of nk"),
195                FixedGenerators::ProofGenerationKey,
196                &nsk,
197                self.params
198            )?;
199        }
200
201        // This is the "viewing key" preimage for CRH^ivk
202        let mut ivk_preimage = vec![];
203
204        // Place ak in the preimage for CRH^ivk
205        ivk_preimage.extend(
206            ak.repr(cs.namespace(|| "representation of ak"))?
207        );
208
209        // This is the nullifier preimage for PRF^nf
210        let mut nf_preimage = vec![];
211
212        // Extend ivk and nf preimages with the representation of
213        // nk.
214        {
215            let repr_nk = nk.repr(
216                cs.namespace(|| "representation of nk")
217            )?;
218
219            ivk_preimage.extend(repr_nk.iter().cloned());
220            nf_preimage.extend(repr_nk);
221        }
222
223        assert_eq!(ivk_preimage.len(), 512);
224        assert_eq!(nf_preimage.len(), 256);
225
226        // Compute the incoming viewing key ivk
227        let mut ivk = blake2s::blake2s(
228            cs.namespace(|| "computation of ivk"),
229            &ivk_preimage,
230            constants::CRH_IVK_PERSONALIZATION
231        )?;
232
233        // drop_5 to ensure it's in the field
234        ivk.truncate(E::Fs::CAPACITY as usize);
235
236        // Witness g_d, checking that it's on the curve.
237        let g_d = {
238            // This binding is to avoid a weird edge case in Rust's
239            // ownership/borrowing rules. self is partially moved
240            // above, but the closure for and_then will have to
241            // move self (or a reference to self) to reference
242            // self.params, so we have to copy self.params here.
243            let params = self.params;
244
245            ecc::EdwardsPoint::witness(
246                cs.namespace(|| "witness g_d"),
247                self.payment_address.as_ref().and_then(|a| a.g_d(params)),
248                self.params
249            )?
250        };
251
252        // Check that g_d is not small order. Technically, this check
253        // is already done in the Output circuit, and this proof ensures
254        // g_d is bound to a product of that check, but for defense in
255        // depth let's check it anyway. It's cheap.
256        g_d.assert_not_small_order(
257            cs.namespace(|| "g_d not small order"),
258            self.params
259        )?;
260
261        // Compute pk_d = g_d^ivk
262        let pk_d = g_d.mul(
263            cs.namespace(|| "compute pk_d"),
264            &ivk,
265            self.params
266        )?;
267
268        // Compute note contents:
269        // value (in big endian) followed by g_d and pk_d
270        let mut note_contents = vec![];
271
272        // Handle the value; we'll need it later for the
273        // dummy input check.
274        let mut value_num = num::Num::zero();
275        {
276            // Get the value in little-endian bit order
277            let value_bits = expose_value_commitment(
278                cs.namespace(|| "value commitment"),
279                self.value_commitment,
280                self.params
281            )?;
282
283            // Compute the note's value as a linear combination
284            // of the bits.
285            let mut coeff = E::Fr::one();
286            for bit in &value_bits {
287                value_num = value_num.add_bool_with_coeff(
288                    CS::one(),
289                    bit,
290                    coeff
291                );
292                coeff.double();
293            }
294
295            // Place the value in the note
296            note_contents.extend(value_bits);
297        }
298
299        // Place g_d in the note
300        note_contents.extend(
301            g_d.repr(cs.namespace(|| "representation of g_d"))?
302        );
303
304        // Place pk_d in the note
305        note_contents.extend(
306            pk_d.repr(cs.namespace(|| "representation of pk_d"))?
307        );
308
309        assert_eq!(
310            note_contents.len(),
311            64 + // value
312            256 + // g_d
313            256 // p_d
314        );
315
316        // Compute the hash of the note contents
317        let mut cm = pedersen_hash::pedersen_hash(
318            cs.namespace(|| "note content hash"),
319            pedersen_hash::Personalization::NoteCommitment,
320            &note_contents,
321            self.params
322        )?;
323
324        {
325            // Booleanize the randomness for the note commitment
326            let rcm = boolean::field_into_boolean_vec_le(
327                cs.namespace(|| "rcm"),
328                self.commitment_randomness
329            )?;
330
331            // Compute the note commitment randomness in the exponent
332            let rcm = ecc::fixed_base_multiplication(
333                cs.namespace(|| "computation of commitment randomness"),
334                FixedGenerators::NoteCommitmentRandomness,
335                &rcm,
336                self.params
337            )?;
338
339            // Randomize the note commitment. Pedersen hashes are not
340            // themselves hiding commitments.
341            cm = cm.add(
342                cs.namespace(|| "randomization of note commitment"),
343                &rcm,
344                self.params
345            )?;
346        }
347
348        // This will store (least significant bit first)
349        // the position of the note in the tree, for use
350        // in nullifier computation.
351        let mut position_bits = vec![];
352
353        // This is an injective encoding, as cur is a
354        // point in the prime order subgroup.
355        let mut cur = cm.get_x().clone();
356
357        // Ascend the merkle tree authentication path
358        for (i, e) in self.auth_path.into_iter().enumerate() {
359            let cs = &mut cs.namespace(|| format!("merkle tree hash {}", i));
360
361            // Determines if the current subtree is the "right" leaf at this
362            // depth of the tree.
363            let cur_is_right = boolean::Boolean::from(boolean::AllocatedBit::alloc(
364                cs.namespace(|| "position bit"),
365                e.map(|e| e.1)
366            )?);
367
368            // Push this boolean for nullifier computation later
369            position_bits.push(cur_is_right.clone());
370
371            // Witness the authentication path element adjacent
372            // at this depth.
373            let path_element = num::AllocatedNum::alloc(
374                cs.namespace(|| "path element"),
375                || {
376                    Ok(e.get()?.0)
377                }
378            )?;
379
380            // Swap the two if the current subtree is on the right
381            let (xl, xr) = num::AllocatedNum::conditionally_reverse(
382                cs.namespace(|| "conditional reversal of preimage"),
383                &cur,
384                &path_element,
385                &cur_is_right
386            )?;
387
388            // We don't need to be strict, because the function is
389            // collision-resistant. If the prover witnesses a congruency,
390            // they will be unable to find an authentication path in the
391            // tree with high probability.
392            let mut preimage = vec![];
393            preimage.extend(xl.into_bits_le(cs.namespace(|| "xl into bits"))?);
394            preimage.extend(xr.into_bits_le(cs.namespace(|| "xr into bits"))?);
395
396            // Compute the new subtree value
397            cur = pedersen_hash::pedersen_hash(
398                cs.namespace(|| "computation of pedersen hash"),
399                pedersen_hash::Personalization::MerkleTree(i),
400                &preimage,
401                self.params
402            )?.get_x().clone(); // Injective encoding
403        }
404
405        {
406            let real_anchor_value = self.anchor;
407
408            // Allocate the "real" anchor that will be exposed.
409            let rt = num::AllocatedNum::alloc(
410                cs.namespace(|| "conditional anchor"),
411                || {
412                    Ok(*real_anchor_value.get()?)
413                }
414            )?;
415
416            // (cur - rt) * value = 0
417            // if value is zero, cur and rt can be different
418            // if value is nonzero, they must be equal
419            cs.enforce(
420                || "conditionally enforce correct root",
421                |lc| lc + cur.get_variable() - rt.get_variable(),
422                |lc| lc + &value_num.lc(E::Fr::one()),
423                |lc| lc
424            );
425
426            // Expose the anchor
427            rt.inputize(cs.namespace(|| "anchor"))?;
428        }
429
430        // Compute the cm + g^position for preventing
431        // faerie gold attacks
432        let mut rho = cm;
433        {
434            // Compute the position in the exponent
435            let position = ecc::fixed_base_multiplication(
436                cs.namespace(|| "g^position"),
437                FixedGenerators::NullifierPosition,
438                &position_bits,
439                self.params
440            )?;
441
442            // Add the position to the commitment
443            rho = rho.add(
444                cs.namespace(|| "faerie gold prevention"),
445                &position,
446                self.params
447            )?;
448        }
449        
450        // Let's compute nf = BLAKE2s(nk || rho)
451        nf_preimage.extend(
452            rho.repr(cs.namespace(|| "representation of rho"))?
453        );
454
455        assert_eq!(nf_preimage.len(), 512);
456        
457        // Compute nf
458        let nf = blake2s::blake2s(
459            cs.namespace(|| "nf computation"),
460            &nf_preimage,
461            constants::PRF_NF_PERSONALIZATION
462        )?;
463
464        multipack::pack_into_inputs(cs.namespace(|| "pack nullifier"), &nf)
465    }
466}
467
468impl<'a, E: JubjubEngine> Circuit<E> for Output<'a, E> {
469    fn synthesize<CS: ConstraintSystem<E>>(self, cs: &mut CS) -> Result<(), SynthesisError>
470    {
471        // Let's start to construct our note, which contains
472        // value (big endian)
473        let mut note_contents = vec![];
474
475        // Expose the value commitment and place the value
476        // in the note.
477        note_contents.extend(expose_value_commitment(
478            cs.namespace(|| "value commitment"),
479            self.value_commitment,
480            self.params
481        )?);
482
483        // Let's deal with g_d
484        {
485            let params = self.params;
486
487            // Prover witnesses g_d, ensuring it's on the
488            // curve.
489            let g_d = ecc::EdwardsPoint::witness(
490                cs.namespace(|| "witness g_d"),
491                self.payment_address.as_ref().and_then(|a| a.g_d(params)),
492                self.params
493            )?;
494
495            // g_d is ensured to be large order. The relationship
496            // between g_d and pk_d ultimately binds ivk to the
497            // note. If this were a small order point, it would
498            // not do this correctly, and the prover could
499            // double-spend by finding random ivk's that satisfy
500            // the relationship.
501            //
502            // Further, if it were small order, epk would be
503            // small order too!
504            g_d.assert_not_small_order(
505                cs.namespace(|| "g_d not small order"),
506                self.params
507            )?;
508
509            // Extend our note contents with the representation of
510            // g_d.
511            note_contents.extend(
512                g_d.repr(cs.namespace(|| "representation of g_d"))?
513            );
514
515            // Booleanize our ephemeral secret key
516            let esk = boolean::field_into_boolean_vec_le(
517                cs.namespace(|| "esk"),
518                self.esk
519            )?;
520
521            // Create the ephemeral public key from g_d.
522            let epk = g_d.mul(
523                cs.namespace(|| "epk computation"),
524                &esk,
525                self.params
526            )?;
527
528            // Expose epk publicly.
529            epk.inputize(cs.namespace(|| "epk"))?;
530        }
531
532        // Now let's deal with pk_d. We don't do any checks and
533        // essentially allow the prover to witness any 256 bits
534        // they would like.
535        {
536            // Just grab pk_d from the witness
537            let pk_d = self.payment_address.as_ref().map(|e| e.pk_d.into_xy());
538
539            // Witness the y-coordinate, encoded as little
540            // endian bits (to match the representation)
541            let y_contents = boolean::field_into_boolean_vec_le(
542                cs.namespace(|| "pk_d bits of y"),
543                pk_d.map(|e| e.1)
544            )?;
545
546            // Witness the sign bit
547            let sign_bit = boolean::Boolean::from(boolean::AllocatedBit::alloc(
548                cs.namespace(|| "pk_d bit of x"),
549                pk_d.map(|e| e.0.into_repr().is_odd())
550            )?);
551
552            // Extend the note with pk_d representation
553            note_contents.extend(y_contents);
554            note_contents.push(sign_bit);
555        }
556
557        assert_eq!(
558            note_contents.len(),
559            64 + // value
560            256 + // g_d
561            256 // pk_d
562        );
563
564        // Compute the hash of the note contents
565        let mut cm = pedersen_hash::pedersen_hash(
566            cs.namespace(|| "note content hash"),
567            pedersen_hash::Personalization::NoteCommitment,
568            &note_contents,
569            self.params
570        )?;
571
572        {
573            // Booleanize the randomness
574            let rcm = boolean::field_into_boolean_vec_le(
575                cs.namespace(|| "rcm"),
576                self.commitment_randomness
577            )?;
578
579            // Compute the note commitment randomness in the exponent
580            let rcm = ecc::fixed_base_multiplication(
581                cs.namespace(|| "computation of commitment randomness"),
582                FixedGenerators::NoteCommitmentRandomness,
583                &rcm,
584                self.params
585            )?;
586
587            // Randomize our note commitment
588            cm = cm.add(
589                cs.namespace(|| "randomization of note commitment"),
590                &rcm,
591                self.params
592            )?;
593        }
594
595        // Only the x-coordinate of the output is revealed,
596        // since we know it is prime order, and we know that
597        // the x-coordinate is an injective encoding for
598        // prime-order elements.
599        cm.get_x().inputize(cs.namespace(|| "commitment"))?;
600
601        Ok(())
602    }
603}
604
605#[test]
606fn test_input_circuit_with_bls12_381() {
607    use bellman::pairing::ff::{Field, BitIterator};
608    use bellman::pairing::bls12_381::*;
609    use rand::{SeedableRng, Rng, XorShiftRng};
610    use ::circuit::test::*;
611    use jubjub::{JubjubBls12, fs, edwards};
612
613    let params = &JubjubBls12::new();
614    let rng = &mut XorShiftRng::from_seed([0x3dbe6259, 0x8d313d76, 0x3237db17, 0xe5bc0654]);
615
616    let tree_depth = 32;
617
618    for _ in 0..10 {
619        let value_commitment = ValueCommitment {
620            value: rng.gen(),
621            randomness: rng.gen()
622        };
623
624        let nsk: fs::Fs = rng.gen();
625        let ak = edwards::Point::rand(rng, params).mul_by_cofactor(params);
626
627        let proof_generation_key = ::primitives::ProofGenerationKey {
628            ak: ak.clone(),
629            nsk: nsk.clone()
630        };
631
632        let viewing_key = proof_generation_key.into_viewing_key(params);
633
634        let payment_address;
635
636        loop {
637            let diversifier = ::primitives::Diversifier(rng.gen());
638
639            if let Some(p) = viewing_key.into_payment_address(
640                diversifier,
641                params
642            )
643            {
644                payment_address = p;
645                break;
646            }
647        }
648
649        let g_d = payment_address.diversifier.g_d(params).unwrap();
650        let commitment_randomness: fs::Fs = rng.gen();
651        let auth_path = vec![Some((rng.gen(), rng.gen())); tree_depth];
652        let ar: fs::Fs = rng.gen();
653
654        {
655            let rk = viewing_key.rk(ar, params).into_xy();
656            let expected_value_cm = value_commitment.cm(params).into_xy();
657            let note = ::primitives::Note {
658                value: value_commitment.value,
659                g_d: g_d.clone(),
660                pk_d: payment_address.pk_d.clone(),
661                r: commitment_randomness.clone()
662            };
663
664            let mut position = 0u64;
665            let cm: Fr = note.cm(params);
666            let mut cur = cm.clone();
667
668            for (i, val) in auth_path.clone().into_iter().enumerate()
669            {
670                let (uncle, b) = val.unwrap();
671
672                let mut lhs = cur;
673                let mut rhs = uncle;
674
675                if b {
676                    ::std::mem::swap(&mut lhs, &mut rhs);
677                }
678
679                let mut lhs: Vec<bool> = BitIterator::new(lhs.into_repr()).collect();
680                let mut rhs: Vec<bool> = BitIterator::new(rhs.into_repr()).collect();
681
682                lhs.reverse();
683                rhs.reverse();
684
685                cur = ::pedersen_hash::pedersen_hash::<Bls12, _>(
686                    ::pedersen_hash::Personalization::MerkleTree(i),
687                    lhs.into_iter()
688                       .take(Fr::NUM_BITS as usize)
689                       .chain(rhs.into_iter().take(Fr::NUM_BITS as usize)),
690                    params
691                ).into_xy().0;
692
693                if b {
694                    position |= 1 << i;
695                }
696            }
697
698            let expected_nf = note.nf(&viewing_key, position, params);
699            let expected_nf = multipack::bytes_to_bits_le(&expected_nf);
700            let expected_nf = multipack::compute_multipacking::<Bls12>(&expected_nf);
701            assert_eq!(expected_nf.len(), 2);
702
703            let mut cs = TestConstraintSystem::<Bls12>::new();
704
705            let instance = Spend {
706                params: params,
707                value_commitment: Some(value_commitment.clone()),
708                proof_generation_key: Some(proof_generation_key.clone()),
709                payment_address: Some(payment_address.clone()),
710                commitment_randomness: Some(commitment_randomness),
711                ar: Some(ar),
712                auth_path: auth_path.clone(),
713                anchor: Some(cur)
714            };
715
716            instance.synthesize(&mut cs).unwrap();
717
718            assert!(cs.is_satisfied());
719            assert_eq!(cs.num_constraints(), 98777);
720            assert_eq!(cs.hash(), "d37c738e83df5d9b0bb6495ac96abf21bcb2697477e2c15c2c7916ff7a3b6a89");
721
722            assert_eq!(cs.get("randomization of note commitment/x3/num"), cm);
723
724            assert_eq!(cs.num_inputs(), 8);
725            assert_eq!(cs.get_input(0, "ONE"), Fr::one());
726            assert_eq!(cs.get_input(1, "rk/x/input variable"), rk.0);
727            assert_eq!(cs.get_input(2, "rk/y/input variable"), rk.1);
728            assert_eq!(cs.get_input(3, "value commitment/commitment point/x/input variable"), expected_value_cm.0);
729            assert_eq!(cs.get_input(4, "value commitment/commitment point/y/input variable"), expected_value_cm.1);
730            assert_eq!(cs.get_input(5, "anchor/input variable"), cur);
731            assert_eq!(cs.get_input(6, "pack nullifier/input 0"), expected_nf[0]);
732            assert_eq!(cs.get_input(7, "pack nullifier/input 1"), expected_nf[1]);
733        }
734    }
735}
736
737#[test]
738fn test_output_circuit_with_bls12_381() {
739    use bellman::pairing::ff::{Field};
740    use bellman::pairing::bls12_381::*;
741    use rand::{SeedableRng, Rng, XorShiftRng};
742    use ::circuit::test::*;
743    use jubjub::{JubjubBls12, fs, edwards};
744
745    let params = &JubjubBls12::new();
746    let rng = &mut XorShiftRng::from_seed([0x3dbe6258, 0x8d313d76, 0x3237db17, 0xe5bc0654]);
747
748    for _ in 0..100 {
749        let value_commitment = ValueCommitment {
750            value: rng.gen(),
751            randomness: rng.gen()
752        };
753
754        let nsk: fs::Fs = rng.gen();
755        let ak = edwards::Point::rand(rng, params).mul_by_cofactor(params);
756
757        let proof_generation_key = ::primitives::ProofGenerationKey {
758            ak: ak.clone(),
759            nsk: nsk.clone()
760        };
761
762        let viewing_key = proof_generation_key.into_viewing_key(params);
763
764        let payment_address;
765
766        loop {
767            let diversifier = ::primitives::Diversifier(rng.gen());
768
769            if let Some(p) = viewing_key.into_payment_address(
770                diversifier,
771                params
772            )
773            {
774                payment_address = p;
775                break;
776            }
777        }
778
779        let commitment_randomness: fs::Fs = rng.gen();
780        let esk: fs::Fs = rng.gen();
781
782        {
783            let mut cs = TestConstraintSystem::<Bls12>::new();
784
785            let instance = Output {
786                params: params,
787                value_commitment: Some(value_commitment.clone()),
788                payment_address: Some(payment_address.clone()),
789                commitment_randomness: Some(commitment_randomness),
790                esk: Some(esk.clone())
791            };
792
793            instance.synthesize(&mut cs).unwrap();
794
795            assert!(cs.is_satisfied());
796            assert_eq!(cs.num_constraints(), 7827);
797            assert_eq!(cs.hash(), "c26d5cdfe6ccd65c03390902c02e11393ea6bb96aae32a7f2ecb12eb9103faee");
798
799            let expected_cm = payment_address.create_note(
800                value_commitment.value,
801                commitment_randomness,
802                params
803            ).expect("should be valid").cm(params);
804
805            let expected_value_cm = value_commitment.cm(params).into_xy();
806
807            let expected_epk = payment_address.g_d(params).expect("should be valid").mul(esk, params);
808            let expected_epk_xy = expected_epk.into_xy();
809
810            assert_eq!(cs.num_inputs(), 6);
811            assert_eq!(cs.get_input(0, "ONE"), Fr::one());
812            assert_eq!(cs.get_input(1, "value commitment/commitment point/x/input variable"), expected_value_cm.0);
813            assert_eq!(cs.get_input(2, "value commitment/commitment point/y/input variable"), expected_value_cm.1);
814            assert_eq!(cs.get_input(3, "epk/x/input variable"), expected_epk_xy.0);
815            assert_eq!(cs.get_input(4, "epk/y/input variable"), expected_epk_xy.1);
816            assert_eq!(cs.get_input(5, "commitment/input variable"), expected_cm);
817        }
818    }
819}