Implement P256 verification via RIP-7212 precompile with Solidity fallback (#4881)

Co-authored-by: Ernesto García <ernestognw@gmail.com>
Co-authored-by: cairo <cairoeth@protonmail.com>
Co-authored-by: sudo rm -rf --no-preserve-root / <pcaversaccio@users.noreply.github.com>
This commit is contained in:
Hadrien Croubois
2024-07-03 09:17:46 +02:00
committed by GitHub
parent ccc110360f
commit 05f218fb66
11 changed files with 4400 additions and 5 deletions

View File

@ -0,0 +1,5 @@
---
'openzeppelin-solidity': minor
---
`P256`: Library for verification and public key recovery of P256 (aka secp256r1) signatures.

View File

@ -25,6 +25,7 @@ import {ERC721Holder} from "../token/ERC721/utils/ERC721Holder.sol";
import {Math} from "../utils/math/Math.sol";
import {MerkleProof} from "../utils/cryptography/MerkleProof.sol";
import {MessageHashUtils} from "../utils/cryptography/MessageHashUtils.sol";
import {P256} from "../utils/cryptography/P256.sol";
import {Panic} from "../utils/Panic.sol";
import {Packing} from "../utils/Packing.sol";
import {RSA} from "../utils/cryptography/RSA.sol";

View File

@ -23,4 +23,9 @@ library Errors {
* @dev The deployment failed.
*/
error FailedDeployment();
/**
* @dev A necessary precompile is missing.
*/
error MissingPrecompile(address);
}

View File

@ -8,6 +8,8 @@ Miscellaneous contracts and libraries containing utility functions you can use t
* {Math}, {SignedMath}: Implementation of various arithmetic functions.
* {SafeCast}: Checked downcasting functions to avoid silent truncation.
* {ECDSA}, {MessageHashUtils}: Libraries for interacting with ECDSA signatures.
* {P256}: Library for verifying and recovering public keys from secp256r1 signatures.
* {RSA}: Library with RSA PKCS#1 v1.5 signature verification utilities.
* {SignatureChecker}: A library helper to support regular ECDSA from EOAs as well as ERC-1271 signatures for smart contracts.
* {Hashes}: Commonly used hash functions.
* {MerkleProof}: Functions for verifying https://en.wikipedia.org/wiki/Merkle_tree[Merkle Tree] proofs.

View File

@ -0,0 +1,316 @@
// SPDX-License-Identifier: MIT
pragma solidity ^0.8.20;
import {Math} from "../math/Math.sol";
import {Errors} from "../Errors.sol";
/**
* @dev Implementation of secp256r1 verification and recovery functions.
*
* The secp256r1 curve (also known as P256) is a NIST standard curve with wide support in modern devices
* and cryptographic standards. Some notable examples include Apple's Secure Enclave and Android's Keystore
* as well as authentication protocols like FIDO2.
*
* Based on the original https://github.com/itsobvioustech/aa-passkeys-wallet/blob/main/src/Secp256r1.sol[implementation of itsobvioustech].
* Heavily inspired in https://github.com/maxrobot/elliptic-solidity/blob/master/contracts/Secp256r1.sol[maxrobot] and
* https://github.com/tdrerup/elliptic-curve-solidity/blob/master/contracts/curves/EllipticCurve.sol[tdrerup] implementations.
*/
library P256 {
struct JPoint {
uint256 x;
uint256 y;
uint256 z;
}
/// @dev Generator (x component)
uint256 internal constant GX = 0x6B17D1F2E12C4247F8BCE6E563A440F277037D812DEB33A0F4A13945D898C296;
/// @dev Generator (y component)
uint256 internal constant GY = 0x4FE342E2FE1A7F9B8EE7EB4A7C0F9E162BCE33576B315ECECBB6406837BF51F5;
/// @dev P (size of the field)
uint256 internal constant P = 0xFFFFFFFF00000001000000000000000000000000FFFFFFFFFFFFFFFFFFFFFFFF;
/// @dev N (order of G)
uint256 internal constant N = 0xFFFFFFFF00000000FFFFFFFFFFFFFFFFBCE6FAADA7179E84F3B9CAC2FC632551;
/// @dev A parameter of the weierstrass equation
uint256 internal constant A = 0xFFFFFFFF00000001000000000000000000000000FFFFFFFFFFFFFFFFFFFFFFFC;
/// @dev B parameter of the weierstrass equation
uint256 internal constant B = 0x5AC635D8AA3A93E7B3EBBD55769886BC651D06B0CC53B0F63BCE3C3E27D2604B;
/// @dev (P + 1) / 4. Useful to compute sqrt
uint256 private constant P1DIV4 = 0x3fffffffc0000000400000000000000000000000400000000000000000000000;
/// @dev N/2 for excluding higher order `s` values
uint256 private constant HALF_N = 0x7fffffff800000007fffffffffffffffde737d56d38bcf4279dce5617e3192a8;
/**
* @dev Verifies a secp256r1 signature using the RIP-7212 precompile and falls back to the Solidity implementation
* if the precompile is not available. This version should work on all chains, but requires the deployment of more
* bytecode.
*
* @param h - hashed message
* @param r - signature half R
* @param s - signature half S
* @param qx - public key coordinate X
* @param qy - public key coordinate Y
*
* IMPORTANT: This function disallows signatures where the `s` value is above `N/2` to prevent malleability.
* To flip the `s` value, compute `s = N - s`.
*/
function verify(bytes32 h, bytes32 r, bytes32 s, bytes32 qx, bytes32 qy) internal view returns (bool) {
(bool valid, bool supported) = _tryVerifyNative(h, r, s, qx, qy);
return supported ? valid : verifySolidity(h, r, s, qx, qy);
}
/**
* @dev Same as {verify}, but it will revert if the required precompile is not available.
*
* Make sure any logic (code or precompile) deployed at that address is the expected one,
* otherwise the returned value may be misinterpreted as a positive boolean.
*/
function verifyNative(bytes32 h, bytes32 r, bytes32 s, bytes32 qx, bytes32 qy) internal view returns (bool) {
(bool valid, bool supported) = _tryVerifyNative(h, r, s, qx, qy);
if (supported) {
return valid;
} else {
revert Errors.MissingPrecompile(address(0x100));
}
}
/**
* @dev Same as {verify}, but it will return false if the required precompile is not available.
*/
function _tryVerifyNative(
bytes32 h,
bytes32 r,
bytes32 s,
bytes32 qx,
bytes32 qy
) private view returns (bool valid, bool supported) {
if (!_isProperSignature(r, s) || !isValidPublicKey(qx, qy)) {
return (false, true); // signature is invalid, and its not because the precompile is missing
}
(bool success, bytes memory returndata) = address(0x100).staticcall(abi.encode(h, r, s, qx, qy));
return (success && returndata.length == 0x20) ? (abi.decode(returndata, (bool)), true) : (false, false);
}
/**
* @dev Same as {verify}, but only the Solidity implementation is used.
*/
function verifySolidity(bytes32 h, bytes32 r, bytes32 s, bytes32 qx, bytes32 qy) internal view returns (bool) {
if (!_isProperSignature(r, s) || !isValidPublicKey(qx, qy)) {
return false;
}
JPoint[16] memory points = _preComputeJacobianPoints(uint256(qx), uint256(qy));
uint256 w = Math.invModPrime(uint256(s), N);
uint256 u1 = mulmod(uint256(h), w, N);
uint256 u2 = mulmod(uint256(r), w, N);
(uint256 x, ) = _jMultShamir(points, u1, u2);
return ((x % N) == uint256(r));
}
/**
* @dev Public key recovery
*
* @param h - hashed message
* @param v - signature recovery param
* @param r - signature half R
* @param s - signature half S
*
* IMPORTANT: This function disallows signatures where the `s` value is above `N/2` to prevent malleability.
* To flip the `s` value, compute `s = N - s` and `v = 1 - v` if (`v = 0 | 1`).
*/
function recovery(bytes32 h, uint8 v, bytes32 r, bytes32 s) internal view returns (bytes32, bytes32) {
if (!_isProperSignature(r, s) || v > 1) {
return (0, 0);
}
uint256 rx = uint256(r);
uint256 ry2 = addmod(mulmod(addmod(mulmod(rx, rx, P), A, P), rx, P), B, P); // weierstrass equation y² = x³ + a.x + b
uint256 ry = Math.modExp(ry2, P1DIV4, P); // This formula for sqrt work because P ≡ 3 (mod 4)
if (mulmod(ry, ry, P) != ry2) return (0, 0); // Sanity check
if (ry % 2 != v % 2) ry = P - ry;
JPoint[16] memory points = _preComputeJacobianPoints(rx, ry);
uint256 w = Math.invModPrime(uint256(r), N);
uint256 u1 = mulmod(N - (uint256(h) % N), w, N);
uint256 u2 = mulmod(uint256(s), w, N);
(uint256 x, uint256 y) = _jMultShamir(points, u1, u2);
return (bytes32(x), bytes32(y));
}
/**
* @dev Checks if (x, y) are valid coordinates of a point on the curve.
* In particular this function checks that x <= P and y <= P.
*/
function isValidPublicKey(bytes32 x, bytes32 y) internal pure returns (bool result) {
assembly ("memory-safe") {
let p := P
let lhs := mulmod(y, y, p) // y^2
let rhs := addmod(mulmod(addmod(mulmod(x, x, p), A, p), x, p), B, p) // ((x^2 + a) * x) + b = x^3 + ax + b
result := and(and(lt(x, p), lt(y, p)), eq(lhs, rhs)) // Should conform with the Weierstrass equation
}
}
/**
* @dev Checks if (r, s) is a proper signature.
* In particular, this checks that `s` is in the "lower-range", making the signature non-malleable.
*/
function _isProperSignature(bytes32 r, bytes32 s) private pure returns (bool) {
return uint256(r) > 0 && uint256(r) < N && uint256(s) > 0 && uint256(s) <= HALF_N;
}
/**
* @dev Reduce from jacobian to affine coordinates
* @param jx - jacobian coordinate x
* @param jy - jacobian coordinate y
* @param jz - jacobian coordinate z
* @return ax - affine coordinate x
* @return ay - affine coordinate y
*/
function _affineFromJacobian(uint256 jx, uint256 jy, uint256 jz) private view returns (uint256 ax, uint256 ay) {
if (jz == 0) return (0, 0);
uint256 zinv = Math.invModPrime(jz, P);
uint256 zzinv = mulmod(zinv, zinv, P);
uint256 zzzinv = mulmod(zzinv, zinv, P);
ax = mulmod(jx, zzinv, P);
ay = mulmod(jy, zzzinv, P);
}
/**
* @dev Point addition on the jacobian coordinates
* Reference: https://www.hyperelliptic.org/EFD/g1p/auto-shortw-jacobian.html#addition-add-1998-cmo-2
*/
function _jAdd(
JPoint memory p1,
uint256 x2,
uint256 y2,
uint256 z2
) private pure returns (uint256 rx, uint256 ry, uint256 rz) {
assembly ("memory-safe") {
let p := P
let z1 := mload(add(p1, 0x40))
let s1 := mulmod(mload(add(p1, 0x20)), mulmod(mulmod(z2, z2, p), z2, p), p) // s1 = y1*z2³
let s2 := mulmod(y2, mulmod(mulmod(z1, z1, p), z1, p), p) // s2 = y2*z1³
let r := addmod(s2, sub(p, s1), p) // r = s2-s1
let u1 := mulmod(mload(p1), mulmod(z2, z2, p), p) // u1 = x1*z2²
let u2 := mulmod(x2, mulmod(z1, z1, p), p) // u2 = x2*z1²
let h := addmod(u2, sub(p, u1), p) // h = u2-u1
let hh := mulmod(h, h, p) // h²
// x' = r²-h³-2*u1*h²
rx := addmod(
addmod(mulmod(r, r, p), sub(p, mulmod(h, hh, p)), p),
sub(p, mulmod(2, mulmod(u1, hh, p), p)),
p
)
// y' = r*(u1*h²-x')-s1*h³
ry := addmod(
mulmod(r, addmod(mulmod(u1, hh, p), sub(p, rx), p), p),
sub(p, mulmod(s1, mulmod(h, hh, p), p)),
p
)
// z' = h*z1*z2
rz := mulmod(h, mulmod(z1, z2, p), p)
}
}
/**
* @dev Point doubling on the jacobian coordinates
* Reference: https://www.hyperelliptic.org/EFD/g1p/auto-shortw-jacobian.html#doubling-dbl-1998-cmo-2
*/
function _jDouble(uint256 x, uint256 y, uint256 z) private pure returns (uint256 rx, uint256 ry, uint256 rz) {
assembly ("memory-safe") {
let p := P
let yy := mulmod(y, y, p)
let zz := mulmod(z, z, p)
let s := mulmod(4, mulmod(x, yy, p), p) // s = 4*x*y²
let m := addmod(mulmod(3, mulmod(x, x, p), p), mulmod(A, mulmod(zz, zz, p), p), p) // m = 3*x²+a*z⁴
let t := addmod(mulmod(m, m, p), sub(p, mulmod(2, s, p)), p) // t = m²-2*s
// x' = t
rx := t
// y' = m*(s-t)-8*y⁴
ry := addmod(mulmod(m, addmod(s, sub(p, t), p), p), sub(p, mulmod(8, mulmod(yy, yy, p), p)), p)
// z' = 2*y*z
rz := mulmod(2, mulmod(y, z, p), p)
}
}
/**
* @dev Compute P·u1 + Q·u2 using the precomputed points for P and Q (see {_preComputeJacobianPoints}).
*
* Uses Strauss Shamir trick for EC multiplication
* https://stackoverflow.com/questions/50993471/ec-scalar-multiplication-with-strauss-shamir-method
* we optimise on this a bit to do with 2 bits at a time rather than a single bit
* the individual points for a single pass are precomputed
* overall this reduces the number of additions while keeping the same number of doublings
*/
function _jMultShamir(JPoint[16] memory points, uint256 u1, uint256 u2) private view returns (uint256, uint256) {
uint256 x = 0;
uint256 y = 0;
uint256 z = 0;
unchecked {
for (uint256 i = 0; i < 128; ++i) {
if (z > 0) {
(x, y, z) = _jDouble(x, y, z);
(x, y, z) = _jDouble(x, y, z);
}
// Read 2 bits of u1, and 2 bits of u2. Combining the two give a lookup index in the table.
uint256 pos = ((u1 >> 252) & 0xc) | ((u2 >> 254) & 0x3);
if (pos > 0) {
if (z == 0) {
(x, y, z) = (points[pos].x, points[pos].y, points[pos].z);
} else {
(x, y, z) = _jAdd(points[pos], x, y, z);
}
}
u1 <<= 2;
u2 <<= 2;
}
}
return _affineFromJacobian(x, y, z);
}
/**
* @dev Precompute a matrice of useful jacobian points associated with a given P. This can be seen as a 4x4 matrix
* that contains combination of P and G (generator) up to 3 times each. See the table below:
*
* ┌────┬─────────────────────┐
* │ i │ 0 1 2 3 │
* ├────┼─────────────────────┤
* │ 0 │ 0 p 2p 3p │
* │ 4 │ g g+p g+2p g+3p │
* │ 8 │ 2g 2g+p 2g+2p 2g+3p │
* │ 12 │ 3g 3g+p 3g+2p 3g+3p │
* └────┴─────────────────────┘
*/
function _preComputeJacobianPoints(uint256 px, uint256 py) private pure returns (JPoint[16] memory points) {
points[0x00] = JPoint(0, 0, 0); // 0,0
points[0x01] = JPoint(px, py, 1); // 1,0 (p)
points[0x04] = JPoint(GX, GY, 1); // 0,1 (g)
points[0x02] = _jDoublePoint(points[0x01]); // 2,0 (2p)
points[0x08] = _jDoublePoint(points[0x04]); // 0,2 (2g)
points[0x03] = _jAddPoint(points[0x01], points[0x02]); // 3,0 (3p)
points[0x05] = _jAddPoint(points[0x01], points[0x04]); // 1,1 (p+g)
points[0x06] = _jAddPoint(points[0x02], points[0x04]); // 2,1 (2p+g)
points[0x07] = _jAddPoint(points[0x03], points[0x04]); // 3,1 (3p+g)
points[0x09] = _jAddPoint(points[0x01], points[0x08]); // 1,2 (p+2g)
points[0x0a] = _jAddPoint(points[0x02], points[0x08]); // 2,2 (2p+2g)
points[0x0b] = _jAddPoint(points[0x03], points[0x08]); // 3,2 (3p+2g)
points[0x0c] = _jAddPoint(points[0x04], points[0x08]); // 0,3 (g+2g)
points[0x0d] = _jAddPoint(points[0x01], points[0x0c]); // 1,3 (p+3g)
points[0x0e] = _jAddPoint(points[0x02], points[0x0c]); // 2,3 (2p+3g)
points[0x0f] = _jAddPoint(points[0x03], points[0x0C]); // 3,3 (3p+3g)
}
function _jAddPoint(JPoint memory p1, JPoint memory p2) private pure returns (JPoint memory) {
(uint256 x, uint256 y, uint256 z) = _jAdd(p1, p2.x, p2.y, p2.z);
return JPoint(x, y, z);
}
function _jDoublePoint(JPoint memory p) private pure returns (JPoint memory) {
(uint256 x, uint256 y, uint256 z) = _jDouble(p.x, p.y, p.z);
return JPoint(x, y, z);
}
}

View File

@ -237,8 +237,8 @@ library Math {
*
* If the input value is not inversible, 0 is returned.
*
* NOTE: If you know for sure that n is (big) a prime, it may be cheaper to use Ferma's little theorem and get the
* inverse using `Math.modExp(a, n - 2, n)`.
* NOTE: If you know for sure that n is (big) a prime, it may be cheaper to use Fermat's little theorem and get the
* inverse using `Math.modExp(a, n - 2, n)`. See {invModPrime}.
*/
function invMod(uint256 a, uint256 n) internal pure returns (uint256) {
unchecked {
@ -288,6 +288,21 @@ library Math {
}
}
/**
* @dev Variant of {invMod}. More efficient, but only works if `p` is known to be a prime greater than `2`.
*
* From https://en.wikipedia.org/wiki/Fermat%27s_little_theorem[Fermat's little theorem], we know that if p is
* prime, then `a**(p-1) ≡ 1 mod p`. As a consequence, we have `a * a**(p-2) ≡ 1 mod p`, which means that
* `a**(p-2)` is the modular multiplicative inverse of a in Fp.
*
* NOTE: this function does NOT check that `p` is a prime greater than `2`.
*/
function invModPrime(uint256 a, uint256 p) internal view returns (uint256) {
unchecked {
return Math.modExp(a, p - 2, p);
}
}
/**
* @dev Returns the modular exponentiation of the specified base, exponent and modulus (b ** e % m)
*

View File

@ -8,9 +8,9 @@ Here are some of the more popular ones.
=== Checking Signatures On-Chain
At a high level, signatures are a set of cryptographic algorithms that allow for a _signer_ to prove himself owner of a _private key_ used to authorize an piece of information (generally a transaction or `UserOperation`). Natively, the EVM supports the Elliptic Curve Digital Signature Algorithm (https://en.wikipedia.org/wiki/Elliptic_Curve_Digital_Signature_Algorithm[ECDSA]) using the secp256r1 field, however other signature algorithms such as RSA are supported.
At a high level, signatures are a set of cryptographic algorithms that allow for a _signer_ to prove himself owner of a _private key_ used to authorize an piece of information (generally a transaction or `UserOperation`). Natively, the EVM supports the Elliptic Curve Digital Signature Algorithm (https://en.wikipedia.org/wiki/Elliptic_Curve_Digital_Signature_Algorithm[ECDSA]) using the secp256k1 curve, however other signature algorithms such as P256 and RSA are supported.
==== Ethereum Signatures (secp256r1)
==== Ethereum Signatures (secp256k1)
xref:api:utils.adoc#ECDSA[`ECDSA`] provides functions for recovering and managing Ethereum account ECDSA signatures. These are often generated via https://web3js.readthedocs.io/en/v1.7.3/web3-eth.html#sign[`web3.eth.sign`], and are a 65 byte array (of type `bytes` in Solidity) arranged the following way: `[[v (1)], [r (32)], [s (32)]]`.
@ -30,6 +30,47 @@ function _verify(bytes32 data, bytes memory signature, address account) internal
WARNING: Getting signature verification right is not trivial: make sure you fully read and understand xref:api:utils.adoc#MessageHashUtils[`MessageHashUtils`]'s and xref:api:utils.adoc#ECDSA[`ECDSA`]'s documentation.
==== P256 Signatures (secp256r1)
P256, also known as secp256r1, is one of the most used signature schemes. P256 signatures are standardized by the National Institute of Standards and Technology (NIST) and it's widely available in consumer hardware and software.
These signatures are different to regular Ethereum Signatures (secp256k1) in that they use a different elliptic curve to perform operations but have similar security guarantees.
[source,solidity]
----
using P256 for bytes32;
function _verify(
bytes32 data,
bytes32 r,
bytes32 s,
bytes32 qx,
bytes32 qy
) internal pure returns (bool) {
return data.verify(data, r, s, qx, qy);
}
----
By default, the `verify` function will try calling the (https://github.com/ethereum/RIPs/blob/master/RIPS/rip-7212.md)[RIP-7212] precompile at address `0x100` and will fallback to an implementation in Solidity if not available. We encourage you to use `verifyNative` if you know the precompile is available on the chain you're working on and on any other chain on which you intend to use the same bytecode in the future. In case of any doubts regarding the implementation roadmap of the native precompile `P256` of potential future target chains, please consider using `verifySolidity`.
[source,solidity]
----
using P256 for bytes32;
function _verify(
bytes32 data,
bytes32 r,
bytes32 s,
bytes32 qx,
bytes32 qy
) internal pure returns (bool) {
// Will only call the precompile at address(0x100)
return data.verifyNative(data, r, s, qx, qy);
}
----
IMPORTANT: The P256 library only allows for `s` values in the lower order of the curve (i.e. `s <= N/2`) to prevent malleability. In case your tooling produces signatures in both sides of the curve, consider flipping the `s` value to keep compatibility.
==== RSA
RSA a public-key cryptosystem that was popularized by corporate and governmental public key infrastructures (https://en.wikipedia.org/wiki/Public_key_infrastructure[PKIs]) and https://en.wikipedia.org/wiki/Domain_Name_System_Security_Extensions[DNSSEC].

View File

@ -10,7 +10,7 @@ hardhat coverage
if [ "${CI:-"false"}" == "true" ]; then
# Foundry coverage
forge coverage --report lcov
forge coverage --report lcov --ir-minimum
# Remove zero hits
sed -i '/,0/d' lcov.info
fi

View File

@ -0,0 +1,135 @@
// SPDX-License-Identifier: MIT
pragma solidity ^0.8.20;
import {Test} from "forge-std/Test.sol";
import {P256} from "@openzeppelin/contracts/utils/cryptography/P256.sol";
import {Math} from "@openzeppelin/contracts/utils/math/Math.sol";
contract P256Test is Test {
/// forge-config: default.fuzz.runs = 512
function testVerify(uint256 seed, bytes32 digest) public {
uint256 privateKey = bound(uint256(keccak256(abi.encode(seed))), 1, P256.N - 1);
(bytes32 x, bytes32 y) = P256PublicKey.getPublicKey(privateKey);
(bytes32 r, bytes32 s) = vm.signP256(privateKey, digest);
s = _ensureLowerS(s);
assertTrue(P256.verify(digest, r, s, x, y));
assertTrue(P256.verifySolidity(digest, r, s, x, y));
}
/// forge-config: default.fuzz.runs = 512
function testRecover(uint256 seed, bytes32 digest) public {
uint256 privateKey = bound(uint256(keccak256(abi.encode(seed))), 1, P256.N - 1);
(bytes32 x, bytes32 y) = P256PublicKey.getPublicKey(privateKey);
(bytes32 r, bytes32 s) = vm.signP256(privateKey, digest);
s = _ensureLowerS(s);
(bytes32 qx0, bytes32 qy0) = P256.recovery(digest, 0, r, s);
(bytes32 qx1, bytes32 qy1) = P256.recovery(digest, 1, r, s);
assertTrue((qx0 == x && qy0 == y) || (qx1 == x && qy1 == y));
}
function _ensureLowerS(bytes32 s) private pure returns (bytes32) {
uint256 _s = uint256(s);
unchecked {
return _s > P256.N / 2 ? bytes32(P256.N - _s) : s;
}
}
}
/**
* @dev Library to derive P256 public key from private key
* Should be removed if Foundry adds this functionality
* See https://github.com/foundry-rs/foundry/issues/7908
*/
library P256PublicKey {
function getPublicKey(uint256 privateKey) internal view returns (bytes32, bytes32) {
(uint256 x, uint256 y, uint256 z) = _jMult(P256.GX, P256.GY, 1, privateKey);
return _affineFromJacobian(x, y, z);
}
function _jMult(
uint256 x,
uint256 y,
uint256 z,
uint256 k
) private pure returns (uint256 rx, uint256 ry, uint256 rz) {
unchecked {
for (uint256 i = 0; i < 256; ++i) {
if (rz > 0) {
(rx, ry, rz) = _jDouble(rx, ry, rz);
}
if (k >> 255 > 0) {
if (rz == 0) {
(rx, ry, rz) = (x, y, z);
} else {
(rx, ry, rz) = _jAdd(rx, ry, rz, x, y, z);
}
}
k <<= 1;
}
}
}
/// From P256.sol
function _affineFromJacobian(uint256 jx, uint256 jy, uint256 jz) private view returns (bytes32 ax, bytes32 ay) {
if (jz == 0) return (0, 0);
uint256 zinv = Math.invModPrime(jz, P256.P);
uint256 zzinv = mulmod(zinv, zinv, P256.P);
uint256 zzzinv = mulmod(zzinv, zinv, P256.P);
ax = bytes32(mulmod(jx, zzinv, P256.P));
ay = bytes32(mulmod(jy, zzzinv, P256.P));
}
function _jDouble(uint256 x, uint256 y, uint256 z) private pure returns (uint256 rx, uint256 ry, uint256 rz) {
uint256 p = P256.P;
uint256 a = P256.A;
assembly ("memory-safe") {
let yy := mulmod(y, y, p)
let zz := mulmod(z, z, p)
let s := mulmod(4, mulmod(x, yy, p), p) // s = 4*x*y²
let m := addmod(mulmod(3, mulmod(x, x, p), p), mulmod(a, mulmod(zz, zz, p), p), p) // m = 3*x²+a*z⁴
let t := addmod(mulmod(m, m, p), sub(p, mulmod(2, s, p)), p) // t = m²-2*s
// x' = t
rx := t
// y' = m*(s-t)-8*y⁴
ry := addmod(mulmod(m, addmod(s, sub(p, t), p), p), sub(p, mulmod(8, mulmod(yy, yy, p), p)), p)
// z' = 2*y*z
rz := mulmod(2, mulmod(y, z, p), p)
}
}
function _jAdd(
uint256 x1,
uint256 y1,
uint256 z1,
uint256 x2,
uint256 y2,
uint256 z2
) private pure returns (uint256 rx, uint256 ry, uint256 rz) {
uint256 p = P256.P;
assembly ("memory-safe") {
let zz1 := mulmod(z1, z1, p) // zz1 = z1²
let zz2 := mulmod(z2, z2, p) // zz2 = z2²
let u1 := mulmod(x1, zz2, p) // u1 = x1*z2²
let u2 := mulmod(x2, zz1, p) // u2 = x2*z1²
let s1 := mulmod(y1, mulmod(zz2, z2, p), p) // s1 = y1*z2³
let s2 := mulmod(y2, mulmod(zz1, z1, p), p) // s2 = y2*z1³
let h := addmod(u2, sub(p, u1), p) // h = u2-u1
let hh := mulmod(h, h, p) // h²
let hhh := mulmod(h, hh, p) // h³
let r := addmod(s2, sub(p, s1), p) // r = s2-s1
// x' = r²-h³-2*u1*h²
rx := addmod(addmod(mulmod(r, r, p), sub(p, hhh), p), sub(p, mulmod(2, mulmod(u1, hh, p), p)), p)
// y' = r*(u1*h²-x')-s1*h³
ry := addmod(mulmod(r, addmod(mulmod(u1, hh, p), sub(p, rx), p), p), sub(p, mulmod(s1, hhh, p)), p)
// z' = h*z1*z2
rz := mulmod(h, mulmod(z1, z2, p), p)
}
}
}

View File

@ -0,0 +1,156 @@
const { ethers } = require('hardhat');
const { expect } = require('chai');
const { secp256r1 } = require('@noble/curves/p256');
const { loadFixture } = require('@nomicfoundation/hardhat-network-helpers');
const N = 0xffffffff00000000ffffffffffffffffbce6faada7179e84f3b9cac2fc632551n;
// As in ECDSA, signatures are malleable and the tooling produce both high and low S values.
// We need to ensure that the s value is in the lower half of the order of the curve.
const ensureLowerOrderS = ({ s, recovery, ...rest }) => {
if (s > N / 2n) {
s = N - s;
recovery = 1 - recovery;
}
return { s, recovery, ...rest };
};
const prepareSignature = (
privateKey = secp256r1.utils.randomPrivateKey(),
messageHash = ethers.hexlify(ethers.randomBytes(0x20)),
) => {
const publicKey = [
secp256r1.getPublicKey(privateKey, false).slice(0x01, 0x21),
secp256r1.getPublicKey(privateKey, false).slice(0x21, 0x41),
].map(ethers.hexlify);
const { r, s, recovery } = ensureLowerOrderS(secp256r1.sign(messageHash.replace(/0x/, ''), privateKey));
const signature = [r, s].map(v => ethers.toBeHex(v, 0x20));
return { privateKey, publicKey, signature, recovery, messageHash };
};
describe('P256', function () {
async function fixture() {
return { mock: await ethers.deployContract('$P256') };
}
beforeEach(async function () {
Object.assign(this, await loadFixture(fixture));
});
describe('with signature', function () {
beforeEach(async function () {
Object.assign(this, prepareSignature());
});
it('verify valid signature', async function () {
expect(await this.mock.$verify(this.messageHash, ...this.signature, ...this.publicKey)).to.be.true;
expect(await this.mock.$verifySolidity(this.messageHash, ...this.signature, ...this.publicKey)).to.be.true;
await expect(this.mock.$verifyNative(this.messageHash, ...this.signature, ...this.publicKey))
.to.be.revertedWithCustomError(this.mock, 'MissingPrecompile')
.withArgs('0x0000000000000000000000000000000000000100');
});
it('recover public key', async function () {
expect(await this.mock.$recovery(this.messageHash, this.recovery, ...this.signature)).to.deep.equal(
this.publicKey,
);
});
it('reject signature with flipped public key coordinates ([x,y] >> [y,x])', async function () {
// flip public key
this.publicKey.reverse();
expect(await this.mock.$verify(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
expect(await this.mock.$verifySolidity(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
expect(await this.mock.$verifyNative(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false; // Flipped public key is not in the curve
});
it('reject signature with flipped signature values ([r,s] >> [s,r])', async function () {
// Preselected signature where `r < N/2` and `s < N/2`
this.signature = [
'0x45350225bad31e89db662fcc4fb2f79f349adbb952b3f652eed1f2aa72fb0356',
'0x513eb68424c42630012309eee4a3b43e0bdc019d179ef0e0c461800845e237ee',
];
// Corresponding hash and public key
this.messageHash = '0x2ad1f900fe63745deeaedfdf396cb6f0f991c4338a9edf114d52f7d1812040a0';
this.publicKey = [
'0x9e30de165e521257996425d9bf12a7d366925614bf204eabbb78172b48e52e59',
'0x94bf0fe72f99654d7beae4780a520848e306d46a1275b965c4f4c2b8e9a2c08d',
];
// Make sure it works
expect(await this.mock.$verify(this.messageHash, ...this.signature, ...this.publicKey)).to.be.true;
// Flip signature
this.signature.reverse();
expect(await this.mock.$verify(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
expect(await this.mock.$verifySolidity(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
await expect(this.mock.$verifyNative(this.messageHash, ...this.signature, ...this.publicKey))
.to.be.revertedWithCustomError(this.mock, 'MissingPrecompile')
.withArgs('0x0000000000000000000000000000000000000100');
expect(await this.mock.$recovery(this.messageHash, this.recovery, ...this.signature)).to.not.deep.equal(
this.publicKey,
);
});
it('reject signature with invalid message hash', async function () {
// random message hash
this.messageHash = ethers.hexlify(ethers.randomBytes(32));
expect(await this.mock.$verify(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
expect(await this.mock.$verifySolidity(this.messageHash, ...this.signature, ...this.publicKey)).to.be.false;
await expect(this.mock.$verifyNative(this.messageHash, ...this.signature, ...this.publicKey))
.to.be.revertedWithCustomError(this.mock, 'MissingPrecompile')
.withArgs('0x0000000000000000000000000000000000000100');
expect(await this.mock.$recovery(this.messageHash, this.recovery, ...this.signature)).to.not.deep.equal(
this.publicKey,
);
});
it('fail to recover signature with invalid recovery bit', async function () {
// flip recovery bit
this.recovery = 1 - this.recovery;
expect(await this.mock.$recovery(this.messageHash, this.recovery, ...this.signature)).to.not.deep.equal(
this.publicKey,
);
});
});
// test cases for https://github.com/C2SP/wycheproof/blob/4672ff74d68766e7785c2cac4c597effccef2c5c/testvectors/ecdsa_secp256r1_sha256_p1363_test.json
describe('wycheproof tests', function () {
for (const { key, tests } of require('./ecdsa_secp256r1_sha256_p1363_test.json').testGroups) {
// parse public key
let [x, y] = [key.wx, key.wy].map(v => ethers.stripZerosLeft('0x' + v, 32));
if (x.length > 66 || y.length > 66) continue;
x = ethers.zeroPadValue(x, 32);
y = ethers.zeroPadValue(y, 32);
// run all tests for this key
for (const { tcId, comment, msg, sig, result } of tests) {
// only keep properly formatted signatures
if (sig.length != 128) continue;
it(`${tcId}: ${comment}`, async function () {
// split signature, and reduce modulo N
let [r, s] = Array(2)
.fill()
.map((_, i) => ethers.toBigInt('0x' + sig.substring(64 * i, 64 * (i + 1))));
// move s to lower part of the curve if needed
if (s <= N && s > N / 2n) s = N - s;
// prepare signature
r = ethers.toBeHex(r, 32);
s = ethers.toBeHex(s, 32);
// hash
const messageHash = ethers.sha256('0x' + msg);
// check verify
expect(await this.mock.$verify(messageHash, r, s, x, y)).to.equal(result == 'valid');
});
}
}
});
});

File diff suppressed because it is too large Load Diff