curve25519_dalek/
montgomery.rs

1// -*- mode: rust; -*-
2//
3// This file is part of curve25519-dalek.
4// Copyright (c) 2016-2021 isis lovecruft
5// Copyright (c) 2016-2019 Henry de Valence
6// See LICENSE for licensing information.
7//
8// Authors:
9// - isis agora lovecruft <isis@patternsinthevoid.net>
10// - Henry de Valence <hdevalence@hdevalence.ca>
11
12//! Scalar multiplication on the Montgomery form of Curve25519.
13//!
14//! To avoid notational confusion with the Edwards code, we use
15//! variables \\( u, v \\) for the Montgomery curve, so that “Montgomery
16//! \\(u\\)” here corresponds to “Montgomery \\(x\\)” elsewhere.
17//!
18//! Montgomery arithmetic works not on the curve itself, but on the
19//! \\(u\\)-line, which discards sign information and unifies the curve
20//! and its quadratic twist.  See [_Montgomery curves and their
21//! arithmetic_][costello-smith] by Costello and Smith for more details.
22//!
23//! The `MontgomeryPoint` struct contains the affine \\(u\\)-coordinate
24//! \\(u\_0(P)\\) of a point \\(P\\) on either the curve or the twist.
25//! Here the map \\(u\_0 : \mathcal M \rightarrow \mathbb F\_p \\) is
26//! defined by \\(u\_0((u,v)) = u\\); \\(u\_0(\mathcal O) = 0\\).  See
27//! section 5.4 of Costello-Smith for more details.
28//!
29//! # Scalar Multiplication
30//!
31//! Scalar multiplication on `MontgomeryPoint`s is provided by the `*`
32//! operator, which implements the Montgomery ladder.
33//!
34//! # Edwards Conversion
35//!
36//! The \\(2\\)-to-\\(1\\) map from the Edwards model to the Montgomery
37//! \\(u\\)-line is provided by `EdwardsPoint::to_montgomery()`.
38//!
39//! To lift a `MontgomeryPoint` to an `EdwardsPoint`, use
40//! `MontgomeryPoint::to_edwards()`, which takes a sign parameter.
41//! This function rejects `MontgomeryPoints` which correspond to points
42//! on the twist.
43//!
44//! [costello-smith]: https://eprint.iacr.org/2017/212.pdf
45
46// We allow non snake_case names because coordinates in projective space are
47// traditionally denoted by the capitalisation of their respective
48// counterparts in affine space.  Yeah, you heard me, rustc, I'm gonna have my
49// affine and projective cakes and eat both of them too.
50#![allow(non_snake_case)]
51
52use core::{
53    hash::{Hash, Hasher},
54    ops::{Mul, MulAssign},
55};
56
57use crate::constants::{APLUS2_OVER_FOUR, MONTGOMERY_A, MONTGOMERY_A_NEG};
58use crate::edwards::{CompressedEdwardsY, EdwardsPoint};
59use crate::field::FieldElement;
60use crate::scalar::{clamp_integer, Scalar};
61
62use crate::traits::Identity;
63
64use subtle::Choice;
65use subtle::ConstantTimeEq;
66use subtle::{ConditionallyNegatable, ConditionallySelectable};
67
68#[cfg(feature = "zeroize")]
69use zeroize::Zeroize;
70
71/// Holds the \\(u\\)-coordinate of a point on the Montgomery form of
72/// Curve25519 or its twist.
73#[derive(Copy, Clone, Debug, Default)]
74#[cfg_attr(feature = "serde", derive(serde::Serialize, serde::Deserialize))]
75pub struct MontgomeryPoint(pub [u8; 32]);
76
77/// Equality of `MontgomeryPoint`s is defined mod p.
78impl ConstantTimeEq for MontgomeryPoint {
79    fn ct_eq(&self, other: &MontgomeryPoint) -> Choice {
80        let self_fe = FieldElement::from_bytes(&self.0);
81        let other_fe = FieldElement::from_bytes(&other.0);
82
83        self_fe.ct_eq(&other_fe)
84    }
85}
86
87impl ConditionallySelectable for MontgomeryPoint {
88    fn conditional_select(a: &Self, b: &Self, choice: Choice) -> Self {
89        Self(<[u8; 32]>::conditional_select(&a.0, &b.0, choice))
90    }
91}
92
93impl PartialEq for MontgomeryPoint {
94    fn eq(&self, other: &MontgomeryPoint) -> bool {
95        self.ct_eq(other).into()
96    }
97}
98
99impl Eq for MontgomeryPoint {}
100
101// Equal MontgomeryPoints must hash to the same value. So we have to get them into a canonical
102// encoding first
103impl Hash for MontgomeryPoint {
104    fn hash<H: Hasher>(&self, state: &mut H) {
105        // Do a round trip through a `FieldElement`. `as_bytes` is guaranteed to give a canonical
106        // 32-byte encoding
107        let canonical_bytes = FieldElement::from_bytes(&self.0).to_bytes();
108        canonical_bytes.hash(state);
109    }
110}
111
112impl Identity for MontgomeryPoint {
113    /// Return the group identity element, which has order 4.
114    fn identity() -> MontgomeryPoint {
115        MontgomeryPoint([0u8; 32])
116    }
117}
118
119#[cfg(feature = "zeroize")]
120impl Zeroize for MontgomeryPoint {
121    fn zeroize(&mut self) {
122        self.0.zeroize();
123    }
124}
125
126impl MontgomeryPoint {
127    /// Fixed-base scalar multiplication (i.e. multiplication by the base point).
128    pub fn mul_base(scalar: &Scalar) -> Self {
129        EdwardsPoint::mul_base(scalar).to_montgomery()
130    }
131
132    /// Multiply this point by `clamp_integer(bytes)`. For a description of clamping, see
133    /// [`clamp_integer`].
134    pub fn mul_clamped(self, bytes: [u8; 32]) -> Self {
135        // We have to construct a Scalar that is not reduced mod l, which breaks scalar invariant
136        // #2. But #2 is not necessary for correctness of variable-base multiplication. All that
137        // needs to hold is invariant #1, i.e., the scalar is less than 2^255. This is guaranteed
138        // by clamping.
139        // Further, we don't do any reduction or arithmetic with this clamped value, so there's no
140        // issues arising from the fact that the curve point is not necessarily in the prime-order
141        // subgroup.
142        let s = Scalar {
143            bytes: clamp_integer(bytes),
144        };
145        s * self
146    }
147
148    /// Multiply the basepoint by `clamp_integer(bytes)`. For a description of clamping, see
149    /// [`clamp_integer`].
150    pub fn mul_base_clamped(bytes: [u8; 32]) -> Self {
151        // See reasoning in Self::mul_clamped why it is OK to make an unreduced Scalar here. We
152        // note that fixed-base multiplication is also defined for all values of `bytes` less than
153        // 2^255.
154        let s = Scalar {
155            bytes: clamp_integer(bytes),
156        };
157        Self::mul_base(&s)
158    }
159
160    /// Given `self` \\( = u\_0(P) \\), and a big-endian bit representation of an integer
161    /// \\(n\\), return \\( u\_0(\[n\]P) \\). This is constant time in the length of `bits`.
162    ///
163    /// **NOTE:** You probably do not want to use this function. Almost every protocol built on
164    /// Curve25519 uses _clamped multiplication_, explained
165    /// [here](https://neilmadden.blog/2020/05/28/whats-the-curve25519-clamping-all-about/).
166    /// When in doubt, use [`Self::mul_clamped`].
167    pub fn mul_bits_be(&self, bits: impl Iterator<Item = bool>) -> MontgomeryPoint {
168        // Algorithm 8 of Costello-Smith 2017
169        let affine_u = FieldElement::from_bytes(&self.0);
170        let mut x0 = ProjectivePoint::identity();
171        let mut x1 = ProjectivePoint {
172            U: affine_u,
173            W: FieldElement::ONE,
174        };
175
176        // Go through the bits from most to least significant, using a sliding window of 2
177        let mut prev_bit = false;
178        for cur_bit in bits {
179            let choice: u8 = (prev_bit ^ cur_bit) as u8;
180
181            debug_assert!(choice == 0 || choice == 1);
182
183            ProjectivePoint::conditional_swap(&mut x0, &mut x1, choice.into());
184            differential_add_and_double(&mut x0, &mut x1, &affine_u);
185
186            prev_bit = cur_bit;
187        }
188        // The final value of prev_bit above is scalar.bits()[0], i.e., the LSB of scalar
189        ProjectivePoint::conditional_swap(&mut x0, &mut x1, Choice::from(prev_bit as u8));
190        // Don't leave the bit in the stack
191        #[cfg(feature = "zeroize")]
192        prev_bit.zeroize();
193
194        x0.as_affine()
195    }
196
197    /// View this `MontgomeryPoint` as an array of bytes.
198    pub const fn as_bytes(&self) -> &[u8; 32] {
199        &self.0
200    }
201
202    /// Convert this `MontgomeryPoint` to an array of bytes.
203    pub const fn to_bytes(&self) -> [u8; 32] {
204        self.0
205    }
206
207    /// Attempt to convert to an `EdwardsPoint`, using the supplied
208    /// choice of sign for the `EdwardsPoint`.
209    ///
210    /// # Inputs
211    ///
212    /// * `sign`: a `u8` donating the desired sign of the resulting
213    ///   `EdwardsPoint`.  `0` denotes positive and `1` negative.
214    ///
215    /// # Return
216    ///
217    /// * `Some(EdwardsPoint)` if `self` is the \\(u\\)-coordinate of a
218    ///   point on (the Montgomery form of) Curve25519;
219    ///
220    /// * `None` if `self` is the \\(u\\)-coordinate of a point on the
221    ///   twist of (the Montgomery form of) Curve25519;
222    ///
223    pub fn to_edwards(&self, sign: u8) -> Option<EdwardsPoint> {
224        // To decompress the Montgomery u coordinate to an
225        // `EdwardsPoint`, we apply the birational map to obtain the
226        // Edwards y coordinate, then do Edwards decompression.
227        //
228        // The birational map is y = (u-1)/(u+1).
229        //
230        // The exceptional points are the zeros of the denominator,
231        // i.e., u = -1.
232        //
233        // But when u = -1, v^2 = u*(u^2+486662*u+1) = 486660.
234        //
235        // Since this is nonsquare mod p, u = -1 corresponds to a point
236        // on the twist, not the curve, so we can reject it early.
237
238        let u = FieldElement::from_bytes(&self.0);
239
240        if u == FieldElement::MINUS_ONE {
241            return None;
242        }
243
244        let one = FieldElement::ONE;
245
246        let y = &(&u - &one) * &(&u + &one).invert();
247
248        let mut y_bytes = y.to_bytes();
249        y_bytes[31] ^= sign << 7;
250
251        CompressedEdwardsY(y_bytes).decompress()
252    }
253}
254
255/// Perform the Elligator2 mapping to a Montgomery point. Returns a Montgomery point and a `Choice`
256/// determining whether eps is a square. This is required by the standard to determine the
257/// sign of the v coordinate.
258///
259/// See <https://www.rfc-editor.org/rfc/rfc9380.html#name-elligator-2-method>
260//
261#[allow(unused)]
262pub(crate) fn elligator_encode(r_0: &FieldElement) -> (MontgomeryPoint, Choice) {
263    let one = FieldElement::ONE;
264    let d_1 = &one + &r_0.square2(); /* 2r^2 */
265
266    let d = &MONTGOMERY_A_NEG * &(d_1.invert()); /* A/(1+2r^2) */
267
268    let d_sq = &d.square();
269    let au = &MONTGOMERY_A * &d;
270
271    let inner = &(d_sq + &au) + &one;
272    let eps = &d * &inner; /* eps = d^3 + Ad^2 + d */
273
274    let (eps_is_sq, _eps) = FieldElement::sqrt_ratio_i(&eps, &one);
275
276    let zero = FieldElement::ZERO;
277    let Atemp = FieldElement::conditional_select(&MONTGOMERY_A, &zero, eps_is_sq); /* 0, or A if nonsquare*/
278    let mut u = &d + &Atemp; /* d, or d+A if nonsquare */
279    u.conditional_negate(!eps_is_sq); /* d, or -d-A if nonsquare */
280
281    (MontgomeryPoint(u.to_bytes()), eps_is_sq)
282}
283
284/// A `ProjectivePoint` holds a point on the projective line
285/// \\( \mathbb P(\mathbb F\_p) \\), which we identify with the Kummer
286/// line of the Montgomery curve.
287#[derive(Copy, Clone, Debug)]
288struct ProjectivePoint {
289    pub U: FieldElement,
290    pub W: FieldElement,
291}
292
293impl Identity for ProjectivePoint {
294    fn identity() -> ProjectivePoint {
295        ProjectivePoint {
296            U: FieldElement::ONE,
297            W: FieldElement::ZERO,
298        }
299    }
300}
301
302impl Default for ProjectivePoint {
303    fn default() -> ProjectivePoint {
304        ProjectivePoint::identity()
305    }
306}
307
308impl ConditionallySelectable for ProjectivePoint {
309    fn conditional_select(
310        a: &ProjectivePoint,
311        b: &ProjectivePoint,
312        choice: Choice,
313    ) -> ProjectivePoint {
314        ProjectivePoint {
315            U: FieldElement::conditional_select(&a.U, &b.U, choice),
316            W: FieldElement::conditional_select(&a.W, &b.W, choice),
317        }
318    }
319}
320
321impl ProjectivePoint {
322    /// Dehomogenize this point to affine coordinates.
323    ///
324    /// # Return
325    ///
326    /// * \\( u = U / W \\) if \\( W \neq 0 \\);
327    /// * \\( 0 \\) if \\( W \eq 0 \\);
328    pub fn as_affine(&self) -> MontgomeryPoint {
329        let u = &self.U * &self.W.invert();
330        MontgomeryPoint(u.to_bytes())
331    }
332}
333
334/// Perform the double-and-add step of the Montgomery ladder.
335///
336/// Given projective points
337/// \\( (U\_P : W\_P) = u(P) \\),
338/// \\( (U\_Q : W\_Q) = u(Q) \\),
339/// and the affine difference
340/// \\(      u\_{P-Q} = u(P-Q) \\), set
341/// $$
342///     (U\_P : W\_P) \gets u(\[2\]P)
343/// $$
344/// and
345/// $$
346///     (U\_Q : W\_Q) \gets u(P + Q).
347/// $$
348#[rustfmt::skip] // keep alignment of explanatory comments
349fn differential_add_and_double(
350    P: &mut ProjectivePoint,
351    Q: &mut ProjectivePoint,
352    affine_PmQ: &FieldElement,
353) {
354    let t0 = &P.U + &P.W;
355    let t1 = &P.U - &P.W;
356    let t2 = &Q.U + &Q.W;
357    let t3 = &Q.U - &Q.W;
358
359    let t4 = t0.square();   // (U_P + W_P)^2 = U_P^2 + 2 U_P W_P + W_P^2
360    let t5 = t1.square();   // (U_P - W_P)^2 = U_P^2 - 2 U_P W_P + W_P^2
361
362    let t6 = &t4 - &t5;     // 4 U_P W_P
363
364    let t7 = &t0 * &t3;     // (U_P + W_P) (U_Q - W_Q) = U_P U_Q + W_P U_Q - U_P W_Q - W_P W_Q
365    let t8 = &t1 * &t2;     // (U_P - W_P) (U_Q + W_Q) = U_P U_Q - W_P U_Q + U_P W_Q - W_P W_Q
366
367    let t9  = &t7 + &t8;    // 2 (U_P U_Q - W_P W_Q)
368    let t10 = &t7 - &t8;    // 2 (W_P U_Q - U_P W_Q)
369
370    let t11 =  t9.square(); // 4 (U_P U_Q - W_P W_Q)^2
371    let t12 = t10.square(); // 4 (W_P U_Q - U_P W_Q)^2
372
373    let t13 = &APLUS2_OVER_FOUR * &t6; // (A + 2) U_P U_Q
374
375    let t14 = &t4 * &t5;    // ((U_P + W_P)(U_P - W_P))^2 = (U_P^2 - W_P^2)^2
376    let t15 = &t13 + &t5;   // (U_P - W_P)^2 + (A + 2) U_P W_P
377
378    let t16 = &t6 * &t15;   // 4 (U_P W_P) ((U_P - W_P)^2 + (A + 2) U_P W_P)
379
380    let t17 = affine_PmQ * &t12; // U_D * 4 (W_P U_Q - U_P W_Q)^2
381    let t18 = t11;               // W_D * 4 (U_P U_Q - W_P W_Q)^2
382
383    P.U = t14;  // U_{P'} = (U_P + W_P)^2 (U_P - W_P)^2
384    P.W = t16;  // W_{P'} = (4 U_P W_P) ((U_P - W_P)^2 + ((A + 2)/4) 4 U_P W_P)
385    Q.U = t18;  // U_{Q'} = W_D * 4 (U_P U_Q - W_P W_Q)^2
386    Q.W = t17;  // W_{Q'} = U_D * 4 (W_P U_Q - U_P W_Q)^2
387}
388
389define_mul_assign_variants!(LHS = MontgomeryPoint, RHS = Scalar);
390
391define_mul_variants!(
392    LHS = MontgomeryPoint,
393    RHS = Scalar,
394    Output = MontgomeryPoint
395);
396define_mul_variants!(
397    LHS = Scalar,
398    RHS = MontgomeryPoint,
399    Output = MontgomeryPoint
400);
401
402/// Multiply this `MontgomeryPoint` by a `Scalar`.
403impl Mul<&Scalar> for &MontgomeryPoint {
404    type Output = MontgomeryPoint;
405
406    /// Given `self` \\( = u\_0(P) \\), and a `Scalar` \\(n\\), return \\( u\_0(\[n\]P) \\)
407    fn mul(self, scalar: &Scalar) -> MontgomeryPoint {
408        // We multiply by the integer representation of the given Scalar. By scalar invariant #1,
409        // the MSB is 0, so we can skip it.
410        self.mul_bits_be(scalar.bits_le().rev().skip(1))
411    }
412}
413
414impl MulAssign<&Scalar> for MontgomeryPoint {
415    fn mul_assign(&mut self, scalar: &Scalar) {
416        *self = (self as &MontgomeryPoint) * scalar;
417    }
418}
419
420impl Mul<&MontgomeryPoint> for &Scalar {
421    type Output = MontgomeryPoint;
422
423    fn mul(self, point: &MontgomeryPoint) -> MontgomeryPoint {
424        point * self
425    }
426}
427
428// ------------------------------------------------------------------------
429// Tests
430// ------------------------------------------------------------------------
431
432#[cfg(test)]
433mod test {
434    use super::*;
435    use crate::constants;
436
437    #[cfg(feature = "alloc")]
438    use alloc::vec::Vec;
439
440    use rand_core::{CryptoRng, RngCore};
441
442    #[test]
443    fn identity_in_different_coordinates() {
444        let id_projective = ProjectivePoint::identity();
445        let id_montgomery = id_projective.as_affine();
446
447        assert!(id_montgomery == MontgomeryPoint::identity());
448    }
449
450    #[test]
451    fn identity_in_different_models() {
452        assert!(EdwardsPoint::identity().to_montgomery() == MontgomeryPoint::identity());
453    }
454
455    #[test]
456    #[cfg(feature = "serde")]
457    fn serde_bincode_basepoint_roundtrip() {
458        use bincode;
459
460        let encoded = bincode::serialize(&constants::X25519_BASEPOINT).unwrap();
461        let decoded: MontgomeryPoint = bincode::deserialize(&encoded).unwrap();
462
463        assert_eq!(encoded.len(), 32);
464        assert_eq!(decoded, constants::X25519_BASEPOINT);
465
466        let raw_bytes = constants::X25519_BASEPOINT.as_bytes();
467        let bp: MontgomeryPoint = bincode::deserialize(raw_bytes).unwrap();
468        assert_eq!(bp, constants::X25519_BASEPOINT);
469    }
470
471    /// Test Montgomery -> Edwards on the X/Ed25519 basepoint
472    #[test]
473    fn basepoint_montgomery_to_edwards() {
474        // sign bit = 0 => basepoint
475        assert_eq!(
476            constants::ED25519_BASEPOINT_POINT,
477            constants::X25519_BASEPOINT.to_edwards(0).unwrap()
478        );
479        // sign bit = 1 => minus basepoint
480        assert_eq!(
481            -constants::ED25519_BASEPOINT_POINT,
482            constants::X25519_BASEPOINT.to_edwards(1).unwrap()
483        );
484    }
485
486    /// Test Edwards -> Montgomery on the X/Ed25519 basepoint
487    #[test]
488    fn basepoint_edwards_to_montgomery() {
489        assert_eq!(
490            constants::ED25519_BASEPOINT_POINT.to_montgomery(),
491            constants::X25519_BASEPOINT
492        );
493    }
494
495    /// Check that Montgomery -> Edwards fails for points on the twist.
496    #[test]
497    fn montgomery_to_edwards_rejects_twist() {
498        let one = FieldElement::ONE;
499
500        // u = 2 corresponds to a point on the twist.
501        let two = MontgomeryPoint((&one + &one).to_bytes());
502
503        assert!(two.to_edwards(0).is_none());
504
505        // u = -1 corresponds to a point on the twist, but should be
506        // checked explicitly because it's an exceptional point for the
507        // birational map.  For instance, libsignal will accept it.
508        let minus_one = MontgomeryPoint((-&one).to_bytes());
509
510        assert!(minus_one.to_edwards(0).is_none());
511    }
512
513    #[test]
514    fn eq_defined_mod_p() {
515        let mut u18_bytes = [0u8; 32];
516        u18_bytes[0] = 18;
517        let u18 = MontgomeryPoint(u18_bytes);
518        let u18_unred = MontgomeryPoint([255; 32]);
519
520        assert_eq!(u18, u18_unred);
521    }
522
523    /// Returns a random point on the prime-order subgroup
524    fn rand_prime_order_point(mut rng: impl RngCore + CryptoRng) -> EdwardsPoint {
525        let s: Scalar = Scalar::random(&mut rng);
526        EdwardsPoint::mul_base(&s)
527    }
528
529    /// Given a bytestring that's little-endian at the byte level, return an iterator over all the
530    /// bits, in little-endian order.
531    fn bytestring_bits_le(x: &[u8]) -> impl DoubleEndedIterator<Item = bool> + Clone + '_ {
532        let bitlen = x.len() * 8;
533        (0..bitlen).map(|i| {
534            // As i runs from 0..256, the bottom 3 bits index the bit, while the upper bits index
535            // the byte. Since self.bytes is little-endian at the byte level, this iterator is
536            // little-endian on the bit level
537            ((x[i >> 3] >> (i & 7)) & 1u8) == 1
538        })
539    }
540
541    #[test]
542    fn montgomery_ladder_matches_edwards_scalarmult() {
543        let mut csprng = rand_core::OsRng;
544
545        for _ in 0..100 {
546            let p_edwards = rand_prime_order_point(csprng);
547            let p_montgomery: MontgomeryPoint = p_edwards.to_montgomery();
548
549            let s: Scalar = Scalar::random(&mut csprng);
550            let expected = s * p_edwards;
551            let result = s * p_montgomery;
552
553            assert_eq!(result, expected.to_montgomery())
554        }
555    }
556
557    // Tests that, on the prime-order subgroup, MontgomeryPoint::mul_bits_be is the same as
558    // multiplying by the Scalar representation of the same bits
559    #[test]
560    fn montgomery_mul_bits_be() {
561        let mut csprng = rand_core::OsRng;
562
563        for _ in 0..100 {
564            // Make a random prime-order point P
565            let p_edwards = rand_prime_order_point(csprng);
566            let p_montgomery: MontgomeryPoint = p_edwards.to_montgomery();
567
568            // Make a random integer b
569            let mut bigint = [0u8; 64];
570            csprng.fill_bytes(&mut bigint[..]);
571            let bigint_bits_be = bytestring_bits_le(&bigint).rev();
572
573            // Check that bP is the same whether calculated as scalar-times-edwards or
574            // integer-times-montgomery.
575            let expected = Scalar::from_bytes_mod_order_wide(&bigint) * p_edwards;
576            let result = p_montgomery.mul_bits_be(bigint_bits_be);
577            assert_eq!(result, expected.to_montgomery())
578        }
579    }
580
581    // Tests that MontgomeryPoint::mul_bits_be is consistent on any point, even ones that might be
582    // on the curve's twist. Specifically, this tests that b₁(b₂P) == b₂(b₁P) for random
583    // integers b₁, b₂ and random (curve or twist) point P.
584    #[test]
585    fn montgomery_mul_bits_be_twist() {
586        let mut csprng = rand_core::OsRng;
587
588        for _ in 0..100 {
589            // Make a random point P on the curve or its twist
590            let p_montgomery = {
591                let mut buf = [0u8; 32];
592                csprng.fill_bytes(&mut buf);
593                MontgomeryPoint(buf)
594            };
595
596            // Compute two big integers b₁ and b₂
597            let mut bigint1 = [0u8; 64];
598            let mut bigint2 = [0u8; 64];
599            csprng.fill_bytes(&mut bigint1[..]);
600            csprng.fill_bytes(&mut bigint2[..]);
601
602            // Compute b₁P and b₂P
603            let bigint1_bits_be = bytestring_bits_le(&bigint1).rev();
604            let bigint2_bits_be = bytestring_bits_le(&bigint2).rev();
605            let prod1 = p_montgomery.mul_bits_be(bigint1_bits_be.clone());
606            let prod2 = p_montgomery.mul_bits_be(bigint2_bits_be.clone());
607
608            // Check that b₁(b₂P) == b₂(b₁P)
609            assert_eq!(
610                prod1.mul_bits_be(bigint2_bits_be),
611                prod2.mul_bits_be(bigint1_bits_be)
612            );
613        }
614    }
615
616    /// Check that mul_base_clamped and mul_clamped agree
617    #[test]
618    fn mul_base_clamped() {
619        let mut csprng = rand_core::OsRng;
620
621        // Test agreement on a large integer. Even after clamping, this is not reduced mod l.
622        let a_bytes = [0xff; 32];
623        assert_eq!(
624            MontgomeryPoint::mul_base_clamped(a_bytes),
625            constants::X25519_BASEPOINT.mul_clamped(a_bytes)
626        );
627
628        // Test agreement on random integers
629        for _ in 0..100 {
630            // This will be reduced mod l with probability l / 2^256 ≈ 6.25%
631            let mut a_bytes = [0u8; 32];
632            csprng.fill_bytes(&mut a_bytes);
633
634            assert_eq!(
635                MontgomeryPoint::mul_base_clamped(a_bytes),
636                constants::X25519_BASEPOINT.mul_clamped(a_bytes)
637            );
638        }
639    }
640
641    #[cfg(feature = "alloc")]
642    const ELLIGATOR_CORRECT_OUTPUT: [u8; 32] = [
643        0x5f, 0x35, 0x20, 0x00, 0x1c, 0x6c, 0x99, 0x36, 0xa3, 0x12, 0x06, 0xaf, 0xe7, 0xc7, 0xac,
644        0x22, 0x4e, 0x88, 0x61, 0x61, 0x9b, 0xf9, 0x88, 0x72, 0x44, 0x49, 0x15, 0x89, 0x9d, 0x95,
645        0xf4, 0x6e,
646    ];
647
648    #[test]
649    #[cfg(feature = "alloc")]
650    fn montgomery_elligator_correct() {
651        let bytes: Vec<u8> = (0u8..32u8).collect();
652        let bits_in: [u8; 32] = (&bytes[..]).try_into().expect("Range invariant broken");
653
654        let fe = FieldElement::from_bytes(&bits_in);
655        let (eg, _) = elligator_encode(&fe);
656        assert_eq!(eg.to_bytes(), ELLIGATOR_CORRECT_OUTPUT);
657    }
658
659    #[test]
660    fn montgomery_elligator_zero_zero() {
661        let zero = [0u8; 32];
662        let fe = FieldElement::from_bytes(&zero);
663        let (eg, _) = elligator_encode(&fe);
664        assert_eq!(eg.to_bytes(), zero);
665    }
666}