ark_ec/hashing/curve_maps/
wb.rs

1use core::marker::PhantomData;
2
3use crate::{models::short_weierstrass::SWCurveConfig, CurveConfig};
4use ark_ff::batch_inversion;
5use ark_poly::{univariate::DensePolynomial, DenseUVPolynomial, Polynomial};
6
7use crate::{
8    hashing::{map_to_curve_hasher::MapToCurve, HashToCurveError},
9    models::short_weierstrass::{Affine, Projective},
10    AffineRepr,
11};
12
13use super::swu::{SWUConfig, SWUMap};
14type BaseField<MP> = <MP as CurveConfig>::BaseField;
15
16/// [`IsogenyMap`] defines an isogeny between curves of
17/// form `Phi(x, y) := (a(x), b(x)*y).
18/// The `x` coordinate of the codomain point only depends on the
19/// `x`-coordinate of the domain point, and the
20/// `y`-coordinate of the codomain point is a multiple of the `y`-coordinate of the domain point.
21/// The multiplier depends on the `x`-coordinate of the domain point.
22/// All isogeny maps of curves of short Weierstrass form can be written in this way. See
23/// [\[Ga18]\]. Theorem 9.7.5 for details.
24///
25/// We assume that `Domain` and `Codomain` have the same `BaseField` but we use both
26/// `BaseField<Domain>` and `BaseField<Codomain>` in our fields' definitions to avoid
27/// using `PhantomData`
28///
29/// - [\[Ga18]\] Galbraith, S. D. (2018). Mathematics of public key cryptography.
30pub struct IsogenyMap<
31    'a,
32    Domain: SWCurveConfig,
33    Codomain: SWCurveConfig<BaseField = BaseField<Domain>>,
34> {
35    pub x_map_numerator: &'a [BaseField<Domain>],
36    pub x_map_denominator: &'a [BaseField<Codomain>],
37
38    pub y_map_numerator: &'a [BaseField<Domain>],
39    pub y_map_denominator: &'a [BaseField<Codomain>],
40}
41
42impl<'a, Domain, Codomain> IsogenyMap<'a, Domain, Codomain>
43where
44    Domain: SWCurveConfig,
45    Codomain: SWCurveConfig<BaseField = BaseField<Domain>>,
46{
47    fn apply(&self, domain_point: Affine<Domain>) -> Result<Affine<Codomain>, HashToCurveError> {
48        match domain_point.xy() {
49            Some((x, y)) => {
50                let x_num = DensePolynomial::from_coefficients_slice(self.x_map_numerator);
51                let x_den = DensePolynomial::from_coefficients_slice(self.x_map_denominator);
52
53                let y_num = DensePolynomial::from_coefficients_slice(self.y_map_numerator);
54                let y_den = DensePolynomial::from_coefficients_slice(self.y_map_denominator);
55
56                let mut v: [BaseField<Domain>; 2] = [x_den.evaluate(&x), y_den.evaluate(&x)];
57                batch_inversion(&mut v);
58                let img_x = x_num.evaluate(&x) * v[0];
59                let img_y = (y_num.evaluate(&x) * y) * v[1];
60                Ok(Affine::<Codomain>::new_unchecked(img_x, img_y))
61            },
62            None => Ok(Affine::identity()),
63        }
64    }
65}
66
67/// Trait defining the necessary parameters for the WB hash-to-curve method
68/// for the curves of Weierstrass form of:
69/// of y^2 = x^3 + a*x + b where b != 0 but `a` can be zero like BLS-381 curve.
70/// From [\[WB2019\]]
71///
72/// - [\[WB2019\]] <http://dx.doi.org/10.46586/tches.v2019.i4.154-179>
73pub trait WBConfig: SWCurveConfig + Sized {
74    // The isogenous curve should be defined over the same base field but it can have
75    // different scalar field type IsogenousCurveScalarField :
76    type IsogenousCurve: SWUConfig<BaseField = BaseField<Self>>;
77
78    const ISOGENY_MAP: IsogenyMap<'static, Self::IsogenousCurve, Self>;
79}
80
81pub struct WBMap<P: WBConfig> {
82    swu_field_curve_hasher: PhantomData<SWUMap<P::IsogenousCurve>>,
83    curve_params: PhantomData<fn() -> P>,
84}
85
86impl<P: WBConfig> MapToCurve<Projective<P>> for WBMap<P> {
87    /// Checks if `P` represents a valid map.
88    fn check_parameters() -> Result<(), HashToCurveError> {
89        match P::ISOGENY_MAP.apply(P::IsogenousCurve::GENERATOR) {
90            Ok(point_on_curve) => {
91                debug_assert!(point_on_curve.is_on_curve(),
92			      "the isogeny maps the generator of its domain: {} into {} which does not belong to its codomain.",P::IsogenousCurve::GENERATOR, point_on_curve);
93            },
94            Err(e) => return Err(e),
95        }
96
97        SWUMap::<P::IsogenousCurve>::check_parameters().unwrap(); // Or ?
98        Ok(())
99    }
100
101    /// Map random field point to a random curve point
102    /// inspired from
103    /// <https://github.com/zcash/pasta_curves/blob/main/src/hashtocurve.rs>
104    fn map_to_curve(
105        element: <Affine<P> as AffineRepr>::BaseField,
106    ) -> Result<Affine<P>, HashToCurveError> {
107        // first we need to map the field point to the isogenous curve
108        let point_on_isogenious_curve = SWUMap::<P::IsogenousCurve>::map_to_curve(element).unwrap();
109        P::ISOGENY_MAP.apply(point_on_isogenious_curve)
110    }
111}
112
113#[cfg(test)]
114mod test {
115    use crate::{
116        hashing::{
117            curve_maps::{
118                swu::SWUConfig,
119                wb::{IsogenyMap, WBConfig, WBMap},
120            },
121            map_to_curve_hasher::MapToCurveBasedHasher,
122            HashToCurve,
123        },
124        models::short_weierstrass::SWCurveConfig,
125        short_weierstrass::{Affine, Projective},
126        CurveConfig,
127    };
128    use ark_ff::{field_hashers::DefaultFieldHasher, fields::Fp64, MontBackend, MontFp};
129
130    #[derive(ark_ff::MontConfig)]
131    #[modulus = "127"]
132    #[generator = "6"]
133    pub struct F127Config;
134    pub type F127 = Fp64<MontBackend<F127Config, 1>>;
135
136    const F127_ZERO: F127 = MontFp!("0");
137    const F127_ONE: F127 = MontFp!("1");
138
139    /// The struct defining our parameters for the target curve of hashing
140    struct TestWBF127MapToCurveConfig;
141
142    impl CurveConfig for TestWBF127MapToCurveConfig {
143        const COFACTOR: &'static [u64] = &[1];
144
145    #[rustfmt::skip]
146        const COFACTOR_INV: F127 = F127_ONE;
147
148        type BaseField = F127;
149        type ScalarField = F127;
150    }
151
152    /// E: Elliptic Curve defined by y^2 = x^3 + 3 over Finite
153    /// Field of size 127
154    impl SWCurveConfig for TestWBF127MapToCurveConfig {
155        /// COEFF_A = 0
156        const COEFF_A: F127 = F127_ZERO;
157
158        /// COEFF_B = 3
159    #[rustfmt::skip]
160        const COEFF_B: F127 = MontFp!("3");
161
162        /// AFFINE_GENERATOR_COEFFS = (G1_GENERATOR_X, G1_GENERATOR_Y)
163        const GENERATOR: Affine<Self> = Affine::new_unchecked(MontFp!("62"), MontFp!("70"));
164    }
165
166    /// Testing WB19 hashing on a small curve
167    /// E_isogenous : Elliptic Curve defined by y^2 = x^3 + 109*x + 124 over Finite
168    /// Field of size 127
169    /// Isogenous to E : y^2 = x^3 + 3
170    struct TestSWU127MapToIsogenousCurveConfig;
171
172    /// First we define the isogenous curve
173    /// sage: E_isogenous.order()
174    /// 127
175    impl CurveConfig for TestSWU127MapToIsogenousCurveConfig {
176        const COFACTOR: &'static [u64] = &[1];
177
178    #[rustfmt::skip]
179        const COFACTOR_INV: F127 = F127_ONE;
180
181        type BaseField = F127;
182        type ScalarField = F127;
183    }
184
185    /// E_isogenous : Elliptic Curve defined by y^2 = x^3 + 109*x + 124 over Finite
186    /// Field of size 127
187    impl SWCurveConfig for TestSWU127MapToIsogenousCurveConfig {
188        /// COEFF_A = 109
189        const COEFF_A: F127 = MontFp!("109");
190
191        /// COEFF_B = 124
192    #[rustfmt::skip]
193        const COEFF_B: F127 = MontFp!("124");
194
195        /// AFFINE_GENERATOR_COEFFS = (G1_GENERATOR_X, G1_GENERATOR_Y)
196        const GENERATOR: Affine<Self> = Affine::new_unchecked(MontFp!("84"), MontFp!("2"));
197    }
198
199    /// SWU parameters for E_isogenous
200    impl SWUConfig for TestSWU127MapToIsogenousCurveConfig {
201        /// NON-SQUARE = - 1
202        const ZETA: F127 = MontFp!("-1");
203    }
204
205    /// E_isogenous : Elliptic Curve defined by y^2 = x^3 + 109*x + 124 over Finite
206    /// Field of size 127
207    /// With psi: E_isogenous -> E
208    /// psi = (psi_x(x,y), psi_y(x,y))
209    /// where
210    /// psi_x: (-57*x^13 - 21*x^12 + 10*x^11 + 34*x^10 + 40*x^9 -
211    /// 13*x^8 + 32*x^7 - 32*x^6 + 23*x^5 - 14*x^4 + 39*x^3 + 23*x^2 + 63*x +
212    /// 4)/(x^12 - 13*x^11 + 11*x^10 - 33*x^9 - 30*x^8 + 30*x^7 + 34*x^6 - 44*x^5 +
213    /// 63*x^4 - 20*x^3 - 10*x^2 + 31*x + 2)
214    ///
215    /// psi_y: (10*x^18*y + 59*x^17*y + 41*x^16*y + 48*x^15*y - 7*x^14*y + 6*x^13*y +
216    /// 5*x^12*y + 62*x^11*y + 12*x^10*y + 36*x^9*y - 49*x^8*y - 18*x^7*y - 63*x^6*y
217    /// - 43*x^5*y - 60*x^4*y - 18*x^3*y + 30*x^2*y - 57*x*y - 34*y)/(x^18 + 44*x^17
218    /// - 63*x^16 + 52*x^15 + 3*x^14 + 38*x^13 - 30*x^12 + 11*x^11 - 42*x^10 - 13*x^9
219    /// - 46*x^8 - 61*x^7 - 16*x^6 - 55*x^5 + 18*x^4 + 23*x^3 - 24*x^2 - 18*x + 32)
220    const ISOGENY_MAP_TESTWBF127: IsogenyMap<
221        '_,
222        TestSWU127MapToIsogenousCurveConfig,
223        TestWBF127MapToCurveConfig,
224    > = IsogenyMap {
225        x_map_numerator: &[
226            MontFp!("4"),
227            MontFp!("63"),
228            MontFp!("23"),
229            MontFp!("39"),
230            MontFp!("-14"),
231            MontFp!("23"),
232            MontFp!("-32"),
233            MontFp!("32"),
234            MontFp!("-13"),
235            MontFp!("40"),
236            MontFp!("34"),
237            MontFp!("10"),
238            MontFp!("-21"),
239            MontFp!("-57"),
240        ],
241
242        x_map_denominator: &[
243            MontFp!("2"),
244            MontFp!("31"),
245            MontFp!("-10"),
246            MontFp!("-20"),
247            MontFp!("63"),
248            MontFp!("-44"),
249            MontFp!("34"),
250            MontFp!("30"),
251            MontFp!("-30"),
252            MontFp!("-33"),
253            MontFp!("11"),
254            MontFp!("-13"),
255            MontFp!("1"),
256        ],
257
258        y_map_numerator: &[
259            MontFp!("-34"),
260            MontFp!("-57"),
261            MontFp!("30"),
262            MontFp!("-18"),
263            MontFp!("-60"),
264            MontFp!("-43"),
265            MontFp!("-63"),
266            MontFp!("-18"),
267            MontFp!("-49"),
268            MontFp!("36"),
269            MontFp!("12"),
270            MontFp!("62"),
271            MontFp!("5"),
272            MontFp!("6"),
273            MontFp!("-7"),
274            MontFp!("48"),
275            MontFp!("41"),
276            MontFp!("59"),
277            MontFp!("10"),
278        ],
279
280        y_map_denominator: &[
281            MontFp!("32"),
282            MontFp!("-18"),
283            MontFp!("-24"),
284            MontFp!("23"),
285            MontFp!("18"),
286            MontFp!("-55"),
287            MontFp!("-16"),
288            MontFp!("-61"),
289            MontFp!("-46"),
290            MontFp!("-13"),
291            MontFp!("-42"),
292            MontFp!("11"),
293            MontFp!("-30"),
294            MontFp!("38"),
295            MontFp!("3"),
296            MontFp!("52"),
297            MontFp!("-63"),
298            MontFp!("44"),
299            MontFp!("1"),
300        ],
301    };
302    impl WBConfig for TestWBF127MapToCurveConfig {
303        type IsogenousCurve = TestSWU127MapToIsogenousCurveConfig;
304
305        const ISOGENY_MAP: super::IsogenyMap<'static, Self::IsogenousCurve, Self> =
306            ISOGENY_MAP_TESTWBF127;
307    }
308
309    /// The point of the test is to get a simple WB compatible curve
310    /// and make simple hash
311    #[test]
312    fn hash_arbitrary_string_to_curve_wb() {
313        use sha2::Sha256;
314        let test_wb_to_curve_hasher = MapToCurveBasedHasher::<
315            Projective<TestWBF127MapToCurveConfig>,
316            DefaultFieldHasher<Sha256, 128>,
317            WBMap<TestWBF127MapToCurveConfig>,
318        >::new(&[1])
319        .unwrap();
320
321        let hash_result = test_wb_to_curve_hasher.hash(b"if you stick a Babel fish in your ear you can instantly understand anything said to you in any form of language.").expect("fail to hash the string to curve");
322
323        assert!(
324            hash_result.x != F127_ZERO && hash_result.y != F127_ZERO,
325            "we assume that not both a and b coefficienst are zero for the test curve"
326        );
327
328        assert!(
329            hash_result.is_on_curve(),
330            "hash results into a point off the curve"
331        );
332    }
333}